utm/abacus.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
header {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     2
  {\em Abacus} (a kind of register machine) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     3
*}
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
theory abacus
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     6
imports Main turing_basic
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     7
begin
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     8
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     9
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    10
  {\em Abacus} instructions:
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    13
datatype abc_inst =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    14
  -- {* @{text "Inc n"} increments the memory cell (or register) with address @{text "n"} by one.
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
     Inc nat
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
     @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    19
      If cell @{text "n"} is already zero, no decrements happens and the executio jumps to
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    20
      the instruction labeled by @{text "label"}.
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
   | Dec nat nat
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    23
  -- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    24
  @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}.
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
   | Goto nat
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    27
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    28
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    29
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    30
  Abacus programs are defined as lists of Abacus instructions.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    31
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    32
type_synonym abc_prog = "abc_inst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    33
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    34
section {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    35
  Sample Abacus programs
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    36
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    37
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    38
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    39
  Abacus for addition and clearance.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    40
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    41
fun plus_clear :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    42
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    43
  "plus_clear m n e = [Dec m e, Inc n, Goto 0]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    44
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    45
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    46
  Abacus for clearing memory untis.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    47
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    48
fun clear :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    49
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    50
  "clear n e = [Dec n e, Goto 0]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    51
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    52
fun plus:: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    53
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    54
  "plus m n p e = [Dec m 4, Inc n, Inc p,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    55
                   Goto 0, Dec p e, Inc m, Goto 4]"
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
fun mult :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    58
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    59
  "mult m1 m2 n p e = [Dec m1 e]@ plus m1 m2 p 1"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    60
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    61
fun expo :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    62
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    63
  "expo n m1 m2 p e = [Inc n, Dec m1 e] @ mult m2 n n p 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    64
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    65
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    66
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    67
  The state of Abacus machine.
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 abc_state = nat
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
  The memory of Abacus machine is defined as a function from address to contents.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    73
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    74
type_synonym abc_mem = "nat \<Rightarrow> nat" *)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    75
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    76
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    77
  The memory of Abacus machine is defined as a list of contents, with 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    78
  every units addressed by index into the list.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    79
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    80
type_synonym abc_lm = "nat list"
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
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    83
  Fetching contents out of memory. Units not represented by list elements are considered
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    84
  as having content @{text "0"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    85
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    86
fun abc_lm_v :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    87
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    88
    "abc_lm_v lm n = (if (n < length lm) then (lm!n) else 0)"         
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    89
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    90
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    91
fun abc_l2m :: "abc_lm \<Rightarrow> abc_mem"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    92
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    93
    "abc_l2m lm = abc_lm_v lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    94
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    95
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    96
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    97
  Set the content of memory unit @{text "n"} to value @{text "v"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    98
  @{text "am"} is the Abacus memory before setting.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    99
  If address @{text "n"} is outside to scope of @{text "am"}, @{text "am"} 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   100
  is extended so that @{text "n"} becomes in scope.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   101
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   102
fun abc_lm_s :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   103
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   104
    "abc_lm_s am n v = (if (n < length am) then (am[n:=v]) else 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   105
                           am@ (replicate (n - length am) 0) @ [v])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   106
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   107
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   108
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   109
  The configuration of Abaucs machines consists of its current state and its
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   110
  current memory:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   111
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   112
type_synonym abc_conf_l = "abc_state \<times> abc_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   113
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   114
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   115
  Fetch instruction out of Abacus program:
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   118
fun abc_fetch :: "nat \<Rightarrow> abc_prog \<Rightarrow> abc_inst option" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   119
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   120
  "abc_fetch s p = (if (s < length p) then Some (p ! s)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   121
                    else None)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   122
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   123
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   124
  Single step execution of Abacus machine. If no instruction is feteched, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   125
  configuration does not change.
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
fun abc_step_l :: "abc_conf_l \<Rightarrow> abc_inst option \<Rightarrow> abc_conf_l"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   128
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   129
  "abc_step_l (s, lm) a = (case a of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   130
               None \<Rightarrow> (s, lm) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   131
               Some (Inc n)  \<Rightarrow> (let nv = abc_lm_v lm n in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   132
                       (s + 1, abc_lm_s lm n (nv + 1))) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   133
               Some (Dec n e) \<Rightarrow> (let nv = abc_lm_v lm n in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   134
                       if (nv = 0) then (e, abc_lm_s lm n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   135
                       else (s + 1,  abc_lm_s lm n (nv - 1))) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   136
               Some (Goto n) \<Rightarrow> (n, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   137
               )"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   138
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   139
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   140
  Multi-step execution of Abacus machine.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   141
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   142
fun abc_steps_l :: "abc_conf_l \<Rightarrow> abc_prog \<Rightarrow> nat \<Rightarrow> abc_conf_l"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   143
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   144
  "abc_steps_l (s, lm) p 0 = (s, lm)" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   145
  "abc_steps_l (s, lm) p (Suc n) = abc_steps_l (abc_step_l (s, lm) (abc_fetch s p)) p n"
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
section {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   148
  Compiling Abacus machines into Truing machines
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   149
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   150
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
subsection {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   153
  Compiling functions
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   154
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   155
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   156
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   157
  @{text "findnth n"} returns the TM which locates the represention of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   158
  memory cell @{text "n"} on the tape and changes representation of zero
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   159
  on the way.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   160
*}
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
fun findnth :: "nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   163
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   164
  "findnth 0 = []" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   165
  "findnth (Suc n) = (findnth n @ [(W1, 2 * n + 1), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   166
           (R, 2 * n + 2), (R, 2 * n + 3), (R, 2 * n + 2)])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   167
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   168
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   169
  @{text "tinc_b"} returns the TM which increments the representation 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   170
  of the memory cell under rw-head by one and move the representation 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   171
  of cells afterwards to the right accordingly.
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   174
definition tinc_b :: "tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   175
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   176
  "tinc_b \<equiv> [(W1, 1), (R, 2), (W1, 3), (R, 2), (W1, 3), (R, 4), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   177
             (L, 7), (W0, 5), (R, 6), (W0, 5), (W1, 3), (R, 6),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   178
             (L, 8), (L, 7), (R, 9), (L, 7), (R, 10), (W0, 9)]" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   179
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   180
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   181
  @{text "tshift tm off"} shifts @{text "tm"} by offset @{text "off"}, leaving 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   182
  instructions concerning state @{text "0"} unchanged, because state @{text "0"} 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   183
  is the end state, which needs not be changed with shift operation.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   184
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   185
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   186
fun tshift :: "tprog \<Rightarrow> nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   187
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   188
  "tshift tp off = (map (\<lambda> (action, state). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   189
       (action, (if state = 0 then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   190
                 else state + off))) tp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   191
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   192
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   193
  @{text "tinc ss n"} returns the TM which simulates the execution of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   194
  Abacus instruction @{text "Inc n"}, assuming that TM is located at
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   195
  location @{text "ss"} in the final TM complied from the whole
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   196
  Abacus program.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   197
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   198
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   199
fun tinc :: "nat \<Rightarrow> nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   200
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   201
  "tinc ss n = tshift (findnth n @ tshift tinc_b (2 * n)) (ss - 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   202
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   203
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   204
  @{text "tinc_b"} returns the TM which decrements the representation 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   205
  of the memory cell under rw-head by one and move the representation 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   206
  of cells afterwards to the left accordingly.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   207
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   208
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   209
definition tdec_b :: "tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   210
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   211
  "tdec_b \<equiv>  [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   212
              (R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   213
              (L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   214
              (R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   215
              (R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   216
              (R, 0), (W0, 16)]"
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
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   219
  @{text "sete tp e"} attaches the termination edges (edges leading to state @{text "0"}) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   220
  of TM @{text "tp"} to the intruction labelled by @{text "e"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   221
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   222
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   223
fun sete :: "tprog \<Rightarrow> nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   224
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   225
  "sete tp e = map (\<lambda> (action, state). (action, if state = 0 then e else state)) tp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   226
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   227
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   228
  @{text "tdec ss n label"} returns the TM which simulates the execution of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   229
  Abacus instruction @{text "Dec n label"}, assuming that TM is located at
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   230
  location @{text "ss"} in the final TM complied from the whole
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   231
  Abacus program.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   232
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   233
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   234
fun tdec :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   235
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   236
  "tdec ss n e = sete (tshift (findnth n @ tshift tdec_b (2 * n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   237
                 (ss - 1)) e"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   238
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   239
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   240
  @{text "tgoto f(label)"} returns the TM simulating the execution of Abacus instruction
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   241
  @{text "Goto label"}, where @{text "f(label)"} is the corresponding location of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   242
  @{text "label"} in the final TM compiled from the overall Abacus program.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   243
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   244
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   245
fun tgoto :: "nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   246
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   247
  "tgoto n = [(Nop, n), (Nop, n)]"
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 layout of the final TM compiled from an Abacus program is represented
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   251
  as a list of natural numbers, where the list element at index @{text "n"} represents the 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   252
  starting state of the TM simulating the execution of @{text "n"}-th instruction
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   253
  in the Abacus program.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   254
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   255
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   256
type_synonym layout = "nat list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   257
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   258
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   259
  @{text "length_of i"} is the length of the 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   260
  TM simulating the Abacus instruction @{text "i"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   261
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   262
fun length_of :: "abc_inst \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   263
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   264
  "length_of i = (case i of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   265
                    Inc n   \<Rightarrow> 2 * n + 9 |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   266
                    Dec n e \<Rightarrow> 2 * n + 16 |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   267
                    Goto n  \<Rightarrow> 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   268
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   269
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   270
  @{text "layout_of ap"} returns the layout of Abacus program @{text "ap"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   271
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   272
fun layout_of :: "abc_prog \<Rightarrow> layout"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   273
  where "layout_of ap = map length_of ap"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   274
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   275
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   276
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   277
  @{text "start_of layout n"} looks out the starting state of @{text "n"}-th
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   278
  TM in the finall TM.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   279
*}
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
fun start_of :: "nat list \<Rightarrow> nat \<Rightarrow> nat"
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
  "start_of ly 0 = Suc 0" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   284
  "start_of ly (Suc as) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   285
        (if as < length ly then start_of ly as + (ly ! as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   286
         else start_of ly as)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   287
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   288
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   289
  @{text "ci lo ss i"} complies Abacus instruction @{text "i"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   290
  assuming the TM of @{text "i"} starts from state @{text "ss"} 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   291
  within the overal layout @{text "lo"}.
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   294
fun ci :: "layout \<Rightarrow> nat \<Rightarrow> abc_inst \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   295
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   296
  "ci ly ss i = (case i of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   297
                    Inc n   \<Rightarrow> tinc ss n |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   298
                    Dec n e \<Rightarrow> tdec ss n (start_of ly e) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   299
                    Goto n  \<Rightarrow> tgoto (start_of ly n))"
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
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   302
  @{text "tpairs_of ap"} transfroms Abacus program @{text "ap"} pairing
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   303
  every instruction with its starting state.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   304
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   305
fun tpairs_of :: "abc_prog \<Rightarrow> (nat \<times> abc_inst) list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   306
  where "tpairs_of ap = (zip (map (start_of (layout_of ap)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   307
                         [0..<(length ap)]) ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   308
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   309
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   310
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   311
  @{text "tms_of ap"} returns the list of TMs, where every one of them simulates
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   312
  the corresponding Abacus intruction in @{text "ap"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   313
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   314
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   315
fun tms_of :: "abc_prog \<Rightarrow> tprog list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   316
  where "tms_of ap = map (\<lambda> (n, tm). ci (layout_of ap) n tm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   317
                         (tpairs_of ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   318
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   319
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   320
  @{text "tm_of ap"} returns the final TM machine compiled from Abacus program @{text "ap"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   321
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   322
fun tm_of :: "abc_prog \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   323
  where "tm_of ap = concat (tms_of ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   324
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   325
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   326
  The following two functions specify the well-formedness of complied TM.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   327
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   328
fun t_ncorrect :: "tprog \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   329
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   330
  "t_ncorrect tp = (length tp mod 2 = 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   331
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   332
fun abc2t_correct :: "abc_prog \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   333
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   334
  "abc2t_correct ap = list_all (\<lambda> (n, tm). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   335
             t_ncorrect (ci (layout_of ap) n tm)) (tpairs_of ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   336
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   337
lemma findnth_length: "length (findnth n) div 2 = 2 * n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   338
apply(induct n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   339
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   340
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   341
lemma ci_length : "length (ci ns n ai) div 2 = length_of ai"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   342
apply(auto simp: ci.simps tinc_b_def tdec_b_def findnth_length
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   343
                 split: abc_inst.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   344
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   345
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   346
subsection {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   347
  Representation of Abacus memory by TM tape
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   348
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   349
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   350
consts tape_of :: "'a \<Rightarrow> block list" ("<_>" 100)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   351
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   352
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   353
  @{text "tape_of_nat_list am"} returns the TM tape representing
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   354
  Abacus memory @{text "am"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   355
  *}
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
fun tape_of_nat_list :: "nat list \<Rightarrow> block list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   358
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   359
  "tape_of_nat_list [] = []" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   360
  "tape_of_nat_list [n] = Oc\<^bsup>n+1\<^esup>" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   361
  "tape_of_nat_list (n#ns) = (Oc\<^bsup>n+1\<^esup>) @ [Bk] @ (tape_of_nat_list ns)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   362
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   363
defs (overloaded)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   364
  tape_of_nl_abv: "<am> \<equiv> tape_of_nat_list am"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   365
  tape_of_nat_abv : "<(n::nat)> \<equiv> Oc\<^bsup>n+1\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   366
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   367
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   368
  @{text "crsp_l acf tcf"} meams the abacus configuration @{text "acf"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   369
  is corretly represented by the TM configuration @{text "tcf"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   370
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   371
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   372
fun crsp_l :: "layout \<Rightarrow> abc_conf_l \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   373
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   374
  "crsp_l ly (as, lm) (ts, (l, r)) inres = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   375
           (ts = start_of ly as \<and> (\<exists> rn. r = <lm> @ Bk\<^bsup>rn\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   376
            \<and> l = Bk # Bk # inres)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   377
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   378
declare crsp_l.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   379
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   380
subsection {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   381
  A more general definition of TM execution. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   382
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   383
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   384
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   385
fun nnth_of :: "(taction \<times> nat) list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (taction \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   386
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   387
  "nnth_of p s b = (if 2*s < length p 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   388
                    then (p ! (2*s + b)) else (Nop, 0))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   389
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   390
thm nth_of.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   391
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   392
fun nfetch :: "tprog \<Rightarrow> nat \<Rightarrow> block \<Rightarrow> taction \<times> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   393
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   394
  "nfetch p 0 b = (Nop, 0)" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   395
  "nfetch p (Suc s) b = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   396
             (case b of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   397
                Bk \<Rightarrow> nnth_of p s 0 |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   398
                Oc \<Rightarrow> nnth_of p s 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   399
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   400
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   401
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   402
  @{text "t_step tcf (tp, ss)"} returns the result of one step exection of TM @{text "tp"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   403
  assuming @{text "tp"} starts from instial state @{text "ss"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   404
*}
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
fun t_step :: "t_conf \<Rightarrow> (tprog \<times> nat) \<Rightarrow> t_conf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   407
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   408
  "t_step c (p, off) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   409
           (let (state, leftn, rightn) = c in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   410
            let (action, next_state) = fetch p (state-off)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   411
                             (case rightn of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   412
                                [] \<Rightarrow> Bk | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   413
                                Bk # xs \<Rightarrow> Bk |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   414
                                Oc # xs \<Rightarrow> Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   415
                             ) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   416
             in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   417
            (next_state, new_tape action (leftn, rightn)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   418
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   419
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   420
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   421
  @{text "t_steps tcf (tp, ss) n"} returns the result of @{text "n"}-step exection 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   422
  of TM @{text "tp"} assuming @{text "tp"} starts from instial state @{text "ss"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   423
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   424
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   425
fun t_steps :: "t_conf \<Rightarrow> (tprog \<times> nat) \<Rightarrow> nat \<Rightarrow> t_conf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   426
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   427
  "t_steps c (p, off) 0 = c" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   428
  "t_steps c (p, off) (Suc n) = t_steps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   429
                     (t_step c (p, off)) (p, off) n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   430
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   431
lemma stepn: "t_steps c (p, off) (Suc n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   432
              t_step (t_steps c (p, off) n) (p, off)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   433
apply(induct n arbitrary: c, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   434
apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   435
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   436
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   437
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   438
  The type of invarints expressing correspondence between 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   439
  Abacus configuration and TM configuration.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   440
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   441
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   442
type_synonym inc_inv_t = "abc_conf_l \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   443
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   444
declare tms_of.simps[simp del] tm_of.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   445
        layout_of.simps[simp del] abc_fetch.simps [simp del]  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   446
        t_step.simps[simp del] t_steps.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   447
        tpairs_of.simps[simp del] start_of.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   448
        fetch.simps [simp del] t_ncorrect.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   449
        new_tape.simps [simp del] ci.simps [simp del] length_of.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   450
        layout_of.simps[simp del] crsp_l.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   451
        abc2t_correct.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   452
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   453
lemma tct_div2: "t_ncorrect tp \<Longrightarrow> (length tp) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   454
apply(simp add: t_ncorrect.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   455
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   456
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   457
lemma t_shift_fetch: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   458
    "\<lbrakk>t_ncorrect tp1; t_ncorrect tp; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   459
      length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   460
    \<Longrightarrow> fetch tp (a - length tp1 div 2) b = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   461
         fetch (tp1 @ tp @ tp2) a b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   462
apply(subgoal_tac "\<exists> x. a = length tp1 div 2 + x", erule exE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   463
apply(case_tac x, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   464
apply(subgoal_tac "length tp1 div 2 + Suc nat = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   465
             Suc (length tp1 div 2 + nat)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   466
apply(simp only: fetch.simps nth_of.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   467
apply(case_tac b, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   468
apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   469
apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   470
apply(simp add: t_ncorrect.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   471
apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   472
apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   473
apply(simp add: t_ncorrect.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   474
apply(rule_tac x = "a - length tp1 div 2" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   475
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   476
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   477
lemma t_shift_in_step:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   478
      "\<lbrakk>t_step (a, aa, ba) (tp, length tp1 div 2) = (s, l, r);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   479
        t_ncorrect tp1; t_ncorrect tp;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   480
        length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   481
       \<Longrightarrow> t_step (a, aa, ba) (tp1 @ tp @ tp2, 0) = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   482
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   483
apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   484
                   Bk | x # xs \<Rightarrow> x)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   485
             = fetch (tp1 @ tp @ tp2) a (case ba of [] \<Rightarrow> Bk | x # xs
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   486
                   \<Rightarrow> x)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   487
apply(case_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow> Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   488
                | x # xs \<Rightarrow> x)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   489
apply(auto intro: t_shift_fetch)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   490
apply(case_tac ba, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   491
apply(case_tac aaa, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   492
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   493
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   494
declare add_Suc_right[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   495
lemma t_step_add: "t_steps c (p, off) (m + n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   496
          t_steps (t_steps c (p, off) m) (p, off) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   497
apply(induct m arbitrary: n,  simp add: t_steps.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   498
apply(subgoal_tac "t_steps c (p, off) (Suc (m + n)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   499
                         t_steps c (p, off) (m + Suc n)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   500
apply(subgoal_tac "t_steps (t_steps c (p, off) m) (p, off) (Suc n) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   501
                t_steps (t_step (t_steps c (p, off) m) (p, off)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   502
                         (p, off) n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   503
apply(simp, simp add: stepn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   504
apply(simp only: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   505
apply(simp only: add_Suc_right)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   506
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   507
declare add_Suc_right[simp]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   508
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   509
lemma s_out_fetch: "\<lbrakk>t_ncorrect tp; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   510
        \<not> (length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   511
         length tp div 2)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   512
      \<Longrightarrow> fetch tp (a - length tp1 div 2) b = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   513
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   514
apply(simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   515
apply(subgoal_tac "\<exists> x. a - length tp1 div 2 = length tp div 2 + x")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   516
apply(erule exE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   517
apply(case_tac x, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   518
apply(auto simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   519
apply(subgoal_tac "2 * (length tp div 2) =  length tp")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   520
apply(auto simp: t_ncorrect.simps split: block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   521
apply(rule_tac x = "a - length tp1 div 2 - length tp div 2" in exI
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   522
     , simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   523
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   524
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   525
lemma conf_keep_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   526
      "\<lbrakk>t_ncorrect tp; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   527
        \<not> (length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   528
       length tp div 2)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   529
      \<Longrightarrow> t_step (a, aa, ba) (tp, length tp1 div 2) = (0, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   530
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   531
apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   532
  Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) = (Nop, 0)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   533
apply(simp add: new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   534
apply(rule s_out_fetch, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   535
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   536
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   537
lemma conf_keep: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   538
      "\<lbrakk>t_ncorrect tp; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   539
        \<not> (length tp1 div 2 < a \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   540
        a \<le> length tp1 div 2 + length tp div 2); n > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   541
      \<Longrightarrow> t_steps (a, aa, ba) (tp, length tp1 div 2) n = (0, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   542
apply(induct n, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   543
apply(case_tac n, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   544
apply(rule_tac conf_keep_step, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   545
apply(subgoal_tac " t_steps (a, aa, ba) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   546
               (tp, length tp1 div 2) (Suc (Suc nat))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   547
         = t_step (t_steps (a, aa, ba) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   548
            (tp, length tp1 div 2) (Suc nat)) (tp, length tp1 div 2)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   549
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   550
apply(rule_tac conf_keep_step, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   551
apply(rule stepn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   552
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   553
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   554
lemma state_bef_inside: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   555
    "\<lbrakk>t_ncorrect tp1; t_ncorrect tp; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   556
      t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   557
      length tp1 div 2 < s0 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   558
         s0 \<le> length tp1 div 2 + length tp div 2;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   559
      length tp1 div 2 < s \<and> s \<le> length tp1 div 2 + length tp div 2; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   560
      n < stp; t_steps (s0, l0, r0) (tp, length tp1 div 2) n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   561
      (a, aa, ba)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   562
      \<Longrightarrow>  length tp1 div 2 < a \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   563
         a \<le> length tp1 div 2 + length tp div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   564
apply(subgoal_tac "\<exists> x. stp = n + x", erule exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   565
apply(simp only: t_step_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   566
apply(rule classical)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   567
apply(subgoal_tac "t_steps (a, aa, ba) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   568
          (tp, length tp1 div 2) x = (0, aa, ba)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   569
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   570
apply(rule conf_keep, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   571
apply(rule_tac x = "stp - n" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   572
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   573
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   574
lemma turing_shift_inside: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   575
       "\<lbrakk>t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   576
         length tp1 div 2 < s0 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   577
         s0 \<le> length tp1 div 2 + length tp div 2; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   578
         t_ncorrect tp1; t_ncorrect tp;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   579
         length tp1 div 2 < s \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   580
         s \<le> length tp1 div 2 + length tp div 2\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   581
       \<Longrightarrow> t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   582
apply(induct stp arbitrary: s l r)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   583
apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   584
apply(subgoal_tac " t_steps (s0, l0, r0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   585
        (tp, length tp1 div 2) (Suc stp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   586
                  = t_step (t_steps (s0, l0, r0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   587
        (tp, length tp1 div 2) stp) (tp, length tp1 div 2)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   588
apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) stp")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   589
apply(subgoal_tac "length tp1 div 2 < a \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   590
            a \<le> length tp1 div 2 + length tp div 2")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   591
apply(subgoal_tac "t_steps (s0, l0, r0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   592
           (tp1 @ tp @ tp2, 0) stp = (a, b, c)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   593
apply(simp only: stepn, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   594
apply(rule_tac t_shift_in_step, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   595
defer
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   596
apply(rule stepn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   597
apply(rule_tac n = stp and stp = "Suc stp" and a = a 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   598
               and aa = b and ba = c in state_bef_inside, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   599
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   600
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   601
lemma take_Suc_last[elim]: "Suc as \<le> length xs \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   602
            take (Suc as) xs = take as xs @ [xs ! as]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   603
apply(induct xs arbitrary: as, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   604
apply(case_tac as, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   605
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   606
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   607
lemma concat_suc: "Suc as \<le> length xs \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   608
       concat (take (Suc as) xs) = concat (take as xs) @ xs! as"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   609
apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   610
by auto
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
lemma concat_take_suc_iff: "Suc n \<le> length tps \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   613
       concat (take n tps) @ (tps ! n) = concat (take (Suc n) tps)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   614
apply(drule_tac concat_suc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   615
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   616
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   617
lemma concat_drop_suc_iff: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   618
   "Suc n < length tps \<Longrightarrow> concat (drop (Suc n) tps) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   619
           tps ! Suc n @ concat (drop (Suc (Suc n)) tps)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   620
apply(induct tps arbitrary: n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   621
apply(case_tac tps, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   622
apply(case_tac n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   623
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   624
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   625
declare append_assoc[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   626
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   627
lemma  tm_append: "\<lbrakk>n < length tps; tp = tps ! n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   628
           \<exists> tp1 tp2. concat tps = tp1 @ tp @ tp2 \<and> tp1 = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   629
              concat (take n tps) \<and> tp2 = concat (drop (Suc n) tps)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   630
apply(rule_tac x = "concat (take n tps)" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   631
apply(rule_tac x = "concat (drop (Suc n) tps)" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   632
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   633
apply(induct n, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   634
apply(case_tac tps, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   635
apply(subgoal_tac "concat (take n tps) @ (tps ! n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   636
               concat (take (Suc n) tps)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   637
apply(simp only: append_assoc[THEN sym], simp only: append_assoc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   638
apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   639
                  concat (drop (Suc (Suc n)) tps)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   640
apply(rule_tac concat_drop_suc_iff, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   641
apply(rule_tac concat_take_suc_iff, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   642
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   643
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   644
declare append_assoc[simp]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   645
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   646
lemma map_of:  "n < length xs \<Longrightarrow> (map f xs) ! n = f (xs ! n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   647
by(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   648
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   649
lemma [simp]: "length (tms_of aprog) = length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   650
apply(auto simp: tms_of.simps tpairs_of.simps)
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   653
lemma ci_nth: "\<lbrakk>ly = layout_of aprog; as < length aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   654
                abc_fetch as aprog = Some ins\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   655
    \<Longrightarrow> ci ly (start_of ly as) ins = tms_of aprog ! as"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   656
apply(simp add: tms_of.simps tpairs_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   657
      abc_fetch.simps  map_of del: map_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   658
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   659
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   660
lemma t_split:"\<lbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   661
        ly = layout_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   662
        as < length aprog; abc_fetch as aprog = Some ins\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   663
      \<Longrightarrow> \<exists> tp1 tp2. concat (tms_of aprog) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   664
            tp1 @ (ci ly (start_of ly as) ins) @ tp2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   665
            \<and> tp1 = concat (take as (tms_of aprog)) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   666
              tp2 = concat (drop (Suc as) (tms_of aprog))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   667
apply(insert tm_append[of "as" "tms_of aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   668
                             "ci ly (start_of ly as) ins"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   669
apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   670
apply(subgoal_tac "length (tms_of aprog) = length aprog", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   671
apply(rule_tac ci_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   672
done
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 math_sub: "\<lbrakk>x >= Suc 0; x - 1 = z\<rbrakk> \<Longrightarrow> x + y - Suc 0 = z + y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   675
by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   676
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   677
lemma start_more_one: "as \<noteq> 0 \<Longrightarrow> start_of ly as >= Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   678
apply(induct as, simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   679
apply(case_tac as, auto simp: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   680
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   681
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   682
lemma tm_ct: "\<lbrakk>abc2t_correct aprog; tp \<in> set (tms_of aprog)\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   683
                           t_ncorrect tp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   684
apply(simp add: abc2t_correct.simps tms_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   685
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   686
apply(simp add:list_all_iff, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   687
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   688
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   689
lemma div_apart: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   690
          \<Longrightarrow> (x + y) div 2 = x div 2 + y div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   691
apply(drule mod_eqD)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   692
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   693
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   694
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   695
lemma div_apart_iff: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   696
           (x + y) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   697
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   698
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   699
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   700
lemma tms_ct: "\<lbrakk>abc2t_correct aprog; n < length aprog\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   701
         t_ncorrect (concat (take n (tms_of aprog)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   702
apply(induct n, simp add: t_ncorrect.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   703
apply(subgoal_tac "concat (take (Suc n) (tms_of aprog)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   704
        concat (take n (tms_of aprog)) @ (tms_of aprog ! n)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   705
apply(simp add: t_ncorrect.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   706
apply(rule_tac div_apart_iff, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   707
apply(subgoal_tac "t_ncorrect (tms_of aprog ! n)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   708
            simp add: t_ncorrect.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   709
apply(rule_tac tm_ct, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   710
apply(rule_tac nth_mem, simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   711
apply(rule_tac concat_suc, simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   712
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   713
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   714
lemma tcorrect_div2: "\<lbrakk>abc2t_correct aprog; Suc as < length aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   715
  \<Longrightarrow> (length (concat (take as (tms_of aprog))) + length (tms_of aprog
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   716
 ! as)) div 2 = length (concat (take as (tms_of aprog))) div 2 + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   717
                 length (tms_of aprog ! as) div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   718
apply(subgoal_tac "t_ncorrect (tms_of aprog ! as)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   719
apply(subgoal_tac "t_ncorrect (concat (take as (tms_of aprog)))")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   720
apply(rule_tac div_apart)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   721
apply(rule tct_div2, simp)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   722
apply(erule_tac tms_ct, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   723
apply(rule_tac tm_ct, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   724
apply(rule_tac nth_mem)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   725
apply(simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   726
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   727
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   728
lemma [simp]: "length (layout_of aprog) = length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   729
apply(auto simp: layout_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   730
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   731
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   732
lemma start_of_ind: "\<lbrakk>as < length aprog; ly = layout_of aprog\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   733
       start_of ly (Suc as) = start_of ly as + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   734
                          length ((tms_of aprog) ! as) div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   735
apply(simp only: start_of.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   736
apply(auto simp: start_of.simps tms_of.simps layout_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   737
                 tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   738
apply(simp add: ci_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   739
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   740
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   741
lemma concat_take_suc: "Suc n \<le> length xs \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   742
  concat (take (Suc n) xs) = concat (take n xs) @ (xs ! n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   743
apply(subgoal_tac "take (Suc n) xs =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   744
                   take n xs @ [xs ! n]")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   745
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   746
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   747
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   748
lemma ci_length_not0: "Suc 0 <= length (ci ly as i) div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   749
apply(subgoal_tac "length (ci ly as i) div 2 = length_of i")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   750
apply(simp add: length_of.simps split: abc_inst.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   751
apply(rule ci_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   752
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   753
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   754
lemma findnth_length2: "length (findnth n) = 4 * n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   755
apply(induct n, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   756
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   757
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   758
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   759
lemma ci_length2: "length (ci ly as i) = 2 * (length_of i)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   760
apply(simp add: ci.simps length_of.simps tinc_b_def tdec_b_def
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   761
              split: abc_inst.splits, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   762
apply(simp add: findnth_length2)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   763
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   764
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   765
lemma tm_mod2: "as < length aprog \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   766
             length (tms_of aprog ! as) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   767
apply(simp add: tms_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   768
apply(subgoal_tac "map (\<lambda>(x, y). ci (layout_of aprog) x y) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   769
              (tpairs_of aprog) ! as
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   770
                = (\<lambda>(x, y). ci (layout_of aprog) x y) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   771
              ((tpairs_of aprog) ! as)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   772
apply(case_tac "(tpairs_of aprog ! as)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   773
apply(subgoal_tac "length (ci (layout_of aprog) a b) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   774
                 2 * (length_of b)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   775
apply(rule ci_length2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   776
apply(rule map_of, simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   777
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   778
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   779
lemma tms_mod2: "as \<le> length aprog \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   780
        length (concat (take as (tms_of aprog))) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   781
apply(induct as, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   782
apply(subgoal_tac "concat (take (Suc as) (tms_of aprog))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   783
                  = concat (take as (tms_of aprog)) @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   784
                       (tms_of aprog ! as)", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   785
apply(rule div_apart_iff, simp, rule tm_mod2, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   786
apply(rule concat_take_suc, simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   787
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   788
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   789
lemma [simp]: "\<lbrakk>as < length aprog; (abc_fetch as aprog) = Some ins\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   790
       \<Longrightarrow> ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   791
          (start_of (layout_of aprog) as) (ins) \<in> set (tms_of aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   792
apply(insert ci_nth[of "layout_of aprog" aprog as], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   793
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   794
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   795
lemma startof_not0: "start_of ly as > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   796
apply(induct as, simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   797
apply(case_tac as, auto simp: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   798
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   799
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   800
declare abc_step_l.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   801
lemma pre_lheq: "\<lbrakk>tp = concat (take as (tms_of aprog));
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   802
   abc2t_correct aprog; as \<le> length aprog\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   803
         start_of (layout_of aprog) as - Suc 0 = length tp div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   804
apply(induct as arbitrary: tp, simp add: start_of.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   805
proof - 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   806
  fix as tp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   807
  assume h1: "\<And>tp. tp = concat (take as (tms_of aprog)) \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   808
     start_of (layout_of aprog) as - Suc 0 = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   809
            length (concat (take as (tms_of aprog))) div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   810
  and h2: " abc2t_correct aprog" "Suc as \<le> length aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   811
  from h2 show "start_of (layout_of aprog) (Suc as) - Suc 0 = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   812
          length (concat (take (Suc as) (tms_of aprog))) div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   813
    apply(insert h1[of "concat (take as (tms_of aprog))"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   814
    apply(insert start_of_ind[of as aprog "layout_of aprog"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   815
    apply(subgoal_tac "(take (Suc as) (tms_of aprog)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   816
            take as (tms_of aprog) @ [(tms_of aprog) ! as]", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   817
    apply(subgoal_tac "(length (concat (take as (tms_of aprog))) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   818
                       length (tms_of aprog ! as)) div 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   819
            = length (concat (take as (tms_of aprog))) div 2 + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   820
              length (tms_of aprog ! as) div 2", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   821
    apply(subgoal_tac "start_of (layout_of aprog) as = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   822
       length (concat (take as (tms_of aprog))) div 2 + Suc 0", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   823
    apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   824
           rule_tac startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   825
    apply(insert tm_mod2[of as aprog], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   826
    apply(insert tms_mod2[of as aprog], simp, arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   827
    apply(rule take_Suc_last, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   828
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   829
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   830
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   831
lemma crsp2stateq: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   832
  "\<lbrakk>as < length aprog; abc2t_correct aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   833
       crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   834
        a = length (concat (take as (tms_of aprog))) div 2 + 1"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   835
apply(simp add: crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   836
apply(insert pre_lheq[of "(concat (take as (tms_of aprog)))" as aprog]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   837
, simp)   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   838
apply(subgoal_tac "start_of (layout_of aprog) as > 0", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   839
      auto intro: startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   840
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   841
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   842
lemma turing_shift_outside: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   843
     "\<lbrakk>t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   844
       s \<noteq> 0; stp > 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   845
       length tp1 div 2 < s0 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   846
       s0 \<le> length tp1 div 2 + length tp div 2; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   847
       t_ncorrect tp1; t_ncorrect tp;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   848
       \<not> (length tp1 div 2 < s \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   849
      s \<le> length tp1 div 2 + length tp div 2)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   850
    \<Longrightarrow> \<exists>stp' > 0. t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp' 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   851
                = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   852
apply(rule_tac x = stp in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   853
apply(case_tac stp, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   854
apply(simp only: stepn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   855
apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) nat")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   856
apply(subgoal_tac "length tp1 div 2 < a \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   857
                   a \<le> length tp1 div 2 + length tp div 2")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   858
apply(subgoal_tac "t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) nat 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   859
                   = (a, b, c)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   860
apply(rule_tac t_shift_in_step, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   861
apply(rule_tac turing_shift_inside, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   862
apply(rule classical)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   863
apply(subgoal_tac "t_step (a,b,c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   864
            (tp, length tp1 div 2) = (0, b, c)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   865
apply(rule_tac conf_keep_step, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   866
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   867
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   868
lemma turing_shift: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   869
  "\<lbrakk>t_steps (s0, (l0, r0)) (tp, (length tp1 div 2)) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   870
   = (s, (l, r)); s \<noteq> 0; stp > 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   871
  (length tp1 div 2 < s0 \<and> s0 <= length tp1 div 2 + length tp div 2);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   872
  t_ncorrect tp1; t_ncorrect tp\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   873
         \<exists> stp' > 0. t_steps (s0, (l0, r0)) (tp1 @ tp @ tp2, 0) stp' =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   874
                    (s, (l, r))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   875
apply(case_tac "s > length tp1 div 2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   876
              s <= length tp1 div 2 + length tp div 2")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   877
apply(subgoal_tac " t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   878
                   (s, l, r)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   879
apply(rule_tac x = stp in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   880
apply(rule_tac turing_shift_inside, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   881
apply(rule_tac turing_shift_outside, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   882
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   883
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   884
lemma inc_startof_not0:  "start_of ly as \<ge> Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   885
apply(induct as, simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   886
apply(simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   887
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   888
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   889
lemma s_crsp:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   890
  "\<lbrakk>as < length aprog; abc_fetch as aprog = Some ins;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   891
  abc2t_correct aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   892
  crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\<rbrakk> \<Longrightarrow>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   893
  length (concat (take as (tms_of aprog))) div 2 < a 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   894
      \<and> a \<le> length (concat (take as (tms_of aprog))) div 2 + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   895
         length (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   896
         ins) div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   897
apply(subgoal_tac "a = length (concat (take as (tms_of aprog))) div 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   898
                   2 + 1", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   899
apply(rule_tac ci_length_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   900
apply(rule crsp2stateq, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   901
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   902
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   903
lemma tms_out_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   904
  "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   905
  abc2t_correct aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   906
  crsp_l ly (as, am) tc inres; as < length aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   907
  abc_fetch as aprog = Some ins;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   908
  t_steps tc (ci ly (start_of ly as) ins, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   909
  (start_of ly as) - 1) n = (s, l, r);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   910
  n > 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   911
  abc_step_l (as, am) (abc_fetch as aprog) = (as', am');
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   912
  s = start_of ly as'
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   913
  \<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   914
  \<Longrightarrow> \<exists> stp > 0. (t_steps tc (tprog, 0) stp = (s, (l, r)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   915
apply(simp only: tm_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   916
apply(subgoal_tac "\<exists> tp1 tp2. concat (tms_of aprog) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   917
      tp1 @ (ci ly (start_of ly as) ins) @ tp2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   918
    \<and> tp1 = concat (take as (tms_of aprog)) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   919
      tp2 = concat (drop (Suc as) (tms_of aprog))")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   920
apply(erule exE, erule exE, erule conjE, erule conjE,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   921
      case_tac tc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   922
apply(rule turing_shift)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   923
apply(subgoal_tac "start_of (layout_of aprog) as - Suc 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   924
                = length tp1 div 2", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   925
apply(rule_tac pre_lheq, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   926
apply(simp add: startof_not0, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   927
apply(rule_tac s_crsp, simp, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   928
apply(rule tms_ct, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   929
apply(rule tm_ct, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   930
apply(subgoal_tac "ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   931
                 (start_of (layout_of aprog) as) ins
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   932
                = (tms_of aprog ! as)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   933
apply(simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   934
apply(simp add: tms_of.simps tpairs_of.simps abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   935
apply(erule_tac t_split, auto simp: tm_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   936
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   937
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   938
subsubsection {* The compilation of @{text "Inc n"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   939
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   940
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   941
  The lemmas in this section lead to the correctness of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   942
  the compilation of @{text "Inc n"} instruction.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   943
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   944
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   945
(*****Begin: inc crsp*******)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   946
fun at_begin_fst_bwtn :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   947
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   948
  "at_begin_fst_bwtn (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   949
      (\<exists> lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \<and> length lm1 = s \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   950
          (if lm1 = [] then l = Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   951
           else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = (Bk\<^bsup>rn\<^esup>))" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   952
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   953
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   954
fun at_begin_fst_awtn :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   955
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   956
  "at_begin_fst_awtn (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   957
      (\<exists> lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \<and> length lm1 = s \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   958
         (if lm1 = []  then l = Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   959
          else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = [Oc]@Bk\<^bsup>rn\<^esup>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   960
  )"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   961
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   962
fun at_begin_norm :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   963
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   964
  "at_begin_norm (as, lm) (s, l, r) ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   965
      (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> length lm1 = s \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   966
        (if lm1 = [] then l = Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   967
         else l = Bk # <rev lm1> @ Bk# Bk # ires ) \<and> r = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   968
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   969
fun in_middle :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   970
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   971
  "in_middle (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   972
      (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   973
       \<and> length lm1 = s \<and> m + 1 = ml + mr \<and>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   974
         ml \<noteq> 0 \<and> tn = s + 1 - length lm \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   975
       (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   976
        else l = (Oc\<^bsup>ml\<^esup>)@[Bk]@<rev lm1>@
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   977
                 Bk # Bk # ires) \<and> (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2>@ (Bk\<^bsup>rn\<^esup>) \<or> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   978
      (lm2 = [] \<and> r = (Oc\<^bsup>mr\<^esup>)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   979
      )"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   980
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   981
fun inv_locate_a :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   982
  where "inv_locate_a (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   983
     (at_begin_norm (as, lm) (s, l, r) ires \<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   984
      at_begin_fst_bwtn (as, lm) (s, l, r) ires \<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   985
      at_begin_fst_awtn (as, lm) (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   986
      )"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   987
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   988
fun inv_locate_b :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   989
  where "inv_locate_b (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   990
        (in_middle (as, lm) (s, l, r)) ires "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   991
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   992
fun inv_after_write :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   993
  where "inv_after_write (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   994
           (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   995
             (if lm1 = [] then l = Oc\<^bsup>m\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   996
              else Oc # l = Oc\<^bsup>Suc m \<^esup>@ Bk # <rev lm1> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   997
                      Bk # Bk # ires) \<and> r = [Oc] @ <lm2> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   998
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   999
fun inv_after_move :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1000
  where "inv_after_move (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1001
      (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1002
        (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1003
         else l = Oc\<^bsup>Suc m\<^esup>@ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1004
        r = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1005
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1006
fun inv_after_clear :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1007
  where "inv_after_clear (as, lm) (s, l, r) ires =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1008
       (\<exists> rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1009
        (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1010
         else l = Oc\<^bsup>Suc m\<^esup>@ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1011
          r = Bk # r' \<and> Oc # r' = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1012
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1013
fun inv_on_right_moving :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1014
  where "inv_on_right_moving (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1015
       (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1016
            ml + mr = m \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1017
          (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1018
          else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1019
         ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2> @ (Bk\<^bsup>rn\<^esup>)) \<or> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1020
          (r = (Oc\<^bsup>mr\<^esup>) \<and> lm2 = [])))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1021
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1022
fun inv_on_left_moving_norm :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1023
  where "inv_on_left_moving_norm (as, lm) (s, l, r) ires =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1024
      (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1025
             ml + mr = Suc m \<and> mr > 0 \<and> (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1026
                                         else l = (Oc\<^bsup>ml\<^esup>) @ Bk # <rev lm1> @ Bk # Bk # ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1027
        \<and> (r = (Oc\<^bsup>mr\<^esup>) @ Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>) \<or> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1028
           (lm2 = [] \<and> r = Oc\<^bsup>mr\<^esup>)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1029
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1030
fun inv_on_left_moving_in_middle_B:: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1031
  where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1032
                (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1033
                     (if lm1 = [] then l = Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1034
                      else l = <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1035
                      r = Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1036
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1037
fun inv_on_left_moving :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1038
  where "inv_on_left_moving (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1039
       (inv_on_left_moving_norm  (as, lm) (s, l, r) ires \<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1040
        inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1041
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1042
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1043
fun inv_check_left_moving_on_leftmost :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1044
  where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1045
                (\<exists> rn. l = ires \<and> r = [Bk, Bk] @ <lm> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1046
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1047
fun inv_check_left_moving_in_middle :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1048
  where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1049
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1050
              (\<exists> lm1 lm2 r' rn. lm = lm1 @ lm2 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1051
                 (Oc # l = <rev lm1> @ Bk # Bk # ires) \<and> r = Oc # Bk # r' \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1052
                           r' = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1053
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1054
fun inv_check_left_moving :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1055
  where "inv_check_left_moving (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1056
             (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1057
             inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1058
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1059
fun inv_after_left_moving :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1060
  where "inv_after_left_moving (as, lm) (s, l, r) ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1061
              (\<exists> rn. l = Bk # ires \<and> r = Bk # <lm> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1062
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1063
fun inv_stop :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1064
  where "inv_stop (as, lm) (s, l, r) ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1065
              (\<exists> rn. l = Bk # Bk # ires \<and> r = <lm> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1066
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1067
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1068
fun inc_inv :: "layout \<Rightarrow> nat \<Rightarrow> inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1069
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1070
  "inc_inv ly n (as, lm) (s, l, r) ires =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1071
              (let ss = start_of ly as in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1072
               let lm' = abc_lm_s lm n ((abc_lm_v lm n)+1) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1073
                if s = 0 then False
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1074
                else if s < ss then False
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1075
                else if s < ss + 2 * n then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1076
                   if (s - ss) mod 2 = 0 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1077
                       inv_locate_a (as, lm) ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1078
                   else inv_locate_b (as, lm) ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1079
                else if s = ss + 2 * n then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1080
                        inv_locate_a (as, lm) (n, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1081
                else if s = ss + 2 * n + 1 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1082
                   inv_locate_b (as, lm) (n, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1083
                else if s = ss + 2 * n + 2 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1084
                   inv_after_write (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1085
                else if s = ss + 2 * n + 3 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1086
                   inv_after_move (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1087
                else if s = ss + 2 * n + 4 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1088
                   inv_after_clear (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1089
                else if s = ss + 2 * n + 5 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1090
                   inv_on_right_moving (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1091
                else if s = ss + 2 * n + 6 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1092
                   inv_on_left_moving (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1093
                else if s = ss + 2 * n + 7 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1094
                   inv_check_left_moving (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1095
                else if s = ss + 2 * n + 8 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1096
                   inv_after_left_moving (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1097
                else if s = ss + 2 * n + 9 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1098
                   inv_stop (as, lm') (s - ss, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1099
                else False) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1100
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1101
lemma fetch_intro: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1102
  "\<lbrakk>\<And>xs.\<lbrakk>ba = Oc # xs\<rbrakk> \<Longrightarrow> P (fetch prog i Oc);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1103
   \<And>xs.\<lbrakk>ba = Bk # xs\<rbrakk> \<Longrightarrow> P (fetch prog i Bk);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1104
   ba = [] \<Longrightarrow> P (fetch prog i Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1105
   \<rbrakk> \<Longrightarrow> P (fetch prog i 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1106
             (case ba of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1107
by (auto split:list.splits block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1108
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1109
lemma length_findnth[simp]: "length (findnth n) = 4 * n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1110
apply(induct n, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1111
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1112
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1113
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1114
declare tshift.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1115
declare findnth.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1116
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1117
lemma findnth_nth: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1118
 "\<lbrakk>n > q; x < 4\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1119
        (findnth n) ! (4 * q + x) = (findnth (Suc q) ! (4 * q + x))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1120
apply(induct n, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1121
apply(case_tac "q < n", simp add: findnth.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1122
apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1123
apply(subgoal_tac "q = n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1124
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1125
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1126
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1127
lemma Suc_pre[simp]: "\<not> a < start_of ly as \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1128
          (Suc a - start_of ly as) = Suc (a - start_of ly as)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1129
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1130
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1131
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1132
lemma fetch_locate_a_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1133
\<And>a  q xs.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1134
    \<lbrakk>\<not> a < start_of (layout_of aprog) as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1135
      a < start_of (layout_of aprog) as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1136
      a - start_of (layout_of aprog) as = 2 * q; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1137
      start_of (layout_of aprog) as > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1138
    \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1139
         (Inc n)) (Suc (2 * q)) Oc) = (R, a+1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1140
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1141
                  nth_of.simps tshift.simps nth_append Suc_pre)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1142
apply(subgoal_tac "(findnth n ! Suc (4 * q)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1143
                 findnth (Suc q) ! (4 * q + 1)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1144
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1145
apply(subgoal_tac " findnth n !(4 * q + 1) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1146
                 findnth (Suc q) ! (4 * q + 1)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1147
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1148
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1149
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1150
lemma fetch_locate_a_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1151
\<And>a  q xs.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1152
    \<lbrakk>abc_fetch as aprog = Some (Inc n);  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1153
     \<not> a < start_of (layout_of aprog) as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1154
     a < start_of (layout_of aprog) as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1155
     a - start_of (layout_of aprog) as = 2 * q; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1156
     start_of (layout_of aprog) as > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1157
    \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1158
      (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * q)) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1159
       = (W1, a)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1160
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1161
                 tshift.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1162
apply(subgoal_tac "(findnth n ! (4 * q)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1163
                           findnth (Suc q) ! (4 * q )")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1164
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1165
apply(subgoal_tac " findnth n !(4 * q + 0) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1166
                            findnth (Suc q) ! (4 * q + 0)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1167
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1168
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1169
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1170
lemma [intro]: "x mod 2 = Suc 0 \<Longrightarrow> \<exists> q. x = Suc (2 * q)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1171
apply(drule mod_eqD, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1172
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1173
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1174
lemma  add3_Suc: "x + 3 = Suc (Suc (Suc x))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1175
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1176
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1177
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1178
declare start_of.simps[simp]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1179
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1180
lemma layout_not0: "start_of ly as > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1181
by(induct as, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1182
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1183
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1184
 "\<lbrakk>\<not> a < start_of (layout_of aprog) as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1185
   a - start_of (layout_of aprog) as = Suc (2 * q); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1186
   abc_fetch as aprog = Some (Inc n); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1187
   start_of (layout_of aprog) as > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1188
    \<Longrightarrow> Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) = a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1189
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1190
"Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1191
              = 2 + 2 * q + start_of (layout_of aprog) as - Suc 0", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1192
  simp, simp add: inc_startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1193
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1194
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1195
lemma fetch_locate_b_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1196
\<And>a  xs.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1197
    \<lbrakk>0 < a; \<not> a < start_of (layout_of aprog) as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1198
  a < start_of (layout_of aprog) as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1199
 (a - start_of (layout_of aprog) as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1200
 start_of (layout_of aprog) as > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1201
    \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1202
      (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Oc) = (R, a)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1203
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1204
                 nth_of.simps tshift.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1205
apply(subgoal_tac "\<exists> q. (a - start_of (layout_of aprog) as) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1206
                         2 * q + 1", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1207
apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1208
                  = findnth (Suc q) ! (4 * q + 3)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1209
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1210
apply(subgoal_tac " findnth n ! (4 * q + 3) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1211
                 findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1212
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1213
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1214
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1215
lemma fetch_locate_b_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1216
\<And>a  xs.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1217
    \<lbrakk>0 < a;  \<not> a < start_of (layout_of aprog) as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1218
     a < start_of (layout_of aprog) as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1219
     (a - start_of (layout_of aprog) as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1220
     start_of (layout_of aprog) as > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1221
    \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1222
        (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Bk) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1223
        = (R, a + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1224
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1225
                  nth_of.simps tshift.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1226
apply(subgoal_tac "\<exists> q. (a - start_of (layout_of aprog) as) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1227
                  2 * q + 1", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1228
apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1229
                    findnth (Suc q) ! (4 * q + 2)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1230
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1231
apply(subgoal_tac " findnth n ! (4 * q + 2) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1232
                    findnth (Suc q) ! (4 * q + 2)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1233
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1234
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1235
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1236
lemma fetch_locate_n_a_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1237
       "start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1238
       \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1239
      (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Oc) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1240
             (R, start_of (layout_of aprog) as + 2 * n + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1241
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1242
                  nth_of.simps tshift.simps nth_append tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1243
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1244
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1245
lemma fetch_locate_n_a_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1246
       start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1247
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1248
    (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Bk) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1249
   = (W1, start_of (layout_of aprog) as + 2 * n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1250
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1251
                  nth_of.simps tshift.simps nth_append tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1252
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1253
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1254
lemma fetch_locate_n_b_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1255
    start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1256
    \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1257
     (Inc n)) (Suc (Suc (2 * n))) Oc) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1258
                      (R, start_of (layout_of aprog) as + 2 * n + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1259
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1260
                  nth_of.simps tshift.simps nth_append tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1261
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1262
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1263
lemma fetch_locate_n_b_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1264
    start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1265
   \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1266
   (Inc n)) (Suc (Suc (2 * n))) Bk) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1267
       (W1, start_of (layout_of aprog) as + 2 * n + 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1268
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1269
                  nth_of.simps tshift.simps nth_append tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1270
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1271
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1272
lemma fetch_after_write_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1273
    start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1274
    \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1275
            (Inc n)) (Suc (Suc (Suc (2 * n)))) Oc) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1276
        (R, start_of (layout_of aprog) as + 2*n + 3)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1277
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1278
                  nth_of.simps tshift.simps nth_append tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1279
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1280
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1281
lemma fetch_after_move_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1282
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1283
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1284
              (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Oc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1285
        = (W0, start_of (layout_of aprog) as + 2 * n + 4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1286
apply(auto simp: ci.simps findnth.simps tshift.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1287
                 tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1288
apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1289
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1290
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1291
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1292
lemma fetch_after_move_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1293
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1294
      \<Longrightarrow>(fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1295
            (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1296
       = (L, start_of (layout_of aprog) as + 2 * n + 6)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1297
apply(auto simp: ci.simps findnth.simps tshift.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1298
                 tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1299
apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1300
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1301
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1302
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1303
lemma fetch_clear_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1304
     start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1305
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1306
              (start_of (layout_of aprog) as) (Inc n)) (5 + 2 * n) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1307
      = (R, start_of (layout_of aprog) as + 2 * n + 5)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1308
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1309
                     tshift.simps tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1310
apply(subgoal_tac "5 + 2*n = Suc (2*n + 4)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1311
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1312
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1313
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1314
lemma fetch_right_move_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1315
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1316
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1317
                (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Oc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1318
      = (R, start_of (layout_of aprog) as + 2 * n + 5)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1319
apply(auto simp: ci.simps findnth.simps tshift.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1320
                 tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1321
apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1322
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1323
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1324
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1325
lemma fetch_right_move_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1326
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1327
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1328
                (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1329
      = (W1, start_of (layout_of aprog) as + 2 * n + 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1330
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1331
                 tshift.simps tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1332
apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1333
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1334
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1335
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1336
lemma fetch_left_move_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1337
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1338
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1339
               (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Oc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1340
      = (L, start_of (layout_of aprog) as + 2 * n + 6)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1341
apply(auto simp: ci.simps findnth.simps tshift.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1342
                 tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1343
apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1344
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1345
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1346
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1347
lemma fetch_left_move_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1348
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1349
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1350
               (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1351
      = (L, start_of (layout_of aprog) as + 2 * n + 7)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1352
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1353
                 tshift.simps tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1354
apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1355
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1356
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1357
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1358
lemma fetch_check_left_move_o: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1359
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1360
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1361
               (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Oc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1362
      = (L, start_of (layout_of aprog) as + 2 * n + 6)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1363
apply(auto simp: ci.simps findnth.simps tshift.simps tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1364
apply(subgoal_tac "8 + 2 * n = Suc (2 * n + 7)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1365
                                  simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1366
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1367
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1368
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1369
lemma fetch_check_left_move_b: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1370
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1371
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1372
              (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1373
      = (R, start_of (layout_of aprog) as + 2 * n + 8)  "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1374
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1375
                 tshift.simps tinc_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1376
apply(subgoal_tac "8 + 2*n= Suc (2*n + 7)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1377
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1378
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1379
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1380
lemma fetch_after_left_move: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1381
      start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1382
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1383
              (start_of (layout_of aprog) as) (Inc n)) (9 + 2*n) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1384
     = (R, start_of (layout_of aprog) as + 2 * n + 9)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1385
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1386
                  nth_of.simps tshift.simps nth_append tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1387
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1388
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1389
lemma fetch_stop: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1390
       start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1391
      \<Longrightarrow> (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1392
             (start_of (layout_of aprog) as) (Inc n)) (10 + 2 *n)  b)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1393
     = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1394
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1395
                  nth_of.simps tshift.simps nth_append tinc_b_def
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1396
            split: block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1397
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1398
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1399
lemma fetch_state0: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1400
       (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1401
               (start_of (layout_of aprog) as) (Inc n)) 0 b)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1402
     = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1403
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1404
                  nth_of.simps tshift.simps nth_append tinc_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1405
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1406
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1407
lemmas fetch_simps = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1408
  fetch_locate_a_o fetch_locate_a_b fetch_locate_b_o fetch_locate_b_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1409
  fetch_locate_n_a_b fetch_locate_n_a_o fetch_locate_n_b_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1410
  fetch_locate_n_b_b fetch_after_write_o fetch_after_move_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1411
  fetch_after_move_b fetch_clear_b fetch_right_move_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1412
  fetch_right_move_b fetch_left_move_o fetch_left_move_b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1413
  fetch_after_left_move fetch_check_left_move_o fetch_stop 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1414
  fetch_state0 fetch_check_left_move_b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1415
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1416
text {* *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1417
declare exponent_def[simp del] tape_of_nat_list.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1418
   at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1419
   at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1420
   abc_lm_s.simps[simp del] abc_lm_v.simps[simp del]  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1421
   ci.simps[simp del] t_step.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1422
   inv_after_move.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1423
   inv_on_left_moving_norm.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1424
   inv_on_left_moving_in_middle_B.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1425
   inv_after_clear.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1426
   inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1427
   inv_on_right_moving.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1428
   inv_check_left_moving.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1429
   inv_check_left_moving_in_middle.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1430
   inv_check_left_moving_on_leftmost.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1431
   inv_after_left_moving.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1432
   inv_stop.simps[simp del] inv_locate_a.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1433
   inv_locate_b.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1434
declare tms_of.simps[simp del] tm_of.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1435
        layout_of.simps[simp del] abc_fetch.simps [simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1436
        t_step.simps[simp del] t_steps.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1437
        tpairs_of.simps[simp del] start_of.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1438
        fetch.simps [simp del] new_tape.simps [simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1439
        nth_of.simps [simp del] ci.simps [simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1440
        length_of.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1441
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1442
(*! Start point *)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1443
lemma [simp]: "Suc (2 * q) mod 2 = Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1444
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1445
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1446
lemma [simp]: "Suc (2 * q) div 2 = q"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1447
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1448
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1449
lemma [simp]: "\<lbrakk> \<not> a < start_of ly as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1450
          a < start_of ly as + 2 * n; a - start_of ly as = 2 * q\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1451
             \<Longrightarrow> Suc a < start_of ly as + 2 * n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1452
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1453
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1454
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1455
lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> (Suc x) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1456
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1457
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1458
lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> (Suc x) div 2 = Suc (x div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1459
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1460
lemma exp_def[simp]: "a\<^bsup>Suc n \<^esup>= a # a\<^bsup>n\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1461
by(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1462
lemma [intro]: "Bk # r = Oc\<^bsup>mr\<^esup> @ r' \<Longrightarrow> mr = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1463
by(case_tac mr, auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1464
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1465
lemma [intro]: "Bk # r = replicate mr Oc \<Longrightarrow> mr = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1466
by(case_tac mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1467
lemma tape_of_nl_abv_cons[simp]: "xs \<noteq> [] \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1468
                   <x # xs> = Oc\<^bsup>Suc x\<^esup>@ Bk # <xs>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1469
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1470
apply(case_tac xs, simp, simp add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1471
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1472
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1473
lemma [simp]: "<[]::nat list> = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1474
by(auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1475
lemma [simp]: "Oc # r = <(lm::nat list)> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> lm \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1476
apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1477
apply(case_tac rn, auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1478
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1479
lemma BkCons_nil: "Bk # xs = <lm::nat list> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> lm = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1480
apply(case_tac lm, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1481
apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1482
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1483
lemma BkCons_nil': "Bk # xs = <lm::nat list> @ Bk\<^bsup>ln\<^esup>\<Longrightarrow> lm = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1484
by(auto intro: BkCons_nil)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1485
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1486
lemma hd_tl_tape_of_nat_list:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1487
   "tl (lm::nat list) \<noteq> [] \<Longrightarrow> <lm> = <hd lm> @ Bk # <tl lm>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1488
apply(frule tape_of_nl_abv_cons[of "tl lm" "hd lm"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1489
apply(simp add: tape_of_nat_abv Bk_def del: tape_of_nl_abv_cons)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1490
apply(subgoal_tac "lm = hd lm # tl lm", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1491
apply(case_tac lm, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1492
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1493
lemma [simp]: "Oc # xs = Oc\<^bsup>mr\<^esup> @ Bk # <lm2> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> mr > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1494
apply(case_tac mr, auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1495
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1496
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1497
lemma tape_of_nat_list_cons: "xs \<noteq> [] \<Longrightarrow> tape_of_nat_list (x # xs) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1498
              replicate (Suc x) Oc @ Bk # tape_of_nat_list xs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1499
apply(drule tape_of_nl_abv_cons[of xs x])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1500
apply(auto simp: tape_of_nl_abv tape_of_nat_abv Oc_def Bk_def exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1501
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1502
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1503
lemma rev_eq: "rev xs = rev ys \<Longrightarrow> xs = ys"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1504
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1505
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1506
lemma tape_of_nat_list_eq: " xs = ys \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1507
        tape_of_nat_list xs = tape_of_nat_list ys"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1508
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1509
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1510
lemma tape_of_nl_nil_eq: "<(lm::nat list)> = [] = (lm = [])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1511
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1512
apply(case_tac lm, simp add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1513
apply(case_tac "list")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1514
apply(auto simp: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1515
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1516
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1517
lemma rep_ind: "replicate (Suc n) a = replicate n a @ [a]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1518
apply(induct n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1519
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1520
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1521
lemma [simp]: "Oc # r = <lm::nat list> @ replicate rn Bk \<Longrightarrow> Suc 0 \<le> length lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1522
apply(rule_tac classical, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1523
apply(case_tac lm, simp, case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1524
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1525
lemma Oc_Bk_Cons: "Oc # Bk # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1526
                   lm \<noteq> [] \<and> hd lm = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1527
apply(case_tac lm, simp, case_tac ln, simp add: exponent_def, simp add: exponent_def, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1528
apply(case_tac lista, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1529
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1530
(*lemma Oc_Oc_Cons: "Oc # Oc # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1531
                  lm \<noteq> [] \<and> hd lm > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1532
apply(case_tac lm, simp add: exponent_def, case_tac ln, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1533
apply(case_tac lista, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1534
        auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1535
apply(case_tac [!] a, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1536
apply(case_tac ln, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1537
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1538
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1539
lemma Oc_nil_zero[simp]: "[Oc] = <lm::nat list> @ Bk\<^bsup>ln\<^esup> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1540
                 \<Longrightarrow> lm = [0] \<and> ln = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1541
apply(case_tac lm, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1542
apply(case_tac ln, auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1543
apply(case_tac [!] list, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1544
        auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1545
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1546
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1547
lemma  [simp]: "Oc # r = <lm2> @ replicate rn Bk \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1548
       (\<exists>rn. r = replicate (hd lm2) Oc @ Bk # <tl lm2> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1549
                      replicate rn Bk) \<or> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1550
          tl lm2 = [] \<and> r = replicate (hd lm2) Oc"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1551
apply(rule_tac disjCI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1552
apply(case_tac "tl lm2 = []", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1553
apply(case_tac lm2, simp add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1554
apply(case_tac rn, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1555
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1556
apply(case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1557
apply(rule_tac x = rn in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1558
apply(simp add: hd_tl_tape_of_nat_list)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1559
apply(simp add: tape_of_nat_abv Oc_def exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1560
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1561
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1562
(*inv: from locate_a to locate_b*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1563
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1564
      "inv_locate_a (as, lm) (q, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1565
       \<Longrightarrow> inv_locate_b (as, lm) (q, Oc # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1566
apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1567
          at_begin_norm.simps at_begin_fst_bwtn.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1568
          at_begin_fst_awtn.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1569
apply(erule disjE, erule exE, erule exE, erule exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1570
apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1571
apply(rule_tac x = "0" in exI, rule_tac x = "hd lm2" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1572
                auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1573
apply(rule_tac x = "Suc 0" in exI, simp add:exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1574
apply(rule_tac x = "lm @ replicate tn 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1575
      rule_tac x = "[]" in exI,    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1576
      rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1577
apply(simp only: rep_ind, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1578
apply(rule_tac x = "Suc 0" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1579
apply(case_tac [1-3] rn, simp_all )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1580
apply(rule_tac x = "lm @ replicate tn 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1581
      rule_tac x = "[]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1582
      rule_tac x = "Suc tn" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1583
      rule_tac x = 0 in exI, simp add: rep_ind del: replicate_Suc split:if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1584
apply(rule_tac x = "Suc 0" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1585
apply(case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1586
apply(rule_tac [!] x = "Suc 0" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1587
apply(case_tac [!] rn, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1588
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1589
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1590
(*inv: from locate_a to _locate_a*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1591
lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1592
       \<Longrightarrow> inv_locate_a (as, am) (q, aaa, Oc # xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1593
apply(simp only: inv_locate_a.simps at_begin_norm.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1594
                 at_begin_fst_bwtn.simps at_begin_fst_awtn.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1595
apply(erule_tac disjE, erule exE, erule exE, erule exE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1596
      rule disjI2, rule disjI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1597
defer
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1598
apply(erule_tac disjE, erule exE, erule exE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1599
      erule exE, rule disjI2, rule disjI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1600
prefer 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1601
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1602
proof-
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1603
  fix lm1 tn rn
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1604
  assume k: "lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> (if lm1 = [] then aaa = Bk # Bk # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1605
    ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Bk # xs = Bk\<^bsup>rn\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1606
  thus "\<exists>lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> (if lm1 = [] then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1607
    aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1608
    (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1609
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1610
    from k have "?P lm1 tn (rn - 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1611
      apply(auto simp: Oc_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1612
      by(case_tac [!] "rn::nat", auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1613
    thus ?thesis by blast
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1614
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1615
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1616
  fix lm1 lm2 rn
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1617
  assume h1: "am = lm1 @ lm2 \<and> length lm1 = q \<and> (if lm1 = [] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1618
    then aaa = Bk # Bk # ires else aaa = Bk # <rev lm1> @ Bk # Bk # ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1619
    Bk # xs = <lm2> @ Bk\<^bsup>rn\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1620
  from h1 have h2: "lm2 = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1621
  proof(rule_tac xs = xs and rn = rn in BkCons_nil, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1622
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1623
  from h1 and h2 show "\<exists>lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1624
    (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1625
    Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1626
    (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1627
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1628
    from h1 and h2  have "?P lm1 0 (rn - 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1629
      apply(auto simp: Oc_def exponent_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1630
                      tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1631
      by(case_tac "rn::nat", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1632
    thus ?thesis by blast
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1633
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1634
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1635
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1636
lemma [intro]: "\<exists>rn. [a] = a\<^bsup>rn\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1637
by(rule_tac x = "Suc 0" in exI, simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1638
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1639
lemma [intro]: "\<exists>tn. [] = a\<^bsup>tn\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1640
apply(rule_tac x = 0 in exI, simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1641
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1642
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1643
lemma [intro]:  "at_begin_norm (as, am) (q, aaa, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1644
             \<Longrightarrow> at_begin_norm (as, am) (q, aaa, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1645
apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1646
apply(rule_tac x = lm1 in exI, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1647
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1648
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1649
lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1650
            \<Longrightarrow> at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1651
apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1652
apply(rule_tac x = "am @ 0\<^bsup>tn\<^esup>" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1653
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1654
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1655
lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1656
           \<Longrightarrow> at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1657
apply(auto simp: at_begin_fst_awtn.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1658
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1659
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1660
lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1661
            \<Longrightarrow> inv_locate_a (as, am) (q, aaa, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1662
apply(simp only: inv_locate_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1663
apply(erule disj_forward)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1664
defer
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1665
apply(erule disj_forward, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1666
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1667
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1668
lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1669
               inv_locate_a (as, am) (q, aaa, [Oc]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1670
apply(insert locate_a_2_locate_a [of as am q aaa "[]"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1671
apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1672
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1673
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1674
(*inv: from locate_b to locate_b*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1675
lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1676
         \<Longrightarrow> inv_locate_b (as, am) (q, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1677
apply(simp only: inv_locate_b.simps in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1678
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1679
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1680
      rule_tac x = tn in exI, rule_tac x = m in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1681
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1682
      rule_tac x = rn in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1683
apply(case_tac mr, simp_all add: exponent_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1684
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1685
lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # <lm::nat list> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1686
                             Bk\<^bsup>rn \<^esup>) \<or> (lm2 = [] \<and> Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1687
       \<Longrightarrow> mr = 0 \<and> lm = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1688
apply(rule context_conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1689
apply(case_tac mr, auto simp:exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1690
apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1691
apply(case_tac n, auto simp: exponent_def Bk_def  tape_of_nl_nil_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1692
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1693
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1694
lemma tape_of_nat_def: "<[m::nat]> =  Oc # Oc\<^bsup>m\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1695
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1696
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1697
lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \<exists>n. xs = Bk\<^bsup>n\<^esup>\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1698
            \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1699
apply(simp add: inv_locate_b.simps inv_locate_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1700
apply(rule_tac disjI2, rule_tac disjI1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1701
apply(simp only: in_middle.simps at_begin_fst_bwtn.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1702
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1703
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1704
apply(subgoal_tac "mr = 0 \<and> lm2 = []")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1705
defer
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1706
apply(rule_tac n = n and mr = mr and lm = "lm2" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1707
               and rn = rn and n = n in zero_and_nil)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1708
apply(auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1709
apply(case_tac "lm1 = []", auto simp: tape_of_nat_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1710
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1711
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1712
lemma length_equal: "xs = ys \<Longrightarrow> length xs = length ys"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1713
by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1714
lemma [simp]: "a\<^bsup>0\<^esup> = []" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1715
by(simp add: exp_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1716
(*inv: from locate_b to locate_a*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1717
lemma [simp]: "length (a\<^bsup>b\<^esup>) = b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1718
apply(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1719
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1720
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1721
lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1722
                \<not> (\<exists>n. xs = Bk\<^bsup>n\<^esup>)\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1723
       \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1724
apply(simp add: inv_locate_b.simps inv_locate_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1725
apply(rule_tac disjI1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1726
apply(simp only: in_middle.simps at_begin_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1727
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1728
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1729
apply(subgoal_tac "tn = 0", simp add: exponent_def , auto split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1730
apply(case_tac [!] mr, simp_all add: tape_of_nat_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1731
apply(case_tac lm2, simp, erule_tac x = rn in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1732
apply(case_tac am, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1733
apply(case_tac lm2, simp, erule_tac x = rn in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1734
apply(drule_tac length_equal, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1735
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1736
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1737
lemma locate_b_2_a[intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1738
       "inv_locate_b (as, am) (q, aaa, Bk # xs) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1739
    \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1740
apply(case_tac "\<exists> n. xs = Bk\<^bsup>n\<^esup>", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1741
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1742
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1743
lemma locate_b_2_locate_a[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1744
    "\<lbrakk>\<not> a < start_of ly as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1745
      a < start_of ly as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1746
      (a - start_of ly as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1747
     inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1748
   \<Longrightarrow> (Suc a < start_of ly as + 2 * n \<longrightarrow> inv_locate_a (as, am)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1749
       (Suc ((a - start_of ly as) div 2), Bk # aaa, xs) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1750
       (\<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1751
                inv_locate_a (as, am) (n, Bk # aaa, xs) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1752
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1753
apply(subgoal_tac "n > 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1754
apply(subgoal_tac "(a - start_of ly as) div 2 = n - 1")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1755
apply(insert locate_b_2_a [of as am "n - 1" aaa xs], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1756
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1757
apply(case_tac n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1758
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1759
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1760
lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1761
           \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1762
apply(simp only: inv_locate_b.simps in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1763
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1764
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1765
      rule_tac x = tn in exI, rule_tac x = m in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1766
      rule_tac x = ml in exI, rule_tac x = mr in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1767
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1768
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1769
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1770
lemma locate_b_2_locate_a_B[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1771
 "\<lbrakk>\<not> a < start_of ly as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1772
   a < start_of ly as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1773
   (a - start_of ly as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1774
   inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1775
   \<Longrightarrow> (Suc a < start_of ly as + 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1776
     inv_locate_a (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1777
            (Suc ((a - start_of ly as) div 2), Bk # aaa, []) ires) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1778
    \<and> (\<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1779
                  inv_locate_a (as, am) (n, Bk # aaa, []) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1780
apply(insert locate_b_2_locate_a [of a ly as n am aaa "[]"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1781
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1782
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1783
(*inv: from locate_b to after_write*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1784
lemma inv_locate_b_2_after_write[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1785
      "inv_locate_b (as, am) (n, aaa, Bk # xs) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1786
      \<Longrightarrow> inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1787
          (Suc (Suc (2 * n)), aaa, Oc # xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1788
apply(auto simp: in_middle.simps inv_after_write.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1789
                 abc_lm_v.simps abc_lm_s.simps  inv_locate_b.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1790
apply(subgoal_tac [!] "mr = 0", auto simp: exponent_def split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1791
apply(subgoal_tac "lm2 = []", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1792
apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1793
      rule_tac x = "lm1" in exI, simp, rule_tac x = "[]" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1794
apply(case_tac "Suc (length lm1) - length am", simp, simp only: rep_ind, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1795
apply(subgoal_tac "length lm1 - length am = nat", simp, arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1796
apply(drule_tac length_equal, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1797
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1798
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1799
lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1800
     inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1801
                     (Suc (Suc (2 * n)), aaa, [Oc]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1802
apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1803
by(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1804
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1805
(*inv: from after_write to after_move*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1806
lemma [simp]: "inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1807
                \<Longrightarrow> inv_after_move (as, lm) (2 * n + 3, Oc # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1808
apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1809
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1810
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1811
lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1812
                )) (Suc (Suc (2 * n)), aaa, Bk # xs) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1813
apply(simp add: inv_after_write.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1814
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1815
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1816
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1817
 "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1818
                        (Suc (Suc (2 * n)), aaa, []) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1819
apply(simp add: inv_after_write.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1820
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1821
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1822
(*inv: from after_move to after_clear*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1823
lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1824
                \<Longrightarrow> inv_after_clear (as, lm) (s', l, Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1825
apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1826
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1827
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1828
(*inv: from after_move to on_leftmoving*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1829
lemma inv_after_move_2_inv_on_left_moving[simp]:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1830
   "inv_after_move (as, lm) (s, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1831
   \<Longrightarrow> (l = [] \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1832
         inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1833
      (l \<noteq> [] \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1834
         inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1835
apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1836
apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1837
                rule disjI1, simp only: inv_on_left_moving_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1838
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1839
apply(subgoal_tac "lm2 = []")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1840
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1841
    rule_tac x = m in exI, rule_tac x = m in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1842
    rule_tac x = 1 in exI,  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1843
    rule_tac x = "rn - 1" in exI, simp, case_tac rn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1844
apply(auto simp: exponent_def  intro: BkCons_nil split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1845
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1846
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1847
lemma [elim]: "[] = <lm::nat list> \<Longrightarrow> lm = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1848
using tape_of_nl_nil_eq[of lm]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1849
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1850
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1851
lemma inv_after_move_2_inv_on_left_moving_B[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1852
    "inv_after_move (as, lm) (s, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1853
      \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1854
          (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1855
apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1856
apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, rule disjI1,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1857
        simp only: inv_on_left_moving_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1858
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1859
apply(subgoal_tac "lm2 = []")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1860
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1861
      rule_tac x = m in exI, rule_tac x = m in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1862
      rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1863
apply(auto simp: exponent_def  tape_of_nl_nil_eq  intro: BkCons_nil  split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1864
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1865
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1866
(*inv: from after_clear to on_right_moving*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1867
lemma [simp]: "Oc # r = replicate rn Bk = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1868
apply(case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1869
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1870
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1871
lemma inv_after_clear_2_inv_on_right_moving[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1872
     "inv_after_clear (as, lm) (2 * n + 4, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1873
      \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1874
apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1875
apply(subgoal_tac "lm2 \<noteq> []")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1876
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1877
      rule_tac x = "hd lm2" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1878
apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1879
apply(simp add: exponent_def, rule conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1880
apply(case_tac [!] "lm2::nat list", auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1881
apply(case_tac rn, auto split: if_splits simp: tape_of_nat_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1882
apply(case_tac list, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1883
     simp add:  tape_of_nl_abv tape_of_nat_list.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1884
apply(erule_tac x = "rn - 1" in allE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1885
      case_tac rn, auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1886
apply(case_tac list, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1887
     simp add:  tape_of_nl_abv tape_of_nat_list.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1888
apply(erule_tac x = "rn - 1" in allE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1889
      case_tac rn, auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1890
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1891
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1892
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1893
lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires\<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1894
               inv_after_clear (as, lm) (2 * n + 4, l, [Bk]) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1895
by(auto simp: inv_after_clear.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1896
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1897
lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1898
             \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, []) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1899
by(insert 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1900
    inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1901
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1902
(*inv: from on_right_moving to on_right_movign*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1903
lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1904
      \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Oc # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1905
apply(auto simp: inv_on_right_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1906
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1907
           rule_tac x = "ml + mr" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1908
apply(rule_tac x = "Suc ml" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1909
           rule_tac x = "mr - 1" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1910
apply(case_tac mr, auto simp: exponent_def )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1911
apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1912
      rule_tac x = "ml + mr" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1913
apply(rule_tac x = "Suc ml" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1914
      rule_tac x = "mr - 1" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1915
apply(case_tac mr, auto split: if_splits simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1916
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1917
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1918
lemma inv_on_right_moving_2_inv_on_right_moving[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1919
     "inv_on_right_moving (as, lm) (2 * n + 5, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1920
     \<Longrightarrow> inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1921
apply(auto simp: inv_on_right_moving.simps inv_after_write.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1922
apply(case_tac mr, auto simp: exponent_def split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1923
apply(case_tac [!] mr, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1924
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1925
      
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1926
lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires\<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1927
             inv_on_right_moving (as, lm) (2 * n + 5, l, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1928
apply(auto simp: inv_on_right_moving.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1929
apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1930
apply (rule_tac x = m in exI, auto split: if_splits simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1931
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1932
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1933
(*inv: from on_right_moving to after_write*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1934
lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1935
       \<Longrightarrow> inv_after_write (as, lm) (Suc (Suc (2 * n)), l, [Oc]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1936
apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1937
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1938
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1939
(*inv: from on_left_moving to on_left_moving*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1940
lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1941
               (s, l, Oc # r) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1942
apply(auto simp: inv_on_left_moving_in_middle_B.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1943
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1944
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1945
lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1946
             = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1947
apply(auto simp: inv_on_left_moving_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1948
apply(case_tac [!] mr, auto simp: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1949
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1950
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1951
lemma [intro]: "\<exists>rna. Oc # Oc\<^bsup>m\<^esup> @ Bk # <lm> @ Bk\<^bsup>rn\<^esup> = <m # lm> @ Bk\<^bsup>rna\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1952
apply(case_tac lm, simp add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1953
apply(rule_tac x = "Suc rn" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1954
apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1955
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1956
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1957
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1958
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1959
  "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1960
    hd l = Bk; l \<noteq> []\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1961
     inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1962
apply(case_tac l, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1963
apply(simp only: inv_on_left_moving_norm.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1964
                 inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1965
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1966
apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1967
apply(case_tac [!] ml, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1968
apply(rule_tac [!] x = 0 in exI, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1969
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1970
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1971
lemma [simp]: "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1972
                hd l = Oc; l \<noteq> []\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1973
            \<Longrightarrow> inv_on_left_moving_norm (as, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1974
                                        (s, tl l, Oc # Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1975
apply(simp only: inv_on_left_moving_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1976
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1977
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1978
      rule_tac x = m in exI, rule_tac x = "ml - 1" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1979
      rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1980
apply(case_tac ml, auto simp: exponent_def split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1981
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1982
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1983
lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1984
     \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1985
apply(auto simp: inv_on_left_moving_norm.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1986
                 inv_on_left_moving_in_middle_B.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1987
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1988
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1989
lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1990
    \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1991
 \<and>  (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1992
apply(simp add: inv_on_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1993
apply(case_tac "l \<noteq> []", rule conjI, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1994
apply(case_tac "hd l", simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1995
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1996
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1997
(*inv: from on_left_moving to check_left_moving*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1998
lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1999
                                      (s, Bk # list, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2000
          \<Longrightarrow> inv_check_left_moving_on_leftmost (as, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2001
                                      (s', list, Bk # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2002
apply(auto simp: inv_on_left_moving_in_middle_B.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2003
                 inv_check_left_moving_on_leftmost.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2004
apply(case_tac [!] "rev lm1", simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2005
apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2006
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2007
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2008
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2009
    "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2010
by(auto simp: inv_check_left_moving_in_middle.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2011
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2012
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2013
 "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2014
  inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2015
apply(auto simp: inv_on_left_moving_in_middle_B.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2016
                 inv_check_left_moving_on_leftmost.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2017
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2018
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2019
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2020
lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2021
                                       (s, list, Oc # r) ires= False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2022
by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2023
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2024
lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2025
                                         (s, Oc # list, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2026
 \<Longrightarrow> inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2027
apply(auto simp: inv_on_left_moving_in_middle_B.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2028
                 inv_check_left_moving_in_middle.simps  split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2029
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2030
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2031
lemma inv_on_left_moving_2_check_left_moving[simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2032
 "inv_on_left_moving (as, lm) (s, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2033
 \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2034
 \<and> (l \<noteq> [] \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2035
      inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2036
apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2037
apply(case_tac l, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2038
apply(case_tac a, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2039
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2040
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2041
lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2042
apply(auto simp: inv_on_left_moving_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2043
by(case_tac [!] mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2044
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2045
lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2046
     inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2047
apply(simp add: inv_on_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2048
apply(auto simp: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2049
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2050
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2051
lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2052
apply(simp add: inv_on_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2053
apply(simp add: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2054
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2055
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2056
lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2057
 \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2058
    (l \<noteq> [] \<longrightarrow> inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2059
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2060
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2061
lemma Oc_Bk_Cons_ex[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2062
 "Oc # Bk # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2063
                             \<exists>ln. list = <tl (lm)> @ Bk\<^bsup>ln\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2064
apply(case_tac "lm", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2065
apply(case_tac ln, simp_all add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2066
apply(case_tac lista, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2067
      auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2068
apply(case_tac [!] a, auto simp: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2069
apply(case_tac ln, simp, rule_tac x = nat in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2070
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2071
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2072
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2073
  "Oc # Bk # list = <rev lm1::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2074
      \<exists>rna. Oc # Bk # <lm2> @ Bk\<^bsup>rn\<^esup> = <hd (rev lm1) # lm2> @ Bk\<^bsup>rna\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2075
apply(frule Oc_Bk_Cons, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2076
apply(case_tac lm2, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2077
     auto simp: tape_of_nl_abv tape_of_nat_list.simps  exponent_def )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2078
apply(rule_tac x = "Suc rn" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2079
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2080
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2081
(*inv: from check_left_moving to on_left_moving*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2082
lemma [intro]: "\<exists>rna. a # a\<^bsup>rn\<^esup> = a\<^bsup>rna\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2083
apply(rule_tac x = "Suc rn" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2084
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2085
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2086
lemma 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2087
inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2088
"inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2089
  \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2090
apply(simp only: inv_check_left_moving_in_middle.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2091
                 inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2092
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2093
apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2094
      rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2095
apply(case_tac [!] "rev lm1",simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2096
apply(case_tac [!] a, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2097
apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2098
apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2099
apply(case_tac [!] lista, simp_all add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2100
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2101
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2102
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2103
 "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2104
     inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2105
apply(auto simp: inv_check_left_moving_in_middle.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2106
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2107
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2108
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2109
 "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2110
   \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2111
apply(insert 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2112
inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2113
                  as lm n "[]" r], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2114
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2115
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2116
lemma [simp]: "a\<^bsup>0\<^esup> = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2117
apply(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2118
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2119
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2120
lemma [simp]: "inv_check_left_moving_in_middle (as, lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2121
                       (s, Oc # list, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2122
   \<Longrightarrow> inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2123
apply(auto simp: inv_check_left_moving_in_middle.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2124
                 inv_on_left_moving_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2125
apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2126
      rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2127
apply(rule_tac conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2128
apply(case_tac "rev lm1", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2129
apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2130
apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2131
apply(case_tac [!] "rev lm1", simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2132
apply(case_tac [!] a, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2133
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2134
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2135
lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2136
\<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2137
   (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2138
apply(case_tac l, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2139
      auto simp: inv_check_left_moving.simps inv_on_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2140
apply(case_tac a, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2141
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2142
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2143
(*inv: check_left_moving to after_left_moving*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2144
lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2145
                \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2146
apply(auto simp: inv_check_left_moving.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2147
 inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2148
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2149
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2150
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2151
lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2152
      \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, []) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2153
by(simp add: inv_check_left_moving.simps  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2154
inv_check_left_moving_in_middle.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2155
inv_check_left_moving_on_leftmost.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2156
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2157
(*inv: after_left_moving to inv_stop*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2158
lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2159
       \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2160
apply(auto simp: inv_after_left_moving.simps inv_stop.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2161
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2162
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2163
lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2164
             \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, []) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2165
by(auto simp: inv_after_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2166
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2167
(*inv: stop to stop*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2168
lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2169
               inv_stop (as, lm) (y, l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2170
apply(simp add: inv_stop.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2171
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2172
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2173
lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2174
apply(auto simp: inv_after_clear.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2175
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2176
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2177
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2178
  "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2179
by(auto simp: inv_after_left_moving.simps  )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2180
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2181
lemma start_of_not0: "as \<noteq> 0 \<Longrightarrow> start_of ly as > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2182
apply(rule startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2183
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2184
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2185
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2186
  The single step currectness of the TM complied from Abacus instruction @{text "Inc n"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2187
  It shows every single step execution of this TM keeps the invariant.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2188
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2189
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2190
lemma inc_inv_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2191
  assumes 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2192
  -- {* Invariant holds on the start *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2193
      h11: "inc_inv ly n (as, am) tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2194
  -- {* The layout of Abacus program @{text "aprog"} is @{text "ly"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2195
  and h12: "ly = layout_of aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2196
  -- {* The instruction at position @{text "as"} is @{text "Inc n"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2197
  and h21: "abc_fetch as aprog = Some (Inc n)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2198
  -- {* TM not yet reach the final state, where @{text "start_of ly as + 2*n + 9"} is the state
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2199
        where the current TM stops and the next TM starts. *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2200
  and h22: "(\<lambda> (s, l, r). s \<noteq> start_of ly as + 2*n + 9) tc"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2201
  shows 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2202
  -- {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2203
  Single step execution of the TM keeps the invaraint, where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2204
  the TM compiled from @{text "Inc n"} is @{text "(ci ly (start_of ly as) (Inc n))"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2205
  @{text "start_of ly as - Suc 0)"} is the offset used to execute this {\em shifted}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2206
  TM.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2207
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2208
  "inc_inv ly n (as, am) (t_step tc (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0)) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2209
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2210
  from h21 h22  have h3 : "start_of (layout_of aprog) as > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2211
    apply(case_tac as, simp add: start_of.simps abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2212
    apply(insert start_of_not0[of as "layout_of aprog"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2213
    done    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2214
  from h11 h12 and h21 h22 and this show ?thesis 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2215
    apply(case_tac tc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2216
    apply(case_tac "a = 0", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2217
      auto split:if_splits simp add:t_step.simps,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2218
      tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2219
    apply (simp_all add:fetch_simps new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2220
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2221
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2222
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2223
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2224
lemma t_steps_ind: "t_steps tc (p, off) (Suc n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2225
                 = t_step (t_steps tc (p, off) n) (p, off)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2226
apply(induct n arbitrary: tc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2227
apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2228
apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2229
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2230
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2231
definition lex_pair :: "((nat \<times> nat) \<times> (nat \<times> nat)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2232
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2233
  "lex_pair \<equiv> less_than <*lex*> less_than"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2234
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2235
definition lex_triple :: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2236
   "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2237
  where "lex_triple \<equiv> less_than <*lex*> lex_pair"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2238
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2239
definition lex_square :: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2240
    "((nat \<times> nat \<times> nat \<times> nat) \<times> (nat \<times> nat \<times> nat \<times> nat)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2241
  where "lex_square \<equiv> less_than <*lex*> lex_triple"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2242
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2243
fun abc_inc_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2244
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2245
  "abc_inc_stage1 (s, l, r) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2246
            (if s = 0 then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2247
             else if s \<le> ss+2*n+1 then 5
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2248
             else if s\<le> ss+2*n+5 then 4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2249
             else if s \<le> ss+2*n+7 then 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2250
             else if s = ss+2*n+8 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2251
             else 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2252
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2253
fun abc_inc_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2254
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2255
  "abc_inc_stage2 (s, l, r) ss n =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2256
                (if s \<le> ss + 2*n + 1 then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2257
                 else if s = ss + 2*n + 2 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2258
                 else if s = ss + 2*n + 3 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2259
                 else if s = ss + 2*n + 4 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2260
                 else if s = ss + 2*n + 5 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2261
                                  if r \<noteq> [] then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2262
                                  else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2263
                 else if s = ss+2*n+6 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2264
                 else if s = ss+2*n+7 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2265
                 else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2266
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2267
fun abc_inc_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow>  nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2268
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2269
  "abc_inc_stage3 (s, l, r) ss n ires = (
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2270
              if s = ss + 2*n + 3 then 4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2271
              else if s = ss + 2*n + 4 then 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2272
              else if s = ss + 2*n + 5 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2273
                   if r \<noteq> [] \<and> hd r = Oc then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2274
                   else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2275
              else if s = ss + 2*n + 2 then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2276
              else if s = ss + 2*n + 6 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2277
                      if l = Bk # ires \<and> r \<noteq> [] \<and>  hd r = Oc then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2278
                      else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2279
              else if s = ss + 2*n + 7 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2280
                      if r \<noteq> [] \<and> hd r = Oc then 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2281
                      else 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2282
              else ss+2*n+9 - s)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2283
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2284
fun abc_inc_stage4 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2285
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2286
  "abc_inc_stage4 (s, l, r) ss n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2287
            (if s \<le> ss+2*n+1 \<and> (s - ss) mod 2 = 0 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2288
                if (r\<noteq>[] \<and> hd r = Oc) then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2289
                else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2290
             else if (s \<le> ss+2*n+1 \<and> (s - ss) mod 2 = Suc 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2291
                                                 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2292
             else if s = ss + 2*n + 6 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2293
                  if l = Bk # ires \<and> hd r = Bk then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2294
                  else Suc (length l)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2295
             else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2296
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2297
fun abc_inc_measure :: "(t_conf \<times> nat \<times> nat \<times> block list) \<Rightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2298
                        (nat \<times> nat \<times> nat \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2299
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2300
  "abc_inc_measure (c, ss, n, ires) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2301
     (abc_inc_stage1 c ss n, abc_inc_stage2 c ss n, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2302
      abc_inc_stage3 c ss n ires, abc_inc_stage4 c ss n ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2303
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2304
definition abc_inc_LE :: "(((nat \<times> block list \<times> block list) \<times> nat \<times> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2305
       nat \<times> block list) \<times> ((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2306
  where "abc_inc_LE \<equiv> (inv_image lex_square abc_inc_measure)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2307
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2308
lemma wf_lex_triple: "wf lex_triple"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2309
by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2310
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2311
lemma wf_lex_square: "wf lex_square"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2312
by (auto intro:wf_lex_triple simp:lex_triple_def lex_square_def lex_pair_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2313
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2314
lemma wf_abc_inc_le[intro]: "wf abc_inc_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2315
by(auto intro:wf_inv_image wf_lex_square simp:abc_inc_LE_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2316
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2317
(********************************************************************)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2318
declare inc_inv.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2319
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2320
lemma halt_lemma2': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2321
  "\<lbrakk>wf LE;  \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2322
    (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); Q (f 0)\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2323
      \<Longrightarrow> \<exists> n. P (f n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2324
apply(intro exCI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2325
apply(subgoal_tac "\<forall> n. Q (f n)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2326
apply(drule_tac f = f in wf_inv_image)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2327
apply(simp add: inv_image_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2328
apply(erule wf_induct, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2329
apply(erule_tac x = x in allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2330
apply(erule_tac x = n in allE, erule_tac x = n in allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2331
apply(erule_tac x = "Suc x" in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2332
apply(rule_tac allI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2333
apply(induct_tac n, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2334
apply(erule_tac x = na in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2335
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2336
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2337
lemma halt_lemma2'': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2338
  "\<lbrakk>P (f n); \<not> P (f (0::nat))\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2339
         \<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2340
apply(induct n rule: nat_less_induct, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2341
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2342
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2343
lemma halt_lemma2''':
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2344
 "\<lbrakk>\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> LE;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2345
                 Q (f 0);  \<forall>i<na. \<not> P (f i)\<rbrakk> \<Longrightarrow> Q (f na)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2346
apply(induct na, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2347
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2348
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2349
lemma halt_lemma2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2350
  "\<lbrakk>wf LE;  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2351
    \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2352
    Q (f 0); \<not> P (f 0)\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2353
  \<Longrightarrow> \<exists> n. P (f n) \<and> Q (f n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2354
apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2355
apply(subgoal_tac "\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2356
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2357
apply(rule_tac x = na in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2358
apply(rule halt_lemma2''', simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2359
apply(erule_tac halt_lemma2'', simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2360
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2361
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2362
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2363
  "\<lbrakk>ly = layout_of aprog; abc_fetch as aprog = Some (Inc n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2364
    \<Longrightarrow> start_of ly (Suc as) = start_of ly as + 2*n +9"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2365
apply(case_tac as, auto simp: abc_fetch.simps start_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2366
          layout_of.simps length_of.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2367
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2368
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2369
lemma inc_inv_init: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2370
 "\<lbrakk>abc_fetch as aprog = Some (Inc n); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2371
   crsp_l ly (as, am) (start_of ly as, l, r) ires; ly = layout_of aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2372
  \<Longrightarrow> inc_inv ly n (as, am) (start_of ly as, l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2373
apply(auto simp: crsp_l.simps inc_inv.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2374
      inv_locate_a.simps at_begin_fst_bwtn.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2375
      at_begin_fst_awtn.simps at_begin_norm.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2376
apply(auto intro: startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2377
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2378
     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2379
lemma inc_inv_stop_pre[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2380
   "\<lbrakk>ly = layout_of aprog; inc_inv ly n (as, am) (s, l, r) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2381
     s = start_of ly as; abc_fetch as aprog = Some (Inc n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2382
    \<Longrightarrow>  (\<forall>na. \<not> (\<lambda>((s, l, r), ss, n', ires'). s = start_of ly (Suc as)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2383
         (t_steps (s, l, r) (ci ly (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2384
          (Inc n), start_of ly as - Suc 0) na, s, n, ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2385
       (\<lambda>((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires')
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2386
         (t_steps (s, l, r) (ci ly (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2387
             (Inc n), start_of ly as - Suc 0) na, s, n, ires) \<longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2388
       (\<lambda>((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires') 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2389
      (t_steps (s, l, r) (ci ly (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2390
              (Inc n), start_of ly as - Suc 0) (Suc na), s, n, ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2391
     ((t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2392
        start_of ly as - Suc 0) (Suc na), s, n, ires), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2393
      t_steps (s, l, r) (ci ly (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2394
        (Inc n), start_of ly as - Suc 0) na, s, n, ires) \<in> abc_inc_LE)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2395
apply(rule allI, rule impI, simp add: t_steps_ind,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2396
       rule conjI, erule_tac conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2397
apply(rule_tac inc_inv_step, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2398
apply(case_tac "t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2399
  (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2400
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2401
  fix na
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2402
  assume h1: "abc_fetch as aprog = Some (Inc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2403
    "\<not> (\<lambda>(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) as + 2 * n + 9)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2404
    (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2405
    (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2406
    (start_of (layout_of aprog) as, n, ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2407
    inc_inv (layout_of aprog) n (as, am) (t_steps (start_of (layout_of aprog) as, l, r)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2408
    (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2409
  from h1 have h2: "start_of (layout_of aprog) as > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2410
    apply(rule_tac startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2411
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2412
  from h1 and h2 show "((t_step (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2413
    (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2414
    (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2415
    start_of (layout_of aprog) as, n, ires),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2416
    t_steps (start_of (layout_of aprog) as, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2417
    (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2418
    start_of (layout_of aprog) as, n, ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2419
            \<in> abc_inc_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2420
    apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2421
               (ci (layout_of aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2422
         (start_of (layout_of aprog) as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2423
           start_of (layout_of aprog) as - Suc 0) na)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2424
    apply(case_tac "a = 0", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2425
     auto split:if_splits simp add:t_step.simps inc_inv.simps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2426
       tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2427
    apply(simp_all add:fetch_simps new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2428
    apply(auto simp add: abc_inc_LE_def  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2429
    lex_square_def lex_triple_def lex_pair_def
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2430
      inv_after_write.simps inv_after_move.simps inv_after_clear.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2431
      inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2432
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2433
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2434
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2435
lemma inc_inv_stop_pre1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2436
  "\<lbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2437
  ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2438
  abc_fetch as aprog = Some (Inc n);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2439
  s = start_of ly as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2440
  inc_inv ly n (as, am) (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2441
  \<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2442
  (\<exists> stp > 0. (\<lambda> (s', l', r').
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2443
           s' = start_of ly (Suc as) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2444
           inc_inv ly n (as, am) (s', l', r') ires) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2445
               (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2446
                        start_of ly as - Suc 0) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2447
apply(insert halt_lemma2[of abc_inc_LE 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2448
    "\<lambda> ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2449
    "(\<lambda> stp. (t_steps (s, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2450
     (ci ly (start_of ly as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2451
     start_of ly as - Suc 0) stp, s, n, ires))" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2452
    "\<lambda> ((s, l, r), ss, n'). inc_inv ly n (as, am) (s, l, r) ires"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2453
apply(insert  wf_abc_inc_le)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2454
apply(insert inc_inv_stop_pre[of ly aprog n as am s l r ires], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2455
apply(simp only: t_steps.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2456
apply(rule_tac x = na in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2457
apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2458
   (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2459
   (Inc n), start_of (layout_of aprog) as - Suc 0) na)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2460
apply(case_tac na, simp add: t_steps.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2461
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2462
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2463
lemma inc_inv_stop: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2464
  assumes program_and_layout: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2465
  -- {* There is an Abacus program @{text "aprog"} and its layout is @{text "ly"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2466
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2467
  and an_instruction:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2468
  -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2469
  "abc_fetch as aprog = Some (Inc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2470
  and the_start_state:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2471
  -- {* According to @{text "ly"} and @{text "as"}, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2472
        the start state of the TM compiled from this
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2473
        @{text "Inc n"} instruction should be @{text "s"}:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2474
     *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2475
  "s = start_of ly as"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2476
  and inv:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2477
  -- {* Invariant holds on configuration @{text "(s, l, r)"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2478
  "inc_inv ly n (as, am) (s, l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2479
  shows  -- {* After @{text "stp"} steps of execution, the compiled 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2480
            TM reaches the start state of next compiled TM and the invariant 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2481
            still holds.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2482
            *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2483
      "(\<exists> stp > 0. (\<lambda> (s', l', r').
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2484
           s' = start_of ly (Suc as) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2485
           inc_inv ly n (as, am) (s', l', r') ires) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2486
               (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2487
                        start_of ly as - Suc 0) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2488
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2489
  from inc_inv_stop_pre1 [OF  program_and_layout an_instruction the_start_state inv] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2490
  show ?thesis .
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2491
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2492
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2493
lemma inc_inv_stop_cond: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2494
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2495
    s' = start_of ly (as + 1); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2496
    inc_inv ly n (as, lm) (s', (l', r')) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2497
    abc_fetch as aprog = Some (Inc n)\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2498
    crsp_l ly (Suc as, abc_lm_s lm n (Suc (abc_lm_v lm n))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2499
                                                (s', l', r') ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2500
apply(subgoal_tac "s' = start_of ly as + 2*n + 9", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2501
apply(auto simp: inc_inv.simps inv_stop.simps crsp_l.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2502
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2503
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2504
lemma inc_crsp_ex_pre:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2505
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2506
    crsp_l ly (as, am) tc ires;  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2507
    abc_fetch as aprog = Some (Inc n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2508
 \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2509
                (t_steps tc (ci ly (start_of ly as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2510
                                start_of ly as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2511
proof(case_tac tc, simp add: abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2512
  fix a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2513
  assume h1: "ly = layout_of aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2514
         "crsp_l (layout_of aprog) (as, am) (a, b, c) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2515
         "abc_fetch as aprog = Some (Inc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2516
  hence h2: "a = start_of ly as"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2517
    by(auto simp: crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2518
  from h1 and h2 have h3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2519
       "inc_inv ly n (as, am) (start_of ly as, b, c) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2520
    by(rule_tac inc_inv_init, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2521
  from h1 and h2 and h3 have h4:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2522
       "(\<exists> stp > 0. (\<lambda> (s', l', r'). s' = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2523
           start_of ly (Suc as) \<and> inc_inv ly n (as, am) (s', l', r') ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2524
         (t_steps (a, b, c) (ci ly (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2525
                 (Inc n), start_of ly as - Suc 0) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2526
    apply(rule_tac inc_inv_stop, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2527
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2528
  from h1 and h2 and h3 and h4 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2529
     "\<exists>stp > 0. crsp_l (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2530
        (Suc as, abc_lm_s am n (Suc (abc_lm_v am n)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2531
       (t_steps (a, b, c) (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2532
          (start_of (layout_of aprog) as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2533
              start_of (layout_of aprog) as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2534
    apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2535
    apply(rule_tac x = stp in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2536
    apply(case_tac "(t_steps (a, b, c) (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2537
         (start_of (layout_of aprog) as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2538
             start_of (layout_of aprog) as - Suc 0) stp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2539
    apply(rule_tac inc_inv_stop_cond, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2540
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2541
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2542
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2543
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2544
  The total correctness of the compilaton of @{text "Inc n"} instruction.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2545
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2546
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2547
lemma inc_crsp_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2548
  assumes layout: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2549
  -- {* For any Abacus program @{text "aprog"}, assuming its layout is @{text "ly"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2550
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2551
  and corresponds: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2552
  -- {* Abacus configuration @{text "(as, am)"} is in correspondence with 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2553
        TM configuration @{text "tc"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2554
  "crsp_l ly (as, am) tc ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2555
  and inc:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2556
  -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2557
  "abc_fetch as aprog = Some (Inc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2558
  shows
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2559
  -- {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2560
  After @{text "stp"} steps of execution, the TM compiled from this @{text "Inc n"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2561
  stops with a configuration which corresponds to the Abacus configuration obtained
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2562
  from the execution of this @{text "Inc n"} instruction.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2563
  *} 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2564
  "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2565
                       (t_steps tc (ci ly (start_of ly as) (Inc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2566
                                start_of ly as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2567
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2568
  from inc_crsp_ex_pre [OF layout corresponds inc] show ?thesis .
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2569
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2570
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2571
(*******End: inc crsp********)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2572
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2573
(*******Begin: dec crsp******)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2574
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2575
subsubsection {* The compilation of @{text "Dec n e"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2576
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2577
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2578
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2579
  The lemmas in this section lead to the correctness of the compilation 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2580
  of @{text "Dec n e"} instruction using the same techniques as 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2581
  @{text "Inc n"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2582
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2583
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2584
type_synonym dec_inv_t = "(nat * nat list) \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow>  bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2585
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2586
fun dec_first_on_right_moving :: "nat \<Rightarrow> dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2587
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2588
  "dec_first_on_right_moving n (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2589
               (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2590
         ml + mr = Suc m \<and> length lm1 = n \<and> ml > 0 \<and> m > 0 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2591
             (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2592
                          else  l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2593
    ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2> @ (Bk\<^bsup>rn\<^esup>)) \<or> (r = (Oc\<^bsup>mr\<^esup>) \<and> lm2 = [])))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2594
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2595
fun dec_on_right_moving :: "dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2596
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2597
  "dec_on_right_moving (as, lm) (s, l, r) ires =  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2598
   (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2599
                             ml + mr = Suc (Suc m) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2600
   (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2601
                else  l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2602
   ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2> @ (Bk\<^bsup>rn\<^esup>)) \<or> (r = (Oc\<^bsup>mr\<^esup>) \<and> lm2 = [])))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2603
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2604
fun dec_after_clear :: "dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2605
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2606
  "dec_after_clear (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2607
              (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2608
                ml + mr = Suc m \<and> ml = Suc m \<and> r \<noteq> [] \<and> r \<noteq> [] \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2609
               (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2610
                            else l = (Oc\<^bsup>ml \<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2611
               (tl r = Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>) \<or> tl r = [] \<and> lm2 = []))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2612
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2613
fun dec_after_write :: "dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2614
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2615
  "dec_after_write (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2616
         (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2617
       ml + mr = Suc m \<and> ml = Suc m \<and> lm2 \<noteq> [] \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2618
       (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2619
                    else l = Bk # (Oc\<^bsup>ml \<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2620
       tl r = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2621
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2622
fun dec_right_move :: "dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2623
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2624
  "dec_right_move (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2625
        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2626
            \<and> ml = Suc m \<and> mr = (0::nat) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2627
              (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2628
                          else l = Bk # Oc\<^bsup>ml\<^esup>@ [Bk] @ <rev lm1> @ Bk # Bk # ires) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2629
           \<and> (r = Bk # <lm2> @ Bk\<^bsup>rn\<^esup>\<or> r = [] \<and> lm2 = []))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2630
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2631
fun dec_check_right_move :: "dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2632
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2633
  "dec_check_right_move (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2634
        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2635
           ml = Suc m \<and> mr = (0::nat) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2636
           (if lm1 = [] then l = Bk # Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2637
                       else l = Bk # Bk # Oc\<^bsup>ml \<^esup>@ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2638
           r = <lm2> @ Bk\<^bsup>rn\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2639
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2640
fun dec_left_move :: "dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2641
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2642
  "dec_left_move (as, lm) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2643
    (\<exists> lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \<and>   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2644
    rn > 0 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2645
   (if lm1 = [] then l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2646
    else l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> r = Bk\<^bsup>rn\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2647
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2648
declare
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2649
  dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2650
  dec_after_write.simps[simp del] dec_left_move.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2651
  dec_check_right_move.simps[simp del] dec_right_move.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2652
  dec_first_on_right_moving.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2653
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2654
fun inv_locate_n_b :: "inc_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2655
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2656
  "inv_locate_n_b (as, lm) (s, l, r) ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2657
    (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2658
     length lm1 = s \<and> m + 1 = ml + mr \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2659
     ml = 1 \<and> tn = s + 1 - length lm \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2660
     (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2661
      else l = Oc\<^bsup>ml\<^esup>@Bk#<rev lm1>@Bk#Bk#ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2662
     (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2>@ (Bk\<^bsup>rn\<^esup>) \<or> (lm2 = [] \<and> r = (Oc\<^bsup>mr\<^esup>)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2663
  )"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2664
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2665
fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2666
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2667
  "dec_inv_1 ly n e (as, am) (s, l, r) ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2668
           (let ss = start_of ly as in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2669
            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2670
            let am'' = abc_lm_s am n (abc_lm_v am n) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2671
              if s = start_of ly e then  inv_stop (as, am'') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2672
              else if s = ss then False
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2673
              else if ss \<le> s \<and> s < ss + 2*n then
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2674
                   if (s - ss) mod 2 = 0 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2675
                        inv_locate_a (as, am) ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2676
                    \<or> inv_locate_a (as, am'') ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2677
                   else 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2678
                     inv_locate_b (as, am) ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2679
                  \<or> inv_locate_b (as, am'') ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2680
              else if s = ss + 2 * n then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2681
                  inv_locate_a (as, am) (n, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2682
                \<or> inv_locate_a (as, am'') (n, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2683
              else if s = ss + 2 * n + 1 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2684
                  inv_locate_b (as, am) (n, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2685
              else if s = ss + 2 * n + 13 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2686
                  inv_on_left_moving (as, am'') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2687
              else if s = ss + 2 * n + 14 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2688
                  inv_check_left_moving (as, am'') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2689
              else if s = ss + 2 * n + 15 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2690
                  inv_after_left_moving (as, am'') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2691
              else False)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2692
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2693
fun dec_inv_2 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2694
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2695
  "dec_inv_2 ly n e (as, am) (s, l, r) ires =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2696
           (let ss = start_of ly as in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2697
            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2698
            let am'' = abc_lm_s am n (abc_lm_v am n) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2699
              if s = 0 then False
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2700
              else if s = ss then False
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2701
              else if ss \<le> s \<and> s < ss + 2*n then
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2702
                   if (s - ss) mod 2 = 0 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2703
                      inv_locate_a (as, am) ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2704
                   else inv_locate_b (as, am) ((s - ss) div 2, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2705
              else if s = ss + 2 * n then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2706
                      inv_locate_a (as, am) (n, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2707
              else if s = ss + 2 * n + 1 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2708
                      inv_locate_n_b (as, am) (n, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2709
              else if s = ss + 2 * n + 2 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2710
                      dec_first_on_right_moving n (as, am'') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2711
              else if s = ss + 2 * n + 3 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2712
                      dec_after_clear (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2713
              else if s = ss + 2 * n + 4 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2714
                      dec_right_move (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2715
              else if s = ss + 2 * n + 5 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2716
                      dec_check_right_move (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2717
              else if s = ss + 2 * n + 6 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2718
                      dec_left_move (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2719
              else if s = ss + 2 * n + 7 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2720
                      dec_after_write (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2721
              else if s = ss + 2 * n + 8 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2722
                      dec_on_right_moving (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2723
              else if s = ss + 2 * n + 9 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2724
                      dec_after_clear (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2725
              else if s = ss + 2 * n + 10 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2726
                      inv_on_left_moving (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2727
              else if s = ss + 2 * n + 11 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2728
                      inv_check_left_moving (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2729
              else if s = ss + 2 * n + 12 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2730
                      inv_after_left_moving (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2731
              else if s = ss + 2 * n + 16 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2732
                      inv_stop (as, am') (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2733
              else False)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2734
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2735
(*begin: dec_fetch lemmas*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2736
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2737
lemma dec_fetch_locate_a_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2738
      "\<lbrakk>start_of ly as \<le> a;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2739
        a < start_of ly as + 2 * n; start_of ly as > 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2740
        a - start_of ly as = 2 * q\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2741
       \<Longrightarrow> fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2742
         (start_of ly as) (Dec n e)) (Suc (2 * q))  Oc = (R, a + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2743
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2744
                  nth_of.simps tshift.simps nth_append Suc_pre)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2745
apply(subgoal_tac "(findnth n ! Suc (4 * q)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2746
                          findnth (Suc q) ! (4 * q + 1)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2747
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2748
apply(subgoal_tac " findnth n !(4 * q + 1) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2749
                          findnth (Suc q) ! (4 * q + 1)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2750
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2751
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2752
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2753
lemma  dec_fetch_locate_a_b:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2754
       "\<lbrakk>start_of ly as \<le> a; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2755
         a < start_of ly as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2756
         start_of ly as > 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2757
         a - start_of ly as = 2 * q\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2758
       \<Longrightarrow> fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2759
              (Suc (2 * q))  Bk = (W1, a)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2760
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2761
                  nth_of.simps tshift.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2762
apply(subgoal_tac "(findnth n ! (4 * q)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2763
                       findnth (Suc q) ! (4 * q )")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2764
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2765
apply(subgoal_tac " findnth n !(4 * q + 0) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2766
                       findnth (Suc q) ! (4 * q + 0)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2767
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2768
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2769
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2770
lemma dec_fetch_locate_b_o:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2771
      "\<lbrakk>start_of ly as \<le> a; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2772
        a < start_of ly as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2773
        (a - start_of ly as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2774
        start_of ly as> 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2775
       \<Longrightarrow> fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2776
                       (Suc (a - start_of ly as)) Oc = (R, a)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2777
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2778
                  nth_of.simps tshift.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2779
apply(subgoal_tac "\<exists> q. (a - start_of ly as) = 2 * q + 1", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2780
apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2781
                                findnth (Suc q) ! (4 * q + 3)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2782
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2783
apply(subgoal_tac " findnth n ! (4 * q + 3) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2784
                 findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2785
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2786
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2787
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2788
lemma dec_fetch_locate_b_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2789
      "\<lbrakk>\<not> a < start_of ly as; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2790
        a < start_of ly as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2791
       (a - start_of ly as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2792
        start_of ly as > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2793
       \<Longrightarrow> fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2794
              (Suc (a - start_of ly as))  Bk = (R, a + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2795
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2796
                  nth_of.simps tshift.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2797
apply(subgoal_tac "\<exists> q. (a - start_of ly as) = 2 * q + 1", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2798
apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2799
                          findnth (Suc q) ! (4 * q + 2)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2800
apply(simp add: findnth.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2801
apply(subgoal_tac " findnth n ! (4 * q + 2) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2802
                          findnth (Suc q) ! (4 * q + 2)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2803
apply(rule_tac findnth_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2804
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2805
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2806
lemma dec_fetch_locate_n_a_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2807
       "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2808
                       (start_of ly as) (Dec n e)) (Suc (2 * n))  Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2809
       = (R, start_of ly as + 2*n + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2810
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2811
                  nth_of.simps tshift.simps nth_append tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2812
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2813
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2814
lemma dec_fetch_locate_n_a_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2815
       "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2816
                       (start_of ly as) (Dec n e)) (Suc (2 * n))  Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2817
       = (W1, start_of ly as + 2*n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2818
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2819
                  nth_of.simps tshift.simps nth_append tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2820
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2821
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2822
lemma dec_fetch_locate_n_b_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2823
       "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2824
            fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2825
                (start_of ly as) (Dec n e)) (Suc (Suc (2 * n)))  Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2826
      = (R, start_of ly as + 2*n + 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2827
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2828
                  nth_of.simps tshift.simps nth_append tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2829
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2830
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2831
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2832
lemma dec_fetch_locate_n_b_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2833
       "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2834
       fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2835
                  (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2836
      = (L, start_of ly as + 2*n + 13)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2837
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2838
                  nth_of.simps tshift.simps nth_append tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2839
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2840
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2841
lemma dec_fetch_first_on_right_move_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2842
      "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2843
       fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2844
             (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n))))  Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2845
     = (R, start_of ly as + 2*n + 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2846
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2847
                  nth_of.simps tshift.simps nth_append tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2848
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2849
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2850
lemma dec_fetch_first_on_right_move_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2851
      "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2852
      fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2853
                             (Suc (Suc (Suc (2 * n))))  Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2854
     = (L, start_of ly as + 2*n + 3)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2855
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2856
                  nth_of.simps tshift.simps nth_append tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2857
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2858
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2859
lemma [simp]: "fetch x (a + 2 * n) b = fetch x (2 * n + a) b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2860
thm arg_cong
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2861
apply(rule_tac x = "a + 2*n" and y = "2*n + a" in arg_cong, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2862
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2863
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2864
lemma dec_fetch_first_after_clear_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2865
     "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2866
                      (start_of ly as) (Dec n e)) (2 * n + 4) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2867
    = (W0, start_of ly as + 2*n + 3)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2868
apply(auto simp: ci.simps findnth.simps tshift.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2869
                          tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2870
apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2871
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2872
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2873
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2874
lemma dec_fetch_first_after_clear_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2875
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2876
     fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2877
                   (start_of ly as) (Dec n e)) (2 * n + 4) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2878
    = (R, start_of ly as + 2*n + 4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2879
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2880
               tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2881
apply(subgoal_tac "2*n + 4= Suc (2*n + 3)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2882
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2883
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2884
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2885
lemma dec_fetch_right_move_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2886
     "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2887
                          (start_of ly as) (Dec n e)) (2 * n + 5) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2888
    = (R, start_of ly as + 2*n + 5)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2889
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2890
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2891
apply(subgoal_tac "2*n + 5= Suc (2*n + 4)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2892
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2893
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2894
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2895
lemma dec_fetch_check_right_move_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2896
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2897
      fetch (ci (layout_of aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2898
                (start_of ly as) (Dec n e)) (2 * n + 6) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2899
    = (L, start_of ly as + 2*n + 6)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2900
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2901
               tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2902
apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2903
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2904
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2905
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2906
lemma dec_fetch_check_right_move_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2907
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2908
    fetch (ci (layout_of aprog) (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2909
                      (Dec n e)) (2 * n + 6) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2910
    = (L, start_of ly as + 2*n + 7)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2911
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2912
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2913
apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2914
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2915
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2916
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2917
lemma dec_fetch_left_move_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2918
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2919
     fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2920
             (start_of ly as) (Dec n e)) (2 * n + 7) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2921
    = (L, start_of ly as + 2*n + 10)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2922
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2923
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2924
apply(subgoal_tac "2*n + 7 = Suc (2*n + 6)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2925
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2926
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2927
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2928
lemma dec_fetch_after_write_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2929
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2930
    fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2931
                   (start_of ly as) (Dec n e)) (2 * n + 8) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2932
    = (W1, start_of ly as + 2*n + 7)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2933
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2934
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2935
apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2936
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2937
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2938
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2939
lemma dec_fetch_after_write_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2940
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2941
     fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2942
                   (start_of ly as) (Dec n e)) (2 * n + 8) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2943
    = (R, start_of ly as + 2*n + 8)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2944
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2945
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2946
apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2947
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2948
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2949
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2950
lemma dec_fetch_on_right_move_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2951
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2952
     fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2953
                   (start_of ly as) (Dec n e)) (2 * n + 9) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2954
    = (L, start_of ly as + 2*n + 9)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2955
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2956
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2957
apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2958
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2959
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2960
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2961
lemma dec_fetch_on_right_move_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2962
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2963
     fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2964
             (start_of ly as) (Dec n e)) (2 * n + 9) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2965
    = (R, start_of ly as + 2*n + 8)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2966
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2967
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2968
apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2969
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2970
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2971
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2972
lemma dec_fetch_after_clear_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2973
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2974
     fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2975
            (start_of ly as) (Dec n e)) (2 * n + 10) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2976
    = (R, start_of ly as + 2*n + 4)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2977
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2978
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2979
apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2980
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2981
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2982
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2983
lemma dec_fetch_after_clear_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2984
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2985
    fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2986
             (start_of ly as) (Dec n e)) (2 * n + 10) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2987
    = (W0, start_of ly as + 2*n + 9)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2988
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2989
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2990
apply(subgoal_tac "2*n + 10= Suc (2*n + 9)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2991
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2992
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2993
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2994
lemma dec_fetch_on_left_move1_o:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2995
      "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2996
    fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2997
           (start_of ly as) (Dec n e)) (2 * n + 11) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2998
    = (L, start_of ly as + 2*n + 10)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2999
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3000
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3001
apply(subgoal_tac "2*n + 11= Suc (2*n + 10)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3002
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3003
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3004
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3005
lemma dec_fetch_on_left_move1_b:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3006
     "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3007
    fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3008
             (start_of ly as) (Dec n e)) (2 * n + 11) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3009
    = (L, start_of ly as + 2*n + 11)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3010
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3011
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3012
apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3013
      simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3014
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3015
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3016
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3017
lemma dec_fetch_check_left_move1_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3018
    "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3019
  fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3020
             (start_of ly as) (Dec n e)) (2 * n + 12) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3021
    = (L, start_of ly as + 2*n + 10)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3022
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3023
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3024
apply(subgoal_tac "2*n + 12= Suc (2*n + 11)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3025
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3026
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3027
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3028
lemma dec_fetch_check_left_move1_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3029
    "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3030
   fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3031
                  (start_of ly as) (Dec n e)) (2 * n + 12) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3032
    = (R, start_of ly as + 2*n + 12)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3033
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3034
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3035
apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3036
      simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3037
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3038
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3039
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3040
lemma dec_fetch_after_left_move1_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3041
  "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3042
  fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3043
                (start_of ly as) (Dec n e)) (2 * n + 13) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3044
    = (R, start_of ly as + 2*n + 16)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3045
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3046
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3047
apply(subgoal_tac "2*n + 13 = Suc (2*n + 12)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3048
      simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3049
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3050
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3051
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3052
lemma dec_fetch_on_left_move2_o:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3053
  "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3054
  fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3055
           (start_of ly as) (Dec n e)) (2 * n + 14) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3056
   = (L, start_of ly as + 2*n + 13)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3057
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3058
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3059
apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3060
      simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3061
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3062
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3063
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3064
lemma dec_fetch_on_left_move2_b:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3065
  "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3066
  fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3067
              (start_of ly as) (Dec n e)) (2 * n + 14) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3068
 = (L, start_of ly as + 2*n + 14)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3069
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3070
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3071
apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3072
      simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3073
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3074
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3075
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3076
lemma dec_fetch_check_left_move2_o:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3077
  "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3078
  fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3079
                (start_of ly as) (Dec n e)) (2 * n + 15)  Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3080
 = (L, start_of ly as + 2*n + 13)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3081
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3082
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3083
apply(subgoal_tac "2*n + 15 = Suc (2*n + 14)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3084
      simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3085
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3086
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3087
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3088
lemma dec_fetch_check_left_move2_b:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3089
  "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3090
  fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3091
                (start_of ly as) (Dec n e)) (2 * n + 15)  Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3092
 = (R, start_of ly as + 2*n + 15)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3093
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3094
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3095
apply(subgoal_tac "2*n + 15= Suc (2*n + 14)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3096
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3097
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3098
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3099
lemma dec_fetch_after_left_move2_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3100
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3101
    abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3102
    start_of ly as > 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3103
     fetch (ci (layout_of aprog) (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3104
              (Dec n e)) (2 * n + 16)  Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3105
 = (R, start_of ly e)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3106
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3107
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3108
apply(subgoal_tac "2*n + 16 = Suc (2*n + 15)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3109
      simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3110
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3111
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3112
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3113
lemma dec_fetch_next_state: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3114
    "start_of ly as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3115
    fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3116
           (start_of ly as) (Dec n e)) (2* n + 17)  b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3117
    = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3118
apply(case_tac b)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3119
apply(auto simp: ci.simps findnth.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3120
                 tshift.simps tdec_b_def add3_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3121
apply(subgoal_tac [!] "2*n + 17 = Suc (2*n + 16)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3122
      simp_all only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3123
apply(auto simp: nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3124
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3125
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3126
(*End: dec_fetch lemmas*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3127
lemmas dec_fetch_simps = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3128
 dec_fetch_locate_a_o dec_fetch_locate_a_b dec_fetch_locate_b_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3129
 dec_fetch_locate_b_b dec_fetch_locate_n_a_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3130
 dec_fetch_locate_n_a_b dec_fetch_locate_n_b_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3131
 dec_fetch_locate_n_b_b dec_fetch_first_on_right_move_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3132
 dec_fetch_first_on_right_move_b dec_fetch_first_after_clear_b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3133
 dec_fetch_first_after_clear_o dec_fetch_right_move_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3134
 dec_fetch_on_right_move_b dec_fetch_on_right_move_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3135
 dec_fetch_after_clear_b dec_fetch_after_clear_o
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3136
 dec_fetch_check_right_move_b dec_fetch_check_right_move_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3137
 dec_fetch_left_move_b dec_fetch_on_left_move1_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3138
 dec_fetch_on_left_move1_o dec_fetch_check_left_move1_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3139
 dec_fetch_check_left_move1_o dec_fetch_after_left_move1_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3140
 dec_fetch_on_left_move2_b dec_fetch_on_left_move2_o
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3141
 dec_fetch_check_left_move2_o dec_fetch_check_left_move2_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3142
 dec_fetch_after_left_move2_b dec_fetch_after_write_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3143
 dec_fetch_after_write_o dec_fetch_next_state
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3144
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3145
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3146
  "\<lbrakk>start_of ly as \<le> a; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3147
    a < start_of ly as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3148
    (a - start_of ly as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3149
    inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3150
     \<Longrightarrow> \<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3151
                inv_locate_a (as, am) (n, Bk # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3152
apply(insert locate_b_2_locate_a[of a ly as n am aaa xs], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3153
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3154
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3155
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3156
  "\<lbrakk>start_of ly as \<le> a; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3157
    a < start_of ly as + 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3158
    (a - start_of ly as) mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3159
    inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3160
   \<Longrightarrow> \<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3161
                  inv_locate_a (as, am) (n, Bk # aaa, []) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3162
apply(insert locate_b_2_locate_a_B[of a ly as n am aaa], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3163
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3164
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3165
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3166
lemma [simp]: "a\<^bsup>0\<^esup>=[]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3167
apply(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3168
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3169
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3170
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3171
lemma exp_ind: "a\<^bsup>Suc b\<^esup> =  a\<^bsup>b\<^esup> @ [a]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3172
apply(simp only: exponent_def rep_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3173
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3174
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3175
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3176
  "inv_locate_b (as, am) (n, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3177
  \<Longrightarrow> dec_first_on_right_moving n (as,  abc_lm_s am n (abc_lm_v am n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3178
                      (Suc (Suc (start_of ly as + 2 * n)), Oc # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3179
apply(simp only: inv_locate_b.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3180
     dec_first_on_right_moving.simps in_middle.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3181
     abc_lm_s.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3182
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3183
apply(erule conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3184
apply(case_tac "n < length am", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3185
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3186
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3187
apply(rule_tac x = "Suc ml" in exI, rule_tac conjI, rule_tac [1-2] impI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3188
prefer 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3189
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3190
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3191
apply(subgoal_tac "Suc n - length am = Suc (n - length am)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3192
      simp only:exponent_def rep_ind, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3193
apply(rule_tac x = "Suc ml" in exI, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3194
apply(rule_tac [!] x = "mr - 1" in exI, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3195
apply(case_tac [!] mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3196
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3197
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3198
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3199
  "\<lbrakk>inv_locate_b (as, am) (n, l, r) ires; l \<noteq> []\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3200
  \<not> inv_on_left_moving_in_middle_B (as, abc_lm_s am n (abc_lm_v am n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3201
    (s, tl l, hd l # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3202
apply(auto simp: inv_locate_b.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3203
                 inv_on_left_moving_in_middle_B.simps in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3204
apply(case_tac [!] ml, auto split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3205
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3206
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3207
lemma [simp]: "inv_locate_b (as, am) (n, l, r) ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3208
apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3209
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3210
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3211
lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (n, l, Bk # r) ires; n < length am\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3212
     \<Longrightarrow> inv_on_left_moving_norm (as, am) (s, tl l, hd l # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3213
apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3214
                 in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3215
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3216
apply(erule_tac conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3217
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3218
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3219
apply(rule_tac x = "ml - 1" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3220
apply(rule_tac [!] x = "Suc mr" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3221
apply(case_tac [!] mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3222
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3223
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3224
lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (n, l, Bk # r) ires; \<not> n < length am\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3225
    \<Longrightarrow> inv_on_left_moving_norm (as, am @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3226
        replicate (n - length am) 0 @ [0]) (s, tl l, hd l # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3227
apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3228
                 in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3229
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3230
apply(erule_tac conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3231
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3232
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3233
apply(subgoal_tac "Suc n - length am = Suc (n - length am)", simp only: rep_ind exponent_def, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3234
apply(rule_tac x = "Suc mr" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3235
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3236
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3237
lemma inv_locate_b_2_on_left_moving[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3238
  "\<lbrakk>inv_locate_b (as, am) (n, l, Bk # r) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3239
   \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as,  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3240
            abc_lm_s am n (abc_lm_v am n)) (s, [], Bk # Bk # r) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3241
       (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as,  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3242
            abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3243
apply(subgoal_tac "l\<noteq>[]")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3244
apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3245
      (as,  abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3246
apply(simp add:inv_on_left_moving.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3247
          abc_lm_s.simps abc_lm_v.simps split: if_splits, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3248
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3249
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3250
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3251
  "inv_locate_b (as, am) (n, l, []) ires \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3252
                   inv_locate_b (as, am) (n, l, [Bk]) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3253
apply(auto simp: inv_locate_b.simps in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3254
apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3255
      rule_tac x = "Suc (length lm1) - length am" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3256
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3257
apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3258
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3259
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3260
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3261
lemma nil_2_nil: "<lm::nat list> = [] \<Longrightarrow> lm = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3262
apply(auto simp: tape_of_nl_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3263
apply(case_tac lm, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3264
apply(case_tac list, auto simp: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3265
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3266
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3267
lemma  inv_locate_b_2_on_left_moving_b[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3268
   "inv_locate_b (as, am) (n, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3269
     \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3270
                  abc_lm_s am n (abc_lm_v am n)) (s, [], [Bk]) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3271
         (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, abc_lm_s am n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3272
                  (abc_lm_v am n)) (s, tl l, [hd l]) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3273
apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3274
apply(simp only: inv_on_left_moving.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3275
apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3276
         (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3277
apply(simp only: inv_on_left_moving_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3278
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3279
apply(erule_tac conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3280
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3281
      rule_tac x = m in exI, rule_tac x = ml in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3282
      rule_tac x = mr in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3283
apply(case_tac mr, simp, simp, case_tac nat, auto intro: nil_2_nil)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3284
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3285
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3286
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3287
 "\<lbrakk>dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3288
   \<Longrightarrow> dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3289
apply(simp only: dec_first_on_right_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3290
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3291
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3292
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3293
apply(rule_tac x = "Suc ml" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3294
      rule_tac x = "mr - 1" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3295
apply(case_tac [!] mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3296
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3297
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3298
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3299
  "dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3300
apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3301
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3302
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3303
lemma [elim]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3304
  "\<lbrakk>\<not> length lm1 < length am; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3305
    am @ replicate (length lm1 - length am) 0 @ [0::nat] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3306
                                                lm1 @ m # lm2;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3307
    0 < m\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3308
   \<Longrightarrow> RR"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3309
apply(subgoal_tac "lm2 = []", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3310
apply(drule_tac length_equal, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3311
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3312
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3313
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3314
 "\<lbrakk>dec_first_on_right_moving n (as, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3315
                   abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3316
\<Longrightarrow> dec_after_clear (as, abc_lm_s am n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3317
                 (abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3318
apply(simp only: dec_first_on_right_moving.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3319
                 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3320
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3321
apply(case_tac "n < length am")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3322
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3323
      rule_tac x = "m - 1" in exI, auto simp: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3324
apply(case_tac [!] mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3325
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3326
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3327
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3328
 "\<lbrakk>dec_first_on_right_moving n (as, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3329
                   abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3330
\<Longrightarrow> (l = [] \<longrightarrow> dec_after_clear (as, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3331
             abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3332
    (l \<noteq> [] \<longrightarrow> dec_after_clear (as, abc_lm_s am n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3333
                      (abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3334
apply(subgoal_tac "l \<noteq> []", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3335
      simp only: dec_first_on_right_moving.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3336
                 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3337
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3338
apply(case_tac "n < length am", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3339
apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3340
apply(case_tac [1-2] mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3341
apply(case_tac [1-2] m, auto simp: dec_first_on_right_moving.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3342
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3343
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3344
lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Oc # r) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3345
                \<Longrightarrow> dec_after_clear (as, am) (s', l, Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3346
apply(auto simp: dec_after_clear.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3347
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3348
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3349
lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Bk # r) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3350
                \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3351
apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3352
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3353
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3354
lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3355
             \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, []) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3356
apply(auto simp: dec_after_clear.simps dec_right_move.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3357
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3358
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3359
lemma [simp]: "\<exists>rn. a::block\<^bsup>rn\<^esup> = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3360
apply(rule_tac x = 0 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3361
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3362
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3363
lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3364
             \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3365
apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3366
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3367
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3368
lemma [simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3369
apply(auto simp: dec_right_move.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3370
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3371
              
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3372
lemma dec_right_move_2_check_right_move[simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3373
     "\<lbrakk>dec_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3374
      \<Longrightarrow> dec_check_right_move (as, am) (s', Bk # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3375
apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3376
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3377
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3378
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3379
 "dec_right_move (as, am) (s, l, []) ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3380
  dec_right_move (as, am) (s, l, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3381
apply(simp add: dec_right_move.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3382
apply(rule_tac iffI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3383
apply(erule_tac [!] exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3384
apply(erule_tac [2] exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3385
apply(rule_tac [!] x = lm1 in exI, rule_tac x = "[]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3386
      rule_tac [!] x = m in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3387
apply(auto intro: nil_2_nil)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3388
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3389
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3390
lemma [simp]: "\<lbrakk>dec_right_move (as, am) (s, l, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3391
             \<Longrightarrow> dec_check_right_move (as, am) (s, Bk # l, []) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3392
apply(insert dec_right_move_2_check_right_move[of as am s l "[]" s'], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3393
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3394
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3395
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3396
lemma [simp]: "dec_check_right_move (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3397
apply(auto simp: dec_check_right_move.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3398
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3399
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3400
lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Oc # r) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3401
             \<Longrightarrow> dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3402
apply(auto simp: dec_check_right_move.simps dec_after_write.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3403
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3404
      rule_tac x = m in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3405
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3406
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3407
lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3408
                \<Longrightarrow> dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3409
apply(auto simp: dec_check_right_move.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3410
                 dec_left_move.simps inv_after_move.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3411
apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3412
apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3413
apply(rule_tac x = "Suc rn" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3414
apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3415
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3416
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3417
lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3418
             \<Longrightarrow> dec_left_move (as, am) (s', tl l, [hd l]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3419
apply(auto simp: dec_check_right_move.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3420
                 dec_left_move.simps inv_after_move.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3421
apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3422
apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3423
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3424
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3425
lemma [simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3426
apply(auto simp: dec_left_move.simps inv_after_move.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3427
apply(case_tac [!] rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3428
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3429
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3430
lemma [simp]: "dec_left_move (as, am) (s, l, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3431
             \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3432
apply(auto simp: dec_left_move.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3433
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3434
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3435
lemma tape_of_nl_abv_cons_ex[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3436
   "\<exists>lna. Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup> = <m # rev lm1> @ Bk\<^bsup>lna\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3437
apply(case_tac "lm1=[]", auto simp: tape_of_nl_abv 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3438
                                    tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3439
apply(rule_tac x = "ln" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3440
apply(simp add:  tape_of_nat_list_cons exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3441
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3442
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3443
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3444
lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3445
                 (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup>, Bk # Bk\<^bsup>rn\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3446
apply(simp only: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3447
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3448
done    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3449
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3450
lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3451
  (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3452
apply(simp add: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3453
apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3454
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3455
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3456
lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3457
  (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3458
apply(simp add: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3459
apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3460
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3461
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3462
lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3463
  inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3464
  Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3465
apply(simp only: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3466
apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3467
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3468
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3469
lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3470
  inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3471
  Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk # Bk # ires, [Bk]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3472
apply(simp only: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3473
apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3474
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3475
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3476
lemma [simp]: "dec_left_move (as, am) (s, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3477
       \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3478
apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3479
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3480
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3481
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3482
lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3483
                        (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup>, [Bk])  ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3484
apply(auto simp: inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3485
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3486
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3487
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3488
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3489
lemma [simp]: "dec_left_move (as, am) (s, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3490
             \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, [hd l]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3491
apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3492
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3493
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3494
lemma [simp]: "dec_after_write (as, am) (s, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3495
       \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3496
apply(auto simp: dec_after_write.simps dec_on_right_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3497
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3498
      rule_tac x = "hd lm2" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3499
apply(rule_tac x = "Suc 0" in exI,rule_tac x =  "Suc (hd lm2)" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3500
apply(case_tac lm2, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3501
apply(case_tac "list = []", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3502
      auto simp: tape_of_nl_abv tape_of_nat_list.simps split: if_splits )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3503
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3504
apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3505
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3506
apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3507
apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3508
apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3509
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3510
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3511
lemma [simp]: "dec_after_write (as, am) (s, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3512
       \<Longrightarrow> dec_after_write (as, am) (s', l, Oc # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3513
apply(auto simp: dec_after_write.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3514
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3515
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3516
lemma [simp]: "dec_after_write (as, am) (s, aaa, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3517
             \<Longrightarrow> dec_after_write (as, am) (s', aaa, [Oc]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3518
apply(auto simp: dec_after_write.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3519
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3520
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3521
lemma [simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3522
       \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3523
apply(simp only: dec_on_right_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3524
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3525
apply(erule conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3526
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3527
      rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3528
      rule_tac x = "mr - 1" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3529
apply(case_tac mr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3530
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3531
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3532
lemma [simp]: "dec_on_right_moving (as, am) (s, l, r) ires\<Longrightarrow>  l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3533
apply(auto simp: dec_on_right_moving.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3534
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3535
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3536
lemma [simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3537
      \<Longrightarrow>  dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3538
apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3539
apply(case_tac [!] mr, auto split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3540
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3541
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3542
lemma [simp]: "dec_on_right_moving (as, am) (s, l, []) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3543
             \<Longrightarrow> dec_after_clear (as, am) (s', tl l, [hd l]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3544
apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3545
apply(case_tac mr, simp_all split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3546
apply(rule_tac x = lm1 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3547
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3548
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3549
lemma start_of_le: "a < b \<Longrightarrow> start_of ly a \<le> start_of ly b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3550
proof(induct b arbitrary: a, simp, case_tac "a = b", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3551
  fix b a
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3552
  show "start_of ly b \<le> start_of ly (Suc b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3553
    apply(case_tac "b::nat", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3554
          simp add: start_of.simps, simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3555
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3556
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3557
  fix b a
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3558
  assume h1: "\<And>a. a < b \<Longrightarrow> start_of ly a \<le> start_of ly b" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3559
             "a < Suc b" "a \<noteq> b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3560
  hence "a < b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3561
    by(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3562
  from h1 and this have h2: "start_of ly a \<le> start_of ly b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3563
    by(drule_tac h1, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3564
  from h2 show "start_of ly a \<le> start_of ly (Suc b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3565
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3566
    have "start_of ly b \<le> start_of ly (Suc b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3567
      apply(case_tac "b::nat", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3568
            simp add: start_of.simps, simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3569
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3570
    from h2 and this show "start_of ly a \<le> start_of ly (Suc b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3571
      by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3572
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3573
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3574
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3575
lemma start_of_dec_length[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3576
  "\<lbrakk>abc_fetch a aprog = Some (Dec n e)\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3577
    start_of (layout_of aprog) (Suc a)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3578
          = start_of (layout_of aprog) a + 2*n + 16"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3579
apply(case_tac a, auto simp: abc_fetch.simps start_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3580
                             layout_of.simps length_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3581
                       split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3582
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3583
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3584
lemma start_of_ge: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3585
 "\<lbrakk>abc_fetch a aprog = Some (Dec n e); a < e\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3586
  start_of (layout_of aprog) e > 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3587
              start_of (layout_of aprog) a + 2*n + 15"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3588
apply(case_tac "e = Suc a", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3589
      simp add: start_of.simps abc_fetch.simps layout_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3590
                length_of.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3591
apply(subgoal_tac "Suc a < e", drule_tac a = "Suc a" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3592
             and ly = "layout_of aprog" in start_of_le)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3593
apply(subgoal_tac "start_of (layout_of aprog) (Suc a)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3594
         = start_of (layout_of aprog) a + 2*n + 16", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3595
apply(rule_tac start_of_dec_length, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3596
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3597
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3598
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3599
lemma starte_not_equal[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3600
 "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3601
   \<Longrightarrow> (start_of ly e \<noteq> Suc (Suc (start_of ly as + 2 * n)) \<and>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3602
        start_of ly e \<noteq> start_of ly as + 2 * n + 3 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3603
        start_of ly e \<noteq> start_of ly as + 2 * n + 4 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3604
        start_of ly e \<noteq> start_of ly as + 2 * n + 5 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3605
        start_of ly e \<noteq> start_of ly as + 2 * n + 6 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3606
        start_of ly e \<noteq> start_of ly as + 2 * n + 7 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3607
        start_of ly e \<noteq> start_of ly as + 2 * n + 8 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3608
        start_of ly e \<noteq> start_of ly as + 2 * n + 9 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3609
        start_of ly e \<noteq> start_of ly as + 2 * n + 10 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3610
        start_of ly e \<noteq> start_of ly as + 2 * n + 11 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3611
        start_of ly e \<noteq> start_of ly as + 2 * n + 12 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3612
        start_of ly e \<noteq> start_of ly as + 2 * n + 13 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3613
        start_of ly e \<noteq> start_of ly as + 2 * n + 14 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3614
        start_of ly e \<noteq> start_of ly as + 2 * n + 15)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3615
apply(case_tac "e = as", simp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3616
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3617
apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3618
apply(drule_tac a = as and e = e in start_of_ge, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3619
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3620
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3621
lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3622
      \<Longrightarrow> (Suc (Suc (start_of ly as + 2 * n)) \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3623
          start_of ly as + 2 * n + 3 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3624
          start_of ly as + 2 * n + 4 \<noteq> start_of ly e \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3625
          start_of ly as + 2 * n + 5 \<noteq>start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3626
          start_of ly as + 2 * n + 6 \<noteq> start_of ly e \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3627
          start_of ly as + 2 * n + 7 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3628
          start_of ly as + 2 * n + 8 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3629
          start_of ly as + 2 * n + 9 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3630
          start_of ly as + 2 * n + 10 \<noteq> start_of ly e \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3631
          start_of ly as + 2 * n + 11 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3632
          start_of ly as + 2 * n + 12 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3633
          start_of ly as + 2 * n + 13 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3634
          start_of ly as + 2 * n + 14 \<noteq> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3635
          start_of ly as + 2 * n + 15 \<noteq> start_of ly e)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3636
apply(insert starte_not_equal[of as aprog n e ly], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3637
                            simp del: starte_not_equal)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3638
apply(erule_tac conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3639
apply(rule_tac conjI, simp del: starte_not_equal)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3640
apply(rule not_sym, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3641
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3642
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3643
lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3644
  fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3645
                       (Dec n as)) (Suc 0) Oc =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3646
 (R, Suc (start_of (layout_of aprog) as))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3647
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3648
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3649
                 nth_of.simps tshift.simps nth_append 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3650
                 Suc_pre tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3651
apply(insert findnth_nth[of 0 n "Suc 0"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3652
apply(simp add: findnth.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3653
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3654
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3655
lemma start_of_inj[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3656
  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); e \<noteq> as; ly = layout_of aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3657
   \<Longrightarrow> start_of ly as \<noteq> start_of ly e"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3658
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3659
apply(case_tac "as", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3660
apply(case_tac "e = nat", simp add: start_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3661
                                    layout_of.simps length_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3662
apply(subgoal_tac "e < length aprog", simp add: length_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3663
                                         split: abc_inst.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3664
apply(simp add: abc_fetch.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3665
apply(subgoal_tac "e < nat", drule_tac a = e and b = nat 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3666
                                   and ly =ly in start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3667
apply(subgoal_tac "start_of ly nat < start_of ly (Suc nat)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3668
          simp, simp add: start_of.simps layout_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3669
apply(subgoal_tac "nat < length aprog", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3670
apply(case_tac "aprog ! nat", auto simp: length_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3671
apply(simp add: abc_fetch.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3672
apply(subgoal_tac "e > as", drule_tac start_of_ge, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3673
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3674
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3675
lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); e < as\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3676
    \<Longrightarrow> Suc (start_of (layout_of aprog) e) - 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3677
                               start_of (layout_of aprog) as = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3678
apply(frule_tac ly = "layout_of aprog" in start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3679
apply(subgoal_tac "start_of (layout_of aprog) as \<noteq> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3680
                            start_of (layout_of aprog) e", arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3681
apply(rule start_of_inj, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3682
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3683
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3684
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3685
   "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3686
     0 < start_of (layout_of aprog) as\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3687
 \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3688
     (Dec n e)) (Suc (start_of (layout_of aprog) e) - 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3689
                 start_of (layout_of aprog) as) Oc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3690
    = (if e = as then (R, start_of (layout_of aprog) as + 1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3691
                 else (Nop, 0))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3692
apply(auto split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3693
apply(case_tac "e < as", simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3694
apply(subgoal_tac " e > as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3695
apply(drule start_of_ge, simp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3696
      auto simp: fetch.simps ci_length nth_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3697
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3698
 "length (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3699
                        (Dec n e)) div 2= length_of (Dec n e)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3700
defer
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3701
apply(simp add: ci_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3702
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3703
 "length (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3704
                  (Dec n e)) mod 2 = 0", auto simp: length_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3705
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3706
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3707
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3708
    "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3709
 fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3710
                                          (Dec n as)) (Suc 0)  Bk 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3711
      = (W1, start_of (layout_of aprog) as)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3712
apply(auto simp: ci.simps findnth.simps fetch.simps nth_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3713
                 tshift.simps nth_append Suc_pre tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3714
apply(insert findnth_nth[of 0 n "0"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3715
apply(simp add: findnth.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3716
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3717
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3718
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3719
 "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3720
   0 < start_of (layout_of aprog) as\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3721
\<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3722
         (Dec n e)) (Suc (start_of (layout_of aprog) e) - 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3723
              start_of (layout_of aprog) as)  Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3724
   = (if e = as then (W1, start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3725
                  else (Nop, 0))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3726
apply(auto split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3727
apply(case_tac "e < as", simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3728
apply(subgoal_tac " e > as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3729
apply(drule start_of_ge, simp, auto simp: fetch.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3730
                                          ci_length nth_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3731
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3732
 "length (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3733
                            (Dec n e)) div 2= length_of (Dec n e)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3734
defer
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3735
apply(simp add: ci_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3736
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3737
 "length (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3738
                   (Dec n e)) mod 2 = 0", auto simp: length_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3739
apply(simp add: ci.simps tshift.simps tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3740
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3741
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3742
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3743
 "inv_stop (as, abc_lm_s am n (abc_lm_v am n)) (s, l, r) ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3744
apply(auto simp: inv_stop.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3745
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3746
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3747
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3748
 "\<lbrakk>abc_fetch as aprog = Some (Dec n e); e \<noteq> as; ly = layout_of aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3749
  \<Longrightarrow> (\<not> (start_of ly as \<le> start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3750
      start_of ly e < start_of ly as + 2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3751
    \<and> start_of ly e \<noteq> start_of ly as + 2*n \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3752
      start_of ly e \<noteq> Suc (start_of ly as + 2*n) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3753
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3754
apply(drule_tac ly = ly in start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3755
apply(case_tac n, simp, drule start_of_inj, simp, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3756
apply(drule_tac start_of_ge, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3757
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3758
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3759
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3760
   "\<lbrakk>abc_fetch as aprog = Some (Dec n e); start_of ly as \<le> s; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3761
     s < start_of ly as + 2 * n; ly = layout_of aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3762
     \<Longrightarrow> Suc s \<noteq> start_of ly e "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3763
apply(case_tac "e = as", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3764
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3765
apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3766
apply(drule_tac start_of_ge, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3767
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3768
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3769
lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3770
                ly = layout_of aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3771
         \<Longrightarrow> Suc (start_of ly as + 2 * n) \<noteq> start_of ly e"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3772
apply(case_tac "e = as", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3773
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3774
apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3775
apply(drule_tac start_of_ge, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3776
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3777
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3778
lemma dec_false_1[simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3779
 "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3780
  \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3781
apply(auto simp: inv_locate_b.simps in_middle.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3782
apply(case_tac "length lm1 \<ge> length am", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3783
apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3784
apply(case_tac mr, auto simp: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3785
apply(subgoal_tac "Suc (length lm1) - length am = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3786
                   Suc (length lm1 - length am)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3787
      simp add: rep_ind del: replicate.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3788
apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3789
                and ys = "lm1 @ m # lm2" in length_equal, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3790
apply(case_tac mr, auto simp: abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3791
apply(case_tac "mr = 0", simp_all add:  exponent_def split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3792
apply(subgoal_tac "Suc (length lm1) - length am = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3793
                       Suc (length lm1 - length am)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3794
      simp add: rep_ind del: replicate.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3795
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3796
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3797
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3798
 "\<lbrakk>inv_locate_b (as, am) (n, aaa, Bk # xs) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3799
   abc_lm_v am n = 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3800
   \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3801
                         (s, tl aaa, hd aaa # Bk # xs) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3802
apply(insert inv_locate_b_2_on_left_moving[of as am n aaa xs ires s], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3803
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3804
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3805
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3806
 "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3807
   \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3808
apply(insert inv_locate_b_2_on_left_moving_b[of as am n aaa ires s], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3809
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3810
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3811
lemma [simp]: "\<lbrakk>am ! n = (0::nat); n < length am\<rbrakk> \<Longrightarrow> am[n := 0] = am"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3812
apply(simp add: list_update_same_conv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3813
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3814
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3815
lemma [simp]: "\<lbrakk>abc_lm_v am n = 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3816
                inv_locate_b (as, abc_lm_s am n 0) (n, Oc # aaa, xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3817
     \<Longrightarrow> inv_locate_b (as, am) (n, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3818
apply(simp only: inv_locate_b.simps in_middle.simps abc_lm_s.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3819
                 abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3820
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3821
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3822
apply(case_tac "n < length am", simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3823
apply(erule_tac conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3824
apply(rule_tac x = tn in exI, rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3825
apply(rule_tac x = ml in exI, rule_tac x = mr in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3826
defer
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3827
apply(rule_tac x = "Suc n - length am" in exI, rule_tac x = m in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3828
apply(subgoal_tac "Suc n - length am = Suc (n - length am)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3829
apply(simp add: exponent_def rep_ind del: replicate.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3830
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3831
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3832
lemma  [intro]: "\<lbrakk>abc_lm_v (a # list) 0 = 0\<rbrakk> \<Longrightarrow> a = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3833
apply(simp add: abc_lm_v.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3834
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3835
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3836
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3837
 "inv_stop (as, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3838
          (start_of (layout_of aprog) e, aaa, Oc # xs) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3839
  \<Longrightarrow> inv_locate_a (as, abc_lm_s am n 0) (0, aaa, Oc # xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3840
apply(simp add: inv_locate_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3841
apply(rule disjI1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3842
apply(auto simp: inv_stop.simps at_begin_norm.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3843
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3844
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3845
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3846
 "\<lbrakk>abc_lm_v am 0 = 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3847
  inv_stop (as, abc_lm_s am 0 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3848
      (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3849
  inv_locate_b (as, am) (0, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3850
apply(auto simp: inv_stop.simps inv_locate_b.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3851
                 in_middle.simps abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3852
apply(case_tac "am = []", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3853
apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3854
      rule_tac x = 0 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3855
apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3856
  simp add: tape_of_nl_abv tape_of_nat_list.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3857
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3858
apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3859
      rule_tac x = "hd am" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3860
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "hd am" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3861
apply(case_tac am, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3862
apply(subgoal_tac "a = 0", case_tac list, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3863
      auto simp: tape_of_nat_list.simps tape_of_nl_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3864
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3865
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3866
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3867
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3868
 "\<lbrakk>inv_stop (as, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3869
          (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3870
  \<Longrightarrow> inv_locate_b (as, am) (0, Oc # aaa, xs) ires \<or> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3871
      inv_locate_b (as, abc_lm_s am n 0) (0, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3872
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3873
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3874
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3875
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3876
"\<lbrakk>abc_lm_v am n = 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3877
  inv_stop (as, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3878
          (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3879
 \<Longrightarrow> \<not> Suc 0 < 2 * n \<longrightarrow> e = as \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3880
            inv_locate_b (as, am) (n, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3881
apply(case_tac n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3882
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3883
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3884
lemma dec_false2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3885
 "inv_stop (as, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3886
  (start_of (layout_of aprog) e, aaa, Bk # xs) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3887
apply(auto simp: inv_stop.simps abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3888
apply(case_tac "am", simp, case_tac n, simp add: tape_of_nl_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3889
apply(case_tac list, simp add: tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3890
apply(simp add: tape_of_nat_list.simps , simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3891
apply(case_tac "list[nat := 0]", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3892
      simp add: tape_of_nat_list.simps  tape_of_nl_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3893
apply(simp add: tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3894
apply(case_tac "am @ replicate (n - length am) 0 @ [0]", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3895
apply(case_tac list, auto simp: tape_of_nl_abv 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3896
                                tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3897
done	
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3898
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3899
lemma dec_false3:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3900
   "inv_stop (as, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3901
              (start_of (layout_of aprog) e, aaa, []) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3902
apply(auto simp: inv_stop.simps abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3903
apply(case_tac "am", case_tac n, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3904
apply(case_tac n, auto simp: tape_of_nl_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3905
apply(case_tac "list::nat list",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3906
            simp add: tape_of_nat_list.simps tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3907
apply(simp add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3908
apply(case_tac "list[nat := 0]", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3909
            simp add: tape_of_nat_list.simps tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3910
apply(simp add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3911
apply(case_tac "(am @ replicate (n - length am) 0 @ [0])", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3912
apply(case_tac list, auto simp: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3913
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3914
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3915
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3916
  "fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3917
       (start_of (layout_of aprog) as) (Dec n e)) 0 b = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3918
by(simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3919
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3920
declare dec_inv_1.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3921
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3922
declare inv_locate_n_b.simps [simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3923
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3924
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3925
"\<lbrakk>0 < abc_lm_v am n; 0 < n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3926
  at_begin_norm (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3927
  \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3928
apply(simp only: at_begin_norm.simps inv_locate_n_b.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3929
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3930
apply(rule_tac x = lm1 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3931
apply(case_tac "length lm2", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3932
apply(case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3933
apply(rule_tac x = "tl lm2" in exI, rule_tac x = "hd lm2" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3934
apply(rule conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3935
apply(case_tac "lm2", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3936
apply(case_tac "lm2", auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3937
apply(case_tac [!] "list", auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3938
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3939
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3940
lemma [simp]: "(\<exists>rn. Oc # xs = Bk\<^bsup>rn\<^esup>) = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3941
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3942
apply(case_tac rn, auto simp: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3943
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3944
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3945
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3946
  "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3947
    at_begin_fst_bwtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3948
 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3949
apply(simp add: at_begin_fst_bwtn.simps inv_locate_n_b.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3950
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3951
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3952
lemma Suc_minus:"length am + tn = n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3953
       \<Longrightarrow> Suc tn = Suc n - length am "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3954
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3955
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3956
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3957
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3958
 "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3959
   at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3960
 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3961
apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3962
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3963
apply(erule conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3964
apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3965
      rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3966
apply(simp add: exponent_def rep_ind del: replicate.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3967
apply(rule conjI)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3968
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3969
apply(case_tac [!] rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3970
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3971
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3972
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3973
 "\<lbrakk>0 < abc_lm_v am n; 0 < n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3974
 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3975
apply(auto simp: inv_locate_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3976
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3977
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3978
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3979
 "\<lbrakk>inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3980
 \<Longrightarrow> dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n))  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3981
                                      (s, Oc # aaa, xs) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3982
apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3983
                 abc_lm_s.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3984
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3985
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3986
apply(rule_tac x = "Suc (Suc 0)" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3987
      rule_tac x = "m - 1" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3988
apply(case_tac m, auto simp:  exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3989
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3990
      rule_tac x = m in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3991
      simp add: Suc_diff_le rep_ind del: replicate.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3992
apply(rule_tac x = "Suc (Suc 0)" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3993
      rule_tac x = "m - 1" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3994
apply(case_tac m, auto simp:  exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3995
apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3996
      rule_tac x = m in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3997
apply(rule_tac x = "Suc (Suc 0)" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3998
      rule_tac x = "m - 1" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3999
apply(case_tac m, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4000
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4001
      rule_tac x = m in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4002
      simp add: Suc_diff_le rep_ind del: replicate.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4003
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4004
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4005
lemma dec_false_2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4006
 "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4007
 \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4008
apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4009
apply(case_tac [!] m, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4010
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4011
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4012
lemma dec_false_2_b:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4013
 "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4014
                                (n, aaa, []) ires\<rbrakk> \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4015
apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4016
apply(case_tac [!] m, auto simp: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4017
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4018
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4019
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4020
(*begin: dec halt1 lemmas*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4021
thm abc_inc_stage1.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4022
fun abc_dec_1_stage1:: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4023
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4024
  "abc_dec_1_stage1 (s, l, r) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4025
       (if s > ss \<and> s \<le> ss + 2*n + 1 then 4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4026
        else if s = ss + 2 * n + 13 \<or> s = ss + 2*n + 14 then 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4027
        else if s = ss + 2*n + 15 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4028
        else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4029
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4030
fun abc_dec_1_stage2:: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4031
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4032
  "abc_dec_1_stage2 (s, l, r) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4033
       (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4034
        else if s = ss + 2*n + 13 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4035
        else if s = ss + 2*n + 14 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4036
        else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4037
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4038
fun abc_dec_1_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4039
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4040
  "abc_dec_1_stage3 (s, l, r) ss n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4041
        (if s \<le> ss + 2*n + 1 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4042
             if (s - ss) mod 2 = 0 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4043
                         if r \<noteq> [] \<and> hd r = Oc then 0 else 1  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4044
                         else length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4045
         else if s = ss + 2 * n + 13 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4046
             if l = Bk # ires \<and> r \<noteq> [] \<and> hd r = Oc then 2 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4047
             else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4048
         else if s = ss + 2 * n + 14 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4049
             if r \<noteq> [] \<and> hd r = Oc then 3 else 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4050
         else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4051
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4052
fun abc_dec_1_measure :: "(t_conf \<times> nat \<times> nat \<times> block list) \<Rightarrow> (nat \<times> nat \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4053
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4054
  "abc_dec_1_measure (c, ss, n, ires) = (abc_dec_1_stage1 c ss n, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4055
                   abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4056
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4057
definition abc_dec_1_LE ::
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4058
  "(((nat \<times> block list \<times> block list) \<times> nat \<times>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4059
  nat \<times> block list) \<times> ((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4060
  where "abc_dec_1_LE \<equiv> (inv_image lex_triple abc_dec_1_measure)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4061
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4062
lemma wf_dec_le: "wf abc_dec_1_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4063
by(auto intro:wf_inv_image wf_lex_triple simp:abc_dec_1_LE_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4064
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4065
declare dec_inv_1.simps[simp del] dec_inv_2.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4066
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4067
lemma [elim]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4068
 "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4069
   start_of (layout_of aprog) as < start_of (layout_of aprog) e;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4070
   start_of (layout_of aprog) e \<le> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4071
         Suc (start_of (layout_of aprog) as + 2 * n)\<rbrakk> \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4072
apply(case_tac "e = as", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4073
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4074
apply(drule_tac a = e and b = as and ly = "layout_of aprog" in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4075
                                                 start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4076
apply(drule_tac start_of_ge, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4077
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4078
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4079
lemma [elim]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4080
                                start_of (layout_of aprog) e 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4081
    = start_of (layout_of aprog) as + 2 * n + 13\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4082
         \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4083
apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4084
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4085
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4086
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4087
lemma [elim]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4088
                 start_of (layout_of aprog) e = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4089
               start_of (layout_of aprog) as + 2 * n + 14\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4090
        \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4091
apply(insert starte_not_equal[of as aprog n e "layout_of aprog"],
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4092
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4093
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4094
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4095
lemma [elim]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4096
 "\<lbrakk>abc_fetch as aprog = Some (Dec n e);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4097
   start_of (layout_of aprog) as < start_of (layout_of aprog) e;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4098
   start_of (layout_of aprog) e \<le> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4099
              Suc (start_of (layout_of aprog) as + 2 * n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4100
   \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4101
apply(case_tac "e = as", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4102
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4103
apply(drule_tac a = e and b = as and ly = "layout_of aprog" in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4104
                                                    start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4105
apply(drule_tac start_of_ge, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4106
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4107
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4108
lemma [elim]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4109
 "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4110
   start_of (layout_of aprog) e = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4111
               start_of (layout_of aprog) as + 2 * n + 13\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4112
    \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4113
apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4114
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4115
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4116
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4117
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4118
 "abc_fetch as aprog = Some (Dec n e) \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4119
   Suc (start_of (layout_of aprog) as) \<noteq> start_of (layout_of aprog) e"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4120
apply(case_tac "e = as", simp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4121
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4122
apply(drule_tac a = e and b = as and ly = "(layout_of aprog)" in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4123
                                                 start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4124
apply(drule_tac a = as and e = e in start_of_ge, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4125
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4126
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4127
lemma [simp]: "inv_on_left_moving (as, am) (s, [], r) ires 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4128
  = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4129
apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4130
                inv_on_left_moving_in_middle_B.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4131
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4132
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4133
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4134
  "inv_check_left_moving (as, abc_lm_s am n 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4135
  (start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4136
 = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4137
apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4138
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4139
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4140
lemma dec_inv_stop1_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4141
    "\<lbrakk>abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4142
      start_of (layout_of aprog) as > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4143
 \<Longrightarrow> \<forall>na. \<not> (\<lambda>(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) e)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4144
            (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4145
              (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4146
                 (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4147
                      (start_of (layout_of aprog) as, n, ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4148
           dec_inv_1 (layout_of aprog) n e (as, am)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4149
            (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4150
              (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4151
                (Dec n e), start_of (layout_of aprog) as - Suc 0) na) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4152
       \<longrightarrow> dec_inv_1 (layout_of aprog) n e (as, am)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4153
            (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4154
              (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4155
                 (Dec n e), start_of (layout_of aprog) as - Suc 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4156
                    (Suc na)) ires \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4157
            ((t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4158
            (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4159
           (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4160
             start_of (layout_of aprog) as, n, ires),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4161
         t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4162
            (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4163
               (Dec n e), start_of (layout_of aprog) as - Suc 0) na,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4164
            start_of (layout_of aprog) as, n, ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4165
           \<in> abc_dec_1_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4166
apply(rule allI, rule impI, simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4167
apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4168
(ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4169
start_of (layout_of aprog) as - Suc 0) na)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4170
apply(auto split:if_splits simp add:t_step.simps dec_inv_1.simps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4171
          tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4172
apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_1.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4173
apply(auto simp add: abc_dec_1_LE_def lex_square_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4174
                     lex_triple_def lex_pair_def  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4175
                split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4176
apply(rule dec_false_1, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4177
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4178
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4179
lemma dec_inv_stop1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4180
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4181
    dec_inv_1 ly n e (as, am) (start_of ly as + 1, l, r) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4182
    abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4183
  (\<exists> stp. (\<lambda> (s', l', r'). s' = start_of ly e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4184
           dec_inv_1 ly n e (as, am) (s', l' , r') ires) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4185
  (t_steps (start_of ly as + 1, l, r)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4186
     (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4187
apply(insert halt_lemma2[of abc_dec_1_LE 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4188
    "\<lambda> ((s, l, r), ss, n', ires'). s = start_of ly e" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4189
     "(\<lambda> stp. (t_steps (start_of ly as + 1, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4190
          (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4191
               stp, start_of ly as, n, ires))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4192
     "\<lambda> ((s, l, r), ss, n, ires'). dec_inv_1 ly n e (as, am) (s, l, r) ires'"],
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4193
     simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4194
apply(insert wf_dec_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4195
apply(insert dec_inv_stop1_pre[of as aprog n e am l r], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4196
apply(subgoal_tac "start_of (layout_of aprog) as > 0", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4197
                                      simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4198
apply(erule_tac exE, rule_tac x = na in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4199
apply(case_tac
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4200
     "(t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4201
         (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4202
           (Dec n e), start_of (layout_of aprog) as - Suc 0) na)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4203
      case_tac b, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4204
apply(rule startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4205
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4206
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4207
(*begin: dec halt2 lemmas*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4208
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4209
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4210
  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4211
    ly = layout_of aprog\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4212
              start_of ly (Suc as) = start_of ly as + 2*n + 16"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4213
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4214
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4215
fun abc_dec_2_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4216
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4217
  "abc_dec_2_stage1 (s, l, r) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4218
              (if s \<le> ss + 2*n + 1 then 7
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4219
               else if s = ss + 2*n + 2 then 6 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4220
               else if s = ss + 2*n + 3 then 5
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4221
               else if s \<ge> ss + 2*n + 4 \<and> s \<le> ss + 2*n + 9 then 4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4222
               else if s = ss + 2*n + 6 then 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4223
               else if s = ss + 2*n + 10 \<or> s = ss + 2*n + 11 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4224
               else if s = ss + 2*n + 12 then 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4225
               else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4226
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4227
thm new_tape.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4228
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4229
fun abc_dec_2_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4230
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4231
  "abc_dec_2_stage2 (s, l, r) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4232
       (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4233
        else if s = ss + 2*n + 10 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4234
        else if s = ss + 2*n + 11 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4235
        else if s = ss + 2*n + 4 then length r - 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4236
        else if s = ss + 2*n + 5 then length r 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4237
        else if s = ss + 2*n + 7 then length r - 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4238
        else if s = ss + 2*n + 8 then  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4239
              length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4240
        else if s = ss + 2*n + 9 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4241
              length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4242
        else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4243
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4244
fun abc_dec_2_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4245
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4246
  "abc_dec_2_stage3 (s, l, r) ss n ires =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4247
        (if s \<le> ss + 2*n + 1 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4248
            if (s - ss) mod 2 = 0 then if r \<noteq> [] \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4249
                                          hd r = Oc then 0 else 1  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4250
            else length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4251
         else if s = ss + 2 * n + 10 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4252
             if l = Bk # ires \<and> r \<noteq> [] \<and> hd r = Oc then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4253
             else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4254
         else if s = ss + 2 * n + 11 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4255
             if r \<noteq> [] \<and> hd r = Oc then 3 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4256
             else 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4257
         else (ss + 2 * n + 16 - s))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4258
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4259
fun abc_dec_2_stage4 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4260
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4261
  "abc_dec_2_stage4 (s, l, r) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4262
          (if s = ss + 2*n + 2 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4263
           else if s = ss + 2*n + 8 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4264
           else if s = ss + 2*n + 3 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4265
               if r \<noteq> [] \<and> hd r = Oc then 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4266
               else 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4267
           else if s = ss + 2*n + 7 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4268
               if r \<noteq> [] \<and> hd r = Oc then 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4269
               else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4270
           else if s = ss + 2*n + 9 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4271
               if r \<noteq> [] \<and> hd r = Oc then 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4272
               else 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4273
           else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4274
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4275
fun abc_dec_2_measure :: "(t_conf \<times> nat \<times> nat \<times> block list) \<Rightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4276
                                    (nat \<times> nat \<times> nat \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4277
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4278
  "abc_dec_2_measure (c, ss, n, ires) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4279
       (abc_dec_2_stage1 c ss n, abc_dec_2_stage2 c ss n,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4280
        abc_dec_2_stage3 c ss n ires, abc_dec_2_stage4 c ss n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4281
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4282
definition abc_dec_2_LE :: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4283
       "(((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list) \<times> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4284
        ((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4285
  where "abc_dec_2_LE \<equiv> (inv_image lex_square abc_dec_2_measure)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4286
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4287
lemma wf_dec_2_le: "wf abc_dec_2_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4288
by(auto intro:wf_inv_image wf_lex_triple wf_lex_square 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4289
   simp:abc_dec_2_LE_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4290
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4291
lemma [simp]: "dec_after_write (as, am) (s, aa, r) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4292
           \<Longrightarrow> takeWhile (\<lambda>a. a = Oc) aa = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4293
apply(simp only : dec_after_write.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4294
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4295
apply(erule_tac conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4296
apply(case_tac aa, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4297
apply(case_tac a, simp only: takeWhile.simps , simp, simp split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4298
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4299
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4300
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4301
     "\<lbrakk>dec_on_right_moving (as, lm) (s, aa, []) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4302
       length (takeWhile (\<lambda>a. a = Oc) (tl aa)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4303
           \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4304
    \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (tl aa)) < 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4305
                       length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4306
apply(simp only: dec_on_right_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4307
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4308
apply(erule_tac conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4309
apply(case_tac mr, auto split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4310
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4311
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4312
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4313
  "dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4314
             (start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4315
 \<Longrightarrow> length xs - Suc 0 < length xs + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4316
                             length (takeWhile (\<lambda>a. a = Oc) aa)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4317
apply(simp only: dec_after_clear.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4318
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4319
apply(erule conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4320
apply(simp split: if_splits )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4321
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4322
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4323
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4324
 "\<lbrakk>dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4325
       (start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4326
    \<Longrightarrow> Suc 0 < length (takeWhile (\<lambda>a. a = Oc) aa)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4327
apply(simp add: dec_after_clear.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4328
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4329
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4330
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4331
 "\<lbrakk>dec_on_right_moving (as, am) (s, aa, Bk # xs) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4332
   Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4333
   \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4334
  \<Longrightarrow> Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4335
    < length (takeWhile (\<lambda>a. a = Oc) aa)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4336
apply(simp only: dec_on_right_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4337
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4338
apply(erule conjE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4339
apply(case_tac ml, auto split: if_splits )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4340
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4341
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4342
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4343
lemma abc_dec_2_wf: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4344
     "\<lbrakk>ly = layout_of aprog; dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r);  abc_fetch as aprog = Dec n e; abc_lm_v am n > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4345
       \<Longrightarrow> \<forall>na. \<not> (\<lambda>(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4346
        (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4347
           (start_of (layout_of aprog) as, n) \<longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4348
        ((t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4349
            start_of (layout_of aprog) as, n),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4350
          t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4351
           start_of (layout_of aprog) as, n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4352
        \<in> abc_dec_2_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4353
proof(rule allI, rule impI, simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4354
  fix na
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4355
  assume h1 :"ly = layout_of aprog" "dec_inv_2 (layout_of aprog) n e (as, am) (Suc (start_of (layout_of aprog) as), l, r)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4356
          "abc_fetch as aprog = Dec n e" "abc_lm_v am n > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4357
         "\<not> (\<lambda>(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4358
             (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4359
             (start_of (layout_of aprog) as, n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4360
  thus "((t_step (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4361
               (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4362
              start_of (layout_of aprog) as, n),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4363
             t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4364
             start_of (layout_of aprog) as, n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4365
            \<in> abc_dec_2_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4366
  proof(insert dec_inv_2_steps[of "layout_of aprog" n e as am "(start_of (layout_of aprog) as + 1, l, r)" aprog na], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4367
        case_tac "(t_steps (start_of (layout_of aprog) as + 1, l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)", case_tac b, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4368
    fix a b aa ba
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4369
    assume "dec_inv_2 (layout_of aprog) n e (as, am) (a, aa, ba)" " a \<noteq> start_of (layout_of aprog) as + 2*n + 16" "abc_lm_v am n > 0" "abc_fetch as aprog = Dec n e "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4370
    thus "((t_step (a, aa, ba) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0), start_of (layout_of aprog) as, n), (a, aa, ba),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4371
                    start_of (layout_of aprog) as, n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4372
                   \<in> abc_dec_2_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4373
      apply(case_tac "a = 0", auto split:if_splits simp add:t_step.simps dec_inv_2.simps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4374
                tactic {* ALLGOALS (resolve_tac (thms "fetch_intro")) *})
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4375
      apply(simp_all add:dec_fetch_simps new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4376
      apply(auto simp add: abc_dec_2_LE_def  lex_square_def lex_triple_def lex_pair_def  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4377
                           split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4378
      
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4379
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4380
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4381
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4382
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4383
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4384
lemma [simp]: "inv_check_left_moving (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4385
  (start_of (layout_of aprog) as + 2 * n + 11, [], Oc # xs) ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4386
apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4387
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4388
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4389
lemma dec_inv_stop2_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4390
  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); abc_lm_v am n > 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4391
    \<forall>na. \<not> (\<lambda>(s, l, r) (ss, n', ires'). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4392
                     s = start_of (layout_of aprog) as + 2 * n + 16)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4393
   (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4394
      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4395
            start_of (layout_of aprog) as - Suc 0) na)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4396
    (start_of (layout_of aprog) as, n, ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4397
 dec_inv_2 (layout_of aprog) n e (as, am)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4398
     (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4399
      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4400
          start_of (layout_of aprog) as - Suc 0) na) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4401
 \<longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4402
 dec_inv_2 (layout_of aprog) n e (as, am)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4403
     (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4404
      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4405
              start_of (layout_of aprog) as - Suc 0) (Suc na)) ires \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4406
 ((t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4407
      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4408
            start_of (layout_of aprog) as - Suc 0) (Suc na),  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4409
              start_of (layout_of aprog) as, n, ires),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4410
  t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4411
      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4412
             start_of (layout_of aprog) as - Suc 0) na,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4413
                          start_of (layout_of aprog) as, n, ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4414
   \<in> abc_dec_2_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4415
apply(subgoal_tac "start_of (layout_of aprog) as > 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4416
apply(rule allI, rule impI, simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4417
apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4418
     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4419
             start_of (layout_of aprog) as - Suc 0) na)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4420
apply(auto split:if_splits simp add:t_step.simps dec_inv_2.simps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4421
           tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4422
apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_2.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4423
apply(auto simp add: abc_dec_2_LE_def lex_square_def lex_triple_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4424
                     lex_pair_def split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4425
apply(auto intro: dec_false_2_b dec_false_2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4426
apply(rule startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4427
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4428
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4429
lemma dec_stop2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4430
 "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4431
   dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4432
   abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4433
   abc_lm_v am n > 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4434
  (\<exists> stp. (\<lambda> (s', l', r'). s' = start_of ly (Suc as) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4435
   dec_inv_2 ly n e (as, am) (s', l', r') ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4436
       (t_steps (start_of ly as+1, l, r) (ci ly (start_of ly as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4437
                           (Dec n e), start_of ly as - Suc 0) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4438
apply(insert halt_lemma2[of abc_dec_2_LE 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4439
      "\<lambda> ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4440
      "(\<lambda> stp. (t_steps (start_of ly as + 1, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4441
       (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4442
                 start_of ly as, n, ires))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4443
      "(\<lambda> ((s, l, r), ss, n, ires'). dec_inv_2 ly n e (as, am) (s, l, r) ires')"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4444
apply(insert wf_dec_2_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4445
apply(insert dec_inv_stop2_pre[of as aprog n e am l r], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4446
      simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4447
apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4448
apply(rule_tac x = na in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4449
apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4450
(ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4451
            start_of (layout_of aprog) as - Suc 0) na)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4452
      case_tac b, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4453
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4454
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4455
lemma dec_inv_stop_cond1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4456
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4457
    dec_inv_1 ly n e (as, lm) (s, (l, r)) ires; s = start_of ly e;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4458
    abc_fetch as aprog = Some (Dec n e); abc_lm_v lm n = 0\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4459
   \<Longrightarrow> crsp_l ly (e, abc_lm_s lm n 0) (s, l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4460
apply(simp add: dec_inv_1.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4461
apply(auto simp: crsp_l.simps inv_stop.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4462
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4463
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4464
lemma dec_inv_stop_cond2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4465
   "\<lbrakk>ly = layout_of aprog; s = start_of ly (Suc as); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4466
     dec_inv_2 ly n e (as, lm) (s, (l, r)) ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4467
     abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4468
     abc_lm_v lm n > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4469
   \<Longrightarrow> crsp_l ly (Suc as,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4470
                  abc_lm_s lm n (abc_lm_v lm n - Suc 0)) (s, l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4471
apply(simp add: dec_inv_2.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4472
apply(auto simp: crsp_l.simps inv_stop.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4473
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4474
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4475
lemma [simp]: "(case Bk\<^bsup>rn\<^esup> of [] \<Rightarrow> Bk |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4476
                 Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) = Bk"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4477
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4478
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4479
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4480
lemma [simp]: "t_steps tc (p,off) (m + n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4481
                   t_steps (t_steps tc (p, off) m) (p, off) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4482
apply(induct m arbitrary: n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4483
apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4484
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4485
  fix m n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4486
  assume h1: "\<And>n. t_steps tc (p, off) (m + n) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4487
                     t_steps (t_steps tc (p, off) m) (p, off) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4488
  hence h2: "t_steps tc (p, off) (Suc m + n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4489
                     t_steps tc (p, off) (m + Suc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4490
    by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4491
  from h1 and this show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4492
    "t_steps tc (p, off) (Suc m + n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4493
         t_steps (t_steps tc (p, off) (Suc m)) (p, off) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4494
  proof(simp only: h2, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4495
    have h3: "(t_step (t_steps tc (p, off) m) (p, off)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4496
                      (t_steps (t_step tc (p, off)) (p, off) m)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4497
      apply(simp add: t_steps.simps[THEN sym] t_steps_ind[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4498
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4499
    from h3 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4500
      "t_steps (t_step (t_steps tc (p, off) m) (p, off)) (p, off) n =          t_steps (t_steps (t_step tc (p, off)) (p, off) m) (p, off) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4501
      by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4502
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4503
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4504
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4505
lemma [simp]: " abc_fetch as aprog = Some (Dec n e) \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4506
          Suc (start_of (layout_of aprog) as) \<noteq> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4507
                           start_of (layout_of aprog) e"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4508
apply(case_tac "e = as", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4509
apply(case_tac "e < as")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4510
apply(drule_tac a = e and b = as and ly = "layout_of aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4511
                                           in start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4512
apply(drule_tac start_of_ge, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4513
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4514
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4515
lemma [simp]: "inv_locate_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4516
apply(auto simp: inv_locate_b.simps in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4517
apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4518
      rule_tac x = 0 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4519
apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4520
apply(case_tac rn, simp, case_tac nat, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4521
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4522
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4523
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4524
       "inv_locate_n_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4525
apply(auto simp: inv_locate_n_b.simps in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4526
apply(case_tac rn, simp, case_tac nat, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4527
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4528
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4529
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4530
"abc_fetch as aprog = Some (Dec n e) \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4531
   dec_inv_1 (layout_of aprog) n e (as, []) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4532
    (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4533
\<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4534
   dec_inv_2 (layout_of aprog) n e (as, []) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4535
    (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4536
apply(simp add: dec_inv_1.simps dec_inv_2.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4537
apply(case_tac n, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4538
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4539
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4540
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4541
 "\<lbrakk>am \<noteq> []; <am> = Oc # r'; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4542
   abc_fetch as aprog = Some (Dec n e)\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4543
 \<Longrightarrow> inv_locate_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4544
apply(auto simp: inv_locate_b.simps in_middle.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4545
apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4546
      rule_tac x = "hd am" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4547
apply(rule_tac x = "Suc 0" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4548
apply(rule_tac x = "hd am" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4549
apply(case_tac am, simp, case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4550
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4551
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4552
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4553
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4554
  "\<lbrakk><am> = Oc # r'; abc_fetch as aprog = Some (Dec n e)\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4555
  inv_locate_n_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4556
apply(auto simp: inv_locate_n_b.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4557
apply(rule_tac x = "tl am" in exI, rule_tac x = "hd am" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4558
apply(case_tac [!] am, auto simp: tape_of_nl_abv tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4559
apply(case_tac [!]list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4560
apply(case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4561
apply(erule_tac x = nat in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4562
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4563
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4564
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4565
   "\<lbrakk>am \<noteq> [];  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4566
     <am> = Oc # r'; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4567
     abc_fetch as aprog = Some (Dec n e)\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4568
    dec_inv_1 (layout_of aprog) n e (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4569
      (Suc (start_of (layout_of aprog) as), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4570
           Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4571
    dec_inv_2 (layout_of aprog) n e (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4572
      (Suc (start_of (layout_of aprog) as), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4573
           Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4574
apply(simp add: dec_inv_1.simps dec_inv_2.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4575
apply(case_tac n, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4576
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4577
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4578
lemma [simp]: "am \<noteq> [] \<Longrightarrow>  \<exists>r'. <am::nat list> = Oc # r'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4579
apply(case_tac am, simp, case_tac list)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4580
apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4581
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4582
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4583
lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4584
      (fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4585
           (start_of (layout_of aprog) as) (Dec n e)) (Suc 0)  Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4586
    = (W1, start_of (layout_of aprog) as)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4587
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4588
             nth_of.simps tshift.simps nth_append Suc_pre tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4589
thm findnth_nth
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4590
apply(insert findnth_nth[of 0 n 0], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4591
apply(simp add: findnth.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4592
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4593
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4594
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4595
    "start_of (layout_of aprog) as > 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4596
   \<Longrightarrow> (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Bk\<^bsup>rn\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4597
    (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4598
                             start_of (layout_of aprog) as - Suc 0))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4599
   = (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn- Suc 0\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4600
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4601
apply(case_tac "start_of (layout_of aprog) as",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4602
      auto simp: new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4603
apply(case_tac rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4604
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4605
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4606
lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4607
 (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4608
         (Dec n e)) (Suc 0)  Oc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4609
  = (R, Suc (start_of (layout_of aprog) as))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4610
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4611
apply(auto simp: ci.simps findnth.simps fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4612
                 nth_of.simps tshift.simps nth_append 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4613
                 Suc_pre tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4614
apply(insert findnth_nth[of 0 n "Suc 0"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4615
apply(simp add: findnth.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4616
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4617
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4618
lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4619
 (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn - Suc 0\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4620
     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4621
        start_of (layout_of aprog) as - Suc 0)) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4622
  (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn-Suc 0\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4623
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4624
apply(case_tac "start_of (layout_of aprog) as", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4625
      auto simp: new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4626
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4627
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4628
lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4629
 t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # r' @ Bk\<^bsup>rn\<^esup>) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4630
      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4631
                 start_of (layout_of aprog) as - Suc 0) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4632
      (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4633
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4634
apply(case_tac "start_of (layout_of aprog) as", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4635
      auto simp: new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4636
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4637
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4638
lemma crsp_next_state:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4639
  "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4640
    abc_fetch as aprog = Some (Dec n e)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4641
  \<Longrightarrow> \<exists> stp' > 0. (\<lambda> (s, l, r). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4642
           (s = Suc (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4643
 \<and> (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r) ires) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4644
 \<and> (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4645
     (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4646
             (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4647
apply(subgoal_tac "start_of (layout_of aprog) as > 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4648
apply(case_tac tc, case_tac b, auto simp: crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4649
apply(case_tac "am = []", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4650
apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4651
proof-
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4652
  fix  rn
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4653
  assume h1: "am \<noteq> []" "abc_fetch as aprog = Some (Dec n e)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4654
             "start_of (layout_of aprog) as > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4655
  hence h2: "\<exists> r'. <am> = Oc # r'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4656
    by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4657
  from h1 and h2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4658
   "\<exists>stp'>0. case t_steps (start_of (layout_of aprog) as, Bk # Bk # ires, <am> @ Bk\<^bsup>rn\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4659
    (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4660
    start_of (layout_of aprog) as - Suc 0) stp' of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4661
    (s, ab) \<Rightarrow> s = Suc (start_of (layout_of aprog) as) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4662
    dec_inv_1 (layout_of aprog) n e (as, am) (s, ab) ires \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4663
    dec_inv_2 (layout_of aprog) n e (as, am) (s, ab) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4664
  proof(erule_tac exE, simp, rule_tac x = "Suc 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4665
        simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4666
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4667
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4668
  assume "abc_fetch as aprog = Some (Dec n e)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4669
  thus "0 < start_of (layout_of aprog) as"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4670
   apply(insert startof_not0[of "layout_of aprog" as], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4671
   done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4672
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4673
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4674
lemma dec_crsp_ex1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4675
  "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4676
  abc_fetch as aprog = Some (Dec n e); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4677
  abc_lm_v am n = 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4678
  \<Longrightarrow> \<exists>stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4679
    (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4680
 (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4681
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4682
  assume h1: "crsp_l (layout_of aprog) (as, am) tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4683
       "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4684
  hence h2: "\<exists> stp' > 0. (\<lambda> (s, l, r). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4685
    (s = Suc (start_of (layout_of aprog) as) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4686
 (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r)) ires)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4687
   (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4688
      (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4689
    apply(insert crsp_next_state[of aprog as am tc ires n e], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4690
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4691
  from h1 and h2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4692
 "\<exists>stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4693
           (t_steps tc (ci (layout_of aprog) (start_of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4694
                (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4695
                    start_of (layout_of aprog) as - Suc 0) stp) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4696
  proof(erule_tac exE, case_tac "(t_steps tc (ci (layout_of aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4697
       (start_of (layout_of aprog) as) (Dec n e), start_of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4698
          (layout_of aprog) as - Suc 0) stp')",  simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4699
    fix stp' a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4700
    assume h3: "stp' > 0 \<and> a = Suc (start_of (layout_of aprog) as) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4701
               dec_inv_1 (layout_of aprog) n e (as, am) (a, b, c) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4702
             "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4703
     "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4704
          (Dec n e), start_of (layout_of aprog) as - Suc 0) stp' 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4705
        = (Suc (start_of (layout_of aprog) as), b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4706
    thus "\<exists>stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4707
     (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4708
           (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4709
    proof(erule_tac conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4710
      assume "dec_inv_1 (layout_of aprog) n e (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4711
                    (Suc (start_of (layout_of aprog) as), b, c) ires"     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4712
             "abc_fetch as aprog = Some (Dec n e)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4713
             "abc_lm_v am n = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4714
             " t_steps tc (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4715
              (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4716
               start_of (layout_of aprog) as - Suc 0) stp' 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4717
             = (Suc (start_of (layout_of aprog) as), b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4718
      hence h4: "\<exists>stp. (\<lambda>(s', l', r'). s' = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4719
                     start_of (layout_of aprog) e \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4720
                dec_inv_1 (layout_of aprog) n e (as, am) (s', l', r') ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4721
                 (t_steps (start_of (layout_of aprog) as + 1, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4722
                  (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4723
                      (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4724
                         start_of (layout_of aprog) as - Suc 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4725
	apply(rule_tac dec_inv_stop1, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4726
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4727
      from  h3 and h4 show ?thesis
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4728
	apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4729
	apply(rule_tac x = "stp' + stp" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4730
	apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4731
                     b, c) (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4732
                     (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4733
                      start_of (layout_of aprog) as - Suc 0) stp)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4734
              simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4735
	apply(rule_tac dec_inv_stop_cond1, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4736
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4737
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4738
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4739
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4740
	  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4741
lemma dec_crsp_ex2:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4742
  "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4743
    abc_fetch as aprog = Some (Dec n e);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4744
    0 < abc_lm_v am n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4745
 \<Longrightarrow> \<exists>stp > 0. crsp_l (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4746
               (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4747
   (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4748
              (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4749
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4750
  assume h1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4751
 "crsp_l (layout_of aprog) (as, am) tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4752
 "abc_fetch as aprog = Some (Dec n e)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4753
  "abc_lm_v am n > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4754
  hence h2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4755
 "\<exists> stp' > 0. (\<lambda> (s, l, r). (s = Suc (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4756
 \<and> (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4757
(t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4758
              (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4759
    apply(insert crsp_next_state[of aprog as am tc ires n e], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4760
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4761
  from h1 and h2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4762
 "\<exists>stp >0. crsp_l (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4763
   (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4764
   (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4765
               (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4766
  proof(erule_tac exE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4767
        case_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4768
 "(t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4769
      (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')",  simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4770
    fix stp' a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4771
    assume h3: "0 < stp' \<and> a = Suc (start_of (layout_of aprog) as) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4772
               dec_inv_2 (layout_of aprog) n e (as, am) (a, b, c) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4773
               "abc_fetch as aprog = Some (Dec n e)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4774
               "abc_lm_v am n > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4775
               "t_steps tc (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4776
                   (start_of (layout_of aprog) as) (Dec n e), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4777
                     start_of (layout_of aprog) as - Suc 0) stp' 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4778
                  = (Suc (start_of (layout_of aprog) as), b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4779
    thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4780
    proof(erule_tac conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4781
      assume 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4782
    "dec_inv_2 (layout_of aprog) n e (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4783
      (Suc (start_of (layout_of aprog) as), b, c) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4784
    "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4785
    "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4786
         (Dec n e), start_of (layout_of aprog) as - Suc 0) stp'
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4787
             = (Suc (start_of (layout_of aprog) as), b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4788
      hence h4: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4789
   "\<exists>stp. (\<lambda>(s', l', r'). s' = start_of (layout_of aprog) (Suc as) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4790
           dec_inv_2 (layout_of aprog) n e (as, am) (s', l', r') ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4791
             (t_steps (start_of (layout_of aprog) as + 1, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4792
              (ci (layout_of aprog) (start_of (layout_of aprog) as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4793
               (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4794
	apply(rule_tac dec_stop2, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4795
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4796
      from  h3 and h4 show ?thesis
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4797
	apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4798
	apply(rule_tac x = "stp' + stp" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4799
	apply(case_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4800
         "(t_steps (Suc (start_of (layout_of aprog) as), b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4801
           (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4802
             (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4803
              ,simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4804
	apply(rule_tac dec_inv_stop_cond2, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4805
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4806
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4807
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4808
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4809
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4810
lemma dec_crsp_ex_pre:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4811
  "\<lbrakk>ly = layout_of aprog; crsp_l ly (as, am) tc ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4812
     abc_fetch as aprog = Some (Dec n e)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4813
 \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4814
      (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4815
                                       start_of ly as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4816
apply(auto simp: abc_step_l.simps intro: dec_crsp_ex2 dec_crsp_ex1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4817
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4818
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4819
lemma dec_crsp_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4820
  assumes layout: -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4821
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4822
  and dec: -- {* There is an @{text "Dec n e"} instruction at postion @{text "as"} of @{text "aprog"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4823
      "abc_fetch as aprog = Some (Dec n e)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4824
  and correspond: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4825
  -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4826
         configuration @{text "tc"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4827
      *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4828
  "crsp_l ly (as, am) tc ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4829
shows 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4830
   "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4831
      (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4832
                                       start_of ly as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4833
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4834
  from dec_crsp_ex_pre layout dec correspond  show ?thesis by blast
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4835
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4836
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4837
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4838
(*******End: dec crsp********)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4839
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4840
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4841
subsubsection {* Compilation of @{text "Goto n"}*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4842
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4843
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4844
(*******Begin: goto crsp********)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4845
lemma goto_fetch: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4846
     "fetch (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4847
         (start_of (layout_of aprog) as) (Goto n)) (Suc 0)  b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4848
     = (Nop, start_of (layout_of aprog) n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4849
apply(auto simp: ci.simps fetch.simps nth_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4850
           split: block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4851
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4852
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4853
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4854
  Correctness of complied @{text "Goto n"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4855
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4856
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4857
lemma goto_crsp_ex_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4858
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4859
    crsp_l ly (as, am) tc ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4860
    abc_fetch as aprog = Some (Goto n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4861
 \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4862
      (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4863
                                        start_of ly as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4864
apply(rule_tac x = 1 in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4865
apply(simp add: abc_step_l.simps t_steps.simps t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4866
apply(case_tac tc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4867
apply(subgoal_tac "a = start_of (layout_of aprog) as", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4868
apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4869
apply(auto simp: goto_fetch new_tape.simps crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4870
apply(rule startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4871
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4872
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4873
lemma goto_crsp_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4874
  assumes layout: "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4875
  and goto: "abc_fetch as aprog = Some (Goto n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4876
  and correspondence: "crsp_l ly (as, am) tc ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4877
  shows "\<exists>stp>0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4878
              (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4879
                                           start_of ly as - Suc 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4880
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4881
  from goto_crsp_ex_pre and layout goto correspondence show "?thesis" by blast
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4882
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4883
(*******End : goto crsp*********)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4884
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4885
subsubsection {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4886
  The correctness of the compiler
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4887
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4888
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4889
declare abc_step_l.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4890
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4891
lemma tm_crsp_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4892
         "\<lbrakk>ly = layout_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4893
           crsp_l ly (as, am) tc ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4894
           as < length aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4895
           abc_fetch as aprog = Some ins\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4896
      \<Longrightarrow> \<exists> n > 0. crsp_l ly (abc_step_l (as,am) (Some ins))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4897
               (t_steps tc (ci (layout_of aprog) (start_of ly as) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4898
                  (ins), (start_of ly as) - (Suc 0)) n) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4899
apply(case_tac "ins", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4900
apply(auto intro: inc_crsp_ex_pre dec_crsp_ex goto_crsp_ex)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4901
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4902
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4903
lemma start_of_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4904
  "n < length aprog \<Longrightarrow> start_of (layout_of aprog) n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4905
                     = start_of (layout_of (butlast aprog)) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4906
apply(induct n, simp add: start_of.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4907
apply(simp add: layout_of.simps start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4908
apply(subgoal_tac "n < length aprog - Suc 0", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4909
apply(subgoal_tac "(aprog ! n) = (butlast aprog ! n)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4910
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4911
  fix n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4912
  assume h1: "Suc n < length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4913
  thus "aprog ! n = butlast aprog ! n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4914
    apply(case_tac "length aprog", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4915
    apply(insert nth_append[of "butlast aprog" "[last aprog]" n])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4916
    apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4917
    apply(simp split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4918
    apply(rule append_butlast_last_id, case_tac aprog, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4919
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4920
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4921
  fix n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4922
  assume "Suc n < length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4923
  thus "n < length aprog - Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4924
    apply(case_tac aprog, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4925
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4926
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4927
    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4928
lemma zip_eq: "xs = ys \<Longrightarrow> zip xs zs = zip ys zs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4929
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4930
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4931
lemma tpairs_of_append_iff: "length aprog = Suc n \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4932
         tpairs_of aprog = tpairs_of (butlast aprog) @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4933
                     [(start_of (layout_of aprog) n, aprog ! n)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4934
apply(simp add: tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4935
apply(insert zip_append[of "map (start_of (layout_of aprog)) [0..<n]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4936
     "butlast aprog" "[start_of (layout_of aprog) n]" "[last aprog]"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4937
apply(simp del: zip_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4938
apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4939
apply(rule_tac zip_eq, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4940
apply(rule_tac start_of_pre, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4941
apply(insert last_conv_nth[of aprog], case_tac aprog, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4942
apply(rule append_butlast_last_id, case_tac aprog, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4943
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4944
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4945
lemma [simp]: "list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4946
         (zip (map (start_of layout) [0..<length aprog]) aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4947
proof(induct "length aprog" arbitrary: aprog, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4948
  fix x aprog
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4949
  assume ind: "\<And>aprog. x = length aprog \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4950
        list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4951
           (zip (map (start_of layout) [0..<length aprog]) aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4952
  and h: "Suc x = length (aprog::abc_inst list)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4953
  have g1: "list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4954
    (zip (map (start_of layout) [0..<length (butlast aprog)]) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4955
                                                 (butlast aprog))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4956
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4957
    apply(rule_tac ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4958
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4959
  have g2: "(map (start_of layout) [0..<length aprog]) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4960
                     map (start_of layout) ([0..<length aprog - 1] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4961
         @ [length aprog - 1])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4962
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4963
    apply(case_tac aprog, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4964
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4965
  have "\<exists> xs a. aprog = xs @ [a]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4966
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4967
    apply(rule_tac x = "butlast aprog" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4968
          rule_tac x = "last aprog" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4969
    apply(case_tac "aprog = []", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4970
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4971
  from this obtain xs where "\<exists> a. aprog = xs @ [a]" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4972
  from this obtain a where g3: "aprog = xs @ [a]" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4973
  from g1 and g2 and g3 show "list_all (\<lambda>(n, tm). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4974
                              abacus.t_ncorrect (ci layout n tm)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4975
              (zip (map (start_of layout) [0..<length aprog]) aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4976
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4977
    apply(auto simp: t_ncorrect.simps ci.simps  tshift.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4978
          tinc_b_def tdec_b_def split: abc_inst.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4979
    apply arith+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4980
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4981
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4982
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4983
lemma [intro]: "abc2t_correct aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4984
apply(simp add: abc2t_correct.simps tpairs_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4985
          split: abc_inst.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4986
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4987
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4988
lemma as_out: "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4989
                crsp_l ly (as, am) tc ires; length aprog \<le> as\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4990
            \<Longrightarrow> abc_step_l (as, am) (abc_fetch as aprog) = (as, am)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4991
apply(simp add: abc_fetch.simps abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4992
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4993
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4994
lemma tm_merge_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4995
  "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4996
    as < length aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4997
    abc_fetch as aprog = Some a; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4998
    abc2t_correct aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4999
    crsp_l (layout_of aprog) (abc_step_l (as, am) (Some a))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5000
     (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5001
         a, start_of (layout_of aprog) as - Suc 0) n) ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5002
    n > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5003
   \<Longrightarrow> \<exists>stp > 0. crsp_l (layout_of aprog) (abc_step_l (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5004
                       (Some a)) (t_steps tc (tm_of aprog, 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5005
apply(case_tac "(t_steps tc (ci (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5006
           (start_of (layout_of aprog) as) a, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5007
            start_of (layout_of aprog) as - Suc 0) n)",  simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5008
apply(case_tac "(abc_step_l (as, am) (Some a))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5009
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5010
  fix aa b c aaa ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5011
  assume h: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5012
  "crsp_l (layout_of aprog) (as, am) tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5013
  "as < length aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5014
  "abc_fetch as aprog = Some a" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5015
  "crsp_l (layout_of aprog) (aaa, ba) (aa, b, c) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5016
  "abc2t_correct aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5017
  "n>0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5018
  "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) a,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5019
      start_of (layout_of aprog) as - Suc 0) n = (aa, b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5020
   "abc_step_l (as, am) (Some a) = (aaa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5021
  hence "aa = start_of (layout_of aprog) aaa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5022
    apply(simp add: crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5023
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5024
  from this and h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5025
  "\<exists>stp > 0. crsp_l (layout_of aprog) (aaa, ba) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5026
                          (t_steps tc (tm_of aprog, 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5027
    apply(insert tms_out_ex[of "layout_of aprog" aprog 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5028
                "tm_of aprog" as am tc ires a n aa b c aaa ba], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5029
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5030
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5031
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5032
lemma crsp_inside: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5033
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5034
    tprog = tm_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5035
    crsp_l ly (as, am) tc ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5036
    as < length aprog\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5037
    (\<exists> stp > 0. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5038
                                         (t_steps tc (tprog, 0) stp) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5039
apply(case_tac "abc_fetch as aprog", simp add: abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5040
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5041
  fix a
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5042
  assume "ly = layout_of aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5043
     "tprog = tm_of aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5044
     "crsp_l ly (as, am) tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5045
     "as < length aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5046
     "abc_fetch as aprog = Some a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5047
  thus "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5048
                 (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5049
    proof(insert tm_crsp_ex[of ly aprog as am tc ires a], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5050
          auto intro: tm_merge_ex)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5051
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5052
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5053
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5054
lemma crsp_outside: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5055
  "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5056
    crsp_l ly (as, am) tc ires; as \<ge> length aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5057
    \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5058
                                         (t_steps tc (tprog, 0) stp) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5059
apply(subgoal_tac "abc_step_l (as, am) (abc_fetch as aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5060
                = (as, am)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5061
apply(rule_tac x = 0 in exI, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5062
apply(rule as_out, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5063
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5064
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5065
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5066
  Single-step correntess of the compiler.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5067
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5068
lemma astep_crsp_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5069
      "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5070
        tprog = tm_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5071
        crsp_l ly (as, am) tc ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5072
       \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5073
                  (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5074
apply(case_tac "as < length aprog")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5075
apply(drule_tac crsp_inside, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5076
apply(rule_tac crsp_outside, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5077
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5078
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5079
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5080
  Single-step correntess of the compiler.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5081
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5082
lemma astep_crsp_pre1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5083
      "\<lbrakk>ly = layout_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5084
        tprog = tm_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5085
        crsp_l ly (as, am) tc ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5086
       \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5087
                  (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5088
apply(case_tac "as < length aprog")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5089
apply(drule_tac crsp_inside, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5090
apply(rule_tac crsp_outside, simp+)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5091
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5092
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5093
lemma astep_crsp:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5094
  assumes layout: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5095
  -- {* There is a Abacus program @{text "aprog"} with layout @{text "ly"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5096
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5097
  and compiled: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5098
  -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5099
  "tprog = tm_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5100
  and corresponds: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5101
  -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM configuration
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5102
   @{text "tc"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5103
  "crsp_l ly (as, am) tc ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5104
  -- {* One step execution of @{text "aprog"} can be simulated by multi-step execution 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5105
  of @{text "tprog"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5106
  shows "(\<exists> stp. crsp_l ly (abc_step_l (as,am) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5107
                  (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5108
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5109
  from astep_crsp_pre1 [OF layout compiled corresponds] show ?thesis .
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5110
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5111
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5112
lemma steps_crsp_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5113
    "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5114
      crsp_l ly ac tc ires; ac' = abc_steps_l ac aprog n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5115
        (\<exists> n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5116
apply(induct n arbitrary: ac' ac tc, simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5117
apply(rule_tac x = 0 in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5118
apply(case_tac ac, simp add: abc_steps_l.simps t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5119
apply(case_tac ac, simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5120
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5121
   "(\<exists> stp. crsp_l ly (abc_step_l (a, b)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5122
            (abc_fetch a aprog)) (t_steps tc (tprog, 0) stp) ires)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5123
apply(erule exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5124
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5125
   "\<exists>n'. crsp_l (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5126
    (abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) aprog n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5127
         (t_steps ((t_steps tc (tprog, 0) stp)) (tm_of aprog, 0) n') ires")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5128
apply(erule exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5129
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5130
    "t_steps (t_steps tc (tprog, 0) stp) (tm_of aprog, 0) n' =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5131
     t_steps tc (tprog, 0) (stp + n')")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5132
apply(rule_tac x = "stp + n'" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5133
apply(auto intro: astep_crsp simp: t_step_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5134
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5135
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5136
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5137
  Multi-step correctess of the compiler.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5138
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5139
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5140
lemma steps_crsp: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5141
  assumes layout: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5142
  -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5143
    "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5144
  and compiled: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5145
  -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5146
  "tprog = tm_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5147
  and correspond: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5148
  -- {* Abacus configuration @{text "ac"} is in correspondence with TM configuration @{text "tc"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5149
      "crsp_l ly ac tc ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5150
  and execution: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5151
  -- {* @{text "ac'"} is the configuration obtained from @{text "n"}-step execution 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5152
      of @{text "aprog"} starting from configuration @{text "ac"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5153
  "ac' = abc_steps_l ac aprog n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5154
  -- {* There exists steps @{text "n'"} steps, after these steps of execution, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5155
  the Turing configuration such obtained is in correspondence with @{text "ac'"} *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5156
  shows "(\<exists> n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5157
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5158
  from steps_crsp_pre [OF layout compiled correspond execution] show ?thesis .
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5159
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5160
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5161
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5162
subsubsection {* The Mop-up machine *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5163
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5164
fun mop_bef :: "nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5165
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5166
  "mop_bef 0 = []" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5167
  "mop_bef (Suc n) = mop_bef n @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5168
       [(R, 2*n + 3), (W0, 2*n + 2), (R, 2*n + 1), (W1, 2*n + 2)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5169
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5170
definition mp_up :: "tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5171
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5172
  "mp_up \<equiv> [(R, 2), (R, 1), (L, 5), (W0, 3), (R, 4), (W0, 3),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5173
            (R, 2), (W0, 3), (L, 5), (L, 6), (R, 0), (L, 6)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5174
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5175
fun tMp :: "nat \<Rightarrow> nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5176
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5177
  "tMp n off = tshift (mop_bef n @ tshift mp_up (2*n)) off"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5178
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5179
declare  mp_up_def[simp del]  tMp.simps[simp del] mop_bef.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5180
(**********Begin: equiv among aba and turing***********)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5181
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5182
lemma tm_append_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5183
 "\<lbrakk>t_ncorrect tp1; t_step tc (tp1, 0) = (s, l, r); s \<noteq> 0\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5184
 \<Longrightarrow> t_step tc (tp1 @ tp2, 0) = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5185
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5186
apply(case_tac tc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5187
apply(case_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5188
       "(fetch tp1 a (case c of [] \<Rightarrow> Bk |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5189
               Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5190
apply(case_tac a, simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5191
apply(simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5192
apply(case_tac c, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5193
apply(case_tac [!] "ab::block")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5194
apply(auto simp: nth_of.simps nth_append t_ncorrect.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5195
           split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5196
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5197
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5198
lemma state0_ind: "t_steps (0, l, r) (tp, 0) stp = (0, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5199
apply(induct stp, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5200
apply(simp add: t_steps.simps t_step.simps fetch.simps new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5201
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5202
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5203
lemma tm_append_steps:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5204
 "\<lbrakk>t_ncorrect tp1; t_steps tc (tp1, 0) stp = (s, l ,r); s \<noteq> 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5205
  \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5206
apply(induct stp arbitrary: tc s l r)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5207
apply(case_tac tc,  simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5208
apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5209
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5210
  fix stp tc s l r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5211
  assume h1: "\<And>tc s l r. \<lbrakk>t_ncorrect tp1; t_steps tc (tp1, 0) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5212
   (s, l, r); s \<noteq> 0\<rbrakk> \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5213
    and h2: "t_steps tc (tp1, 0) (Suc stp) = (s, l, r)" "s \<noteq> 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5214
            "t_ncorrect tp1"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5215
  thus "t_steps tc (tp1 @ tp2, 0) (Suc stp) = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5216
    apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5217
    apply(case_tac "(t_step tc (tp1, 0))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5218
    proof-
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5219
      fix a b c 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5220
      assume g1: "\<And>tc s l r. \<lbrakk>t_steps tc (tp1, 0) stp = (s, l, r); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5221
                0 < s\<rbrakk> \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5222
	and g2: "t_step tc (tp1, 0) = (a, b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5223
                "t_steps (a, b, c) (tp1, 0) stp = (s, l, r)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5224
                "0 < s" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5225
                "t_ncorrect tp1"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5226
      hence g3: "a > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5227
	apply(case_tac "a::nat", auto simp: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5228
	apply(simp add: state0_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5229
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5230
      from g1 and g2 and this have g4: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5231
                    "(t_step tc (tp1 @ tp2, 0)) = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5232
	apply(rule_tac tm_append_step, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5233
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5234
      from g1 and g2 and g3 and g4 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5235
          "t_steps (t_step tc (tp1 @ tp2, 0)) (tp1 @ tp2, 0) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5236
                                                         = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5237
	apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5238
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5239
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5240
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5241
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5242
lemma shift_fetch: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5243
 "\<lbrakk>n < length tp; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5244
  (tp:: (taction \<times> nat) list) ! n = (aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5245
   ba \<noteq> 0\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5246
   \<Longrightarrow> (tshift tp (length tp div 2)) ! n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5247
                     (aa , ba + length tp div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5248
apply(simp add: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5249
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5250
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5251
lemma tshift_length_equal: "length (tshift tp q) = length tp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5252
apply(auto simp: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5253
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5254
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5255
thm nth_of.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5256
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5257
lemma [simp]: "t_ncorrect tp \<Longrightarrow> 2 * (length tp div 2) = length tp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5258
apply(auto simp: t_ncorrect.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5259
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5260
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5261
lemma  tm_append_step_equal': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5262
   "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5263
    (\<lambda> (s, l, r). ((\<lambda> (s', l', r'). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5264
      (s'\<noteq> 0 \<longrightarrow> (s = s' + off \<and> l = l' \<and> r = r'))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5265
         (t_step (a, b, c) (tp', 0))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5266
               (t_step (a + off, b, c) (tp @ tshift tp' off, 0))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5267
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5268
apply(case_tac a, simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5269
apply(case_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5270
"(fetch tp' a (case c of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5271
 simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5272
apply(case_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5273
"(fetch (tp @ tshift tp' (length tp div 2))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5274
        (Suc (nat + length tp div 2)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5275
           (case c of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5276
 simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5277
apply(case_tac "(new_tape aa (b, c))",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5278
      case_tac "(new_tape aaa (b, c))", simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5279
      rule impI, simp add: fetch.simps split: block.splits option.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5280
apply (auto simp: nth_of.simps t_ncorrect.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5281
                      nth_append tshift_length_equal tshift.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5282
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5283
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5284
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5285
lemma  tm_append_step_equal: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5286
 "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5287
   t_step (a, b, c) (tp', 0) = (aa, ab, bb);  aa \<noteq> 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5288
 \<Longrightarrow> t_step (a + length tp div 2, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5289
        (tp @ tshift tp' (length tp div 2), 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5290
                          = (aa + length tp div 2, ab, bb)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5291
apply(insert tm_append_step_equal'[of tp tp' off a b c], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5292
apply(case_tac "(t_step (a + length tp div 2, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5293
                   (tp @ tshift tp' (length tp div 2), 0))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5294
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5295
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5296
lemma tm_append_steps_equal: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5297
 "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5298
   (\<lambda> (s, l, r). ((\<lambda> (s', l', r'). ((s'\<noteq> 0 \<longrightarrow> s = s' + off \<and> l = l'
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5299
                     \<and> r = r'))) (t_steps (a, b, c) (tp', 0) stp)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5300
   (t_steps (a + off, b, c) (tp @ tshift tp' off, 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5301
apply(induct stp arbitrary : a b c, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5302
apply(simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5303
apply(case_tac "(t_step (a, b, c) (tp', 0))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5304
apply(case_tac "aa = 0", simp add: state0_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5305
apply(subgoal_tac "(t_step (a + length tp div 2, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5306
                      (tp @ tshift tp' (length tp div 2), 0)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5307
  = (aa + length tp div 2, ba, ca)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5308
apply(rule tm_append_step_equal, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5309
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5310
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5311
(*********Begin: mop_up***************)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5312
type_synonym mopup_type = "t_conf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5313
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5314
fun mopup_stop :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5315
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5316
  "mopup_stop (s, l, r) lm n ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5317
        (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = <abc_lm_v lm n> @ Bk\<^bsup>rn\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5318
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5319
fun mopup_bef_erase_a :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5320
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5321
  "mopup_bef_erase_a (s, l, r) lm n ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5322
         (\<exists> ln m rn. l = Bk \<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5323
                  r = Oc\<^bsup>m \<^esup>@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<^bsup>rn\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5324
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5325
fun mopup_bef_erase_b :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5326
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5327
  "mopup_bef_erase_b (s, l, r) lm n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5328
      (\<exists> ln m rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = Bk # Oc\<^bsup>m\<^esup> @ Bk # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5329
                                      <(drop (s div 2) lm)> @ Bk\<^bsup>rn\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5330
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5331
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5332
fun mopup_jump_over1 :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5333
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5334
  "mopup_jump_over1 (s, l, r) lm n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5335
      (\<exists> ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5336
        l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5337
     (r = Oc\<^bsup>m2\<^esup> @ Bk # <(drop (Suc n) lm)> @ Bk\<^bsup>rn \<^esup>\<or> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5338
     (r = Oc\<^bsup>m2\<^esup> \<and> (drop (Suc n) lm) = [])))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5339
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5340
fun mopup_aft_erase_a :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5341
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5342
  "mopup_aft_erase_a (s, l, r) lm n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5343
      (\<exists> lnl lnr rn (ml::nat list) m. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5344
          m = Suc (abc_lm_v lm n) \<and> l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5345
                                   (r = <ml> @ Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5346
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5347
fun mopup_aft_erase_b :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5348
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5349
  "mopup_aft_erase_b (s, l, r) lm n ires= 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5350
   (\<exists> lnl lnr rn (ml::nat list) m. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5351
      m = Suc (abc_lm_v lm n) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5352
      l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5353
     (r = Bk # <ml> @ Bk\<^bsup>rn \<^esup>\<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5354
      r = Bk # Bk # <ml> @ Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5355
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5356
fun mopup_aft_erase_c :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5357
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5358
  "mopup_aft_erase_c (s, l, r) lm n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5359
 (\<exists> lnl lnr rn (ml::nat list) m. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5360
     m = Suc (abc_lm_v lm n) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5361
     l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5362
    (r = <ml> @ Bk\<^bsup>rn \<^esup>\<or> r = Bk # <ml> @ Bk\<^bsup>rn\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5363
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5364
fun mopup_left_moving :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5365
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5366
  "mopup_left_moving (s, l, r) lm n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5367
  (\<exists> lnl lnr rn m.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5368
     m = Suc (abc_lm_v lm n) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5369
   ((l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> r = Bk\<^bsup>rn\<^esup>) \<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5370
    (l = Oc\<^bsup>m - 1\<^esup> @ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> r = Oc # Bk\<^bsup>rn\<^esup>)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5371
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5372
fun mopup_jump_over2 :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5373
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5374
  "mopup_jump_over2 (s, l, r) lm n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5375
     (\<exists> ln rn m1 m2.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5376
          m1 + m2 = Suc (abc_lm_v lm n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5377
        \<and> r \<noteq> [] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5378
        \<and> (hd r = Oc \<longrightarrow> (l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = Oc\<^bsup>m2\<^esup> @ Bk\<^bsup>rn\<^esup>)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5379
        \<and> (hd r = Bk \<longrightarrow> (l = Bk\<^bsup>ln\<^esup> @ Bk # ires \<and> r = Bk # Oc\<^bsup>m1 + m2\<^esup> @ Bk\<^bsup>rn\<^esup>)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5380
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5381
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5382
fun mopup_inv :: "mopup_type"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5383
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5384
  "mopup_inv (s, l, r) lm n ires = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5385
      (if s = 0 then mopup_stop (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5386
       else if s \<le> 2*n then
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5387
               if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5388
                   else mopup_bef_erase_b (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5389
            else if s = 2*n + 1 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5390
                mopup_jump_over1 (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5391
            else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5392
            else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5393
            else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5394
            else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5395
            else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5396
            else False)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5397
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5398
declare 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5399
  mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5400
  mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5401
  mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5402
  mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5403
  mopup_stop.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5404
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5405
lemma mopup_fetch_0[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5406
     "(fetch (mop_bef n @ tshift mp_up (2 * n)) 0 b) = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5407
by(simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5408
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5409
lemma mop_bef_length[simp]: "length (mop_bef n) = 4 * n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5410
apply(induct n, simp add: mop_bef.simps, simp add: mop_bef.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5411
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5412
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5413
thm findnth_nth
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5414
lemma mop_bef_nth: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5415
  "\<lbrakk>q < n; x < 4\<rbrakk> \<Longrightarrow> mop_bef n ! (4 * q + x) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5416
                             mop_bef (Suc q) ! ((4 * q) + x)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5417
apply(induct n, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5418
apply(case_tac "q < n", simp add: mop_bef.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5419
apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5420
apply(subgoal_tac "q = n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5421
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5422
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5423
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5424
lemma fetch_bef_erase_a_o[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5425
 "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5426
  \<Longrightarrow> (fetch (mop_bef n @ tshift mp_up (2 * n)) s Oc) = (W0, s + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5427
apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5428
apply(subgoal_tac "length (mop_bef n) = 4*n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5429
apply(auto simp: fetch.simps nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5430
apply(subgoal_tac "mop_bef n ! (4 * q + 1) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5431
                      mop_bef (Suc q) ! ((4 * q) + 1)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5432
      simp add: mop_bef.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5433
apply(rule mop_bef_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5434
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5435
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5436
lemma fetch_bef_erase_a_b[simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5437
  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5438
   \<Longrightarrow>  (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s + 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5439
apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5440
apply(subgoal_tac "length (mop_bef n) = 4*n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5441
apply(auto simp: fetch.simps nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5442
apply(subgoal_tac "mop_bef n ! (4 * q + 0) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5443
                       mop_bef (Suc q) ! ((4 * q + 0))", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5444
      simp add: mop_bef.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5445
apply(rule mop_bef_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5446
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5447
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5448
lemma fetch_bef_erase_b_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5449
  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5450
     (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s - 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5451
apply(subgoal_tac "\<exists> q. s = 2 * q", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5452
apply(case_tac qa, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5453
apply(auto simp: fetch.simps nth_of.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5454
apply(subgoal_tac "mop_bef n ! (4 * nat + 2) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5455
                     mop_bef (Suc nat) ! ((4 * nat) + 2)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5456
      simp add: mop_bef.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5457
apply(rule mop_bef_nth, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5458
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5459
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5460
lemma fetch_jump_over1_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5461
 "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5462
  = (R, Suc (2 * n))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5463
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5464
apply(auto simp: fetch.simps nth_of.simps mp_up_def nth_append 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5465
                 tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5466
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5467
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5468
lemma fetch_jump_over1_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5469
 "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Bk 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5470
 = (R, Suc (Suc (2 * n)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5471
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5472
apply(auto simp: fetch.simps nth_of.simps mp_up_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5473
                 nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5474
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5475
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5476
lemma fetch_aft_erase_a_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5477
 "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Oc 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5478
 = (W0, Suc (2 * n + 2))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5479
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5480
apply(auto simp: fetch.simps nth_of.simps mp_up_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5481
                 nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5482
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5483
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5484
lemma fetch_aft_erase_a_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5485
 "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5486
  = (L, Suc (2 * n + 4))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5487
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5488
apply(auto simp: fetch.simps nth_of.simps mp_up_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5489
                 nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5490
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5491
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5492
lemma fetch_aft_erase_b_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5493
 "fetch (mop_bef n @ tshift mp_up (2 * n)) (2*n + 3) Bk
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5494
  = (R, Suc (2 * n + 3))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5495
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5496
apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5497
apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5498
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5499
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5500
lemma fetch_aft_erase_c_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5501
 "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Oc 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5502
 = (W0, Suc (2 * n + 2))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5503
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5504
apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5505
apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5506
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5507
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5508
lemma fetch_aft_erase_c_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5509
 "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Bk 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5510
 = (R, Suc (2 * n + 1))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5511
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5512
apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5513
apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5514
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5515
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5516
lemma fetch_left_moving_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5517
 "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Oc) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5518
 = (L, 2*n + 6)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5519
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5520
apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5521
apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5522
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5523
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5524
lemma fetch_left_moving_b: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5525
 "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5526
  = (L, 2*n + 5)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5527
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5528
apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5529
apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5530
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5531
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5532
lemma fetch_jump_over2_b:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5533
  "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Bk) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5534
 = (R, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5535
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5536
apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5537
apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5538
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5539
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5540
lemma fetch_jump_over2_o: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5541
"(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Oc) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5542
 = (L, 2*n + 6)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5543
apply(subgoal_tac "length (mop_bef n) = 4 * n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5544
apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5545
apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5546
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5547
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5548
lemmas mopupfetchs = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5549
fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5550
fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5551
fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5552
fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5553
fetch_jump_over2_b fetch_jump_over2_o
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5554
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5555
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5556
"\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5557
  mopup_bef_erase_a (s, l, Oc # xs) lm n ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5558
  Suc s \<le> 2 * n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5559
  mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5560
apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5561
apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5562
apply(case_tac m, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5563
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5564
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5565
lemma mopup_false1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5566
  "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0;  \<not> Suc s \<le> 2 * n\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5567
  \<Longrightarrow> RR"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5568
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5569
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5570
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5571
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5572
 "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5573
   mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5574
 \<Longrightarrow> (Suc s \<le> 2 * n \<longrightarrow> mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires)  \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5575
     (\<not> Suc s \<le> 2 * n \<longrightarrow> mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5576
apply(auto elim: mopup_false1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5577
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5578
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5579
lemma drop_abc_lm_v_simp[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5580
   "n < length lm \<Longrightarrow> drop n lm = abc_lm_v lm n # drop (Suc n) lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5581
apply(auto simp: abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5582
apply(drule drop_Suc_conv_tl, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5583
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5584
lemma [simp]: "(\<exists>rna. Bk\<^bsup>rn\<^esup> = Bk # Bk\<^bsup>rna\<^esup>) \<or> Bk\<^bsup>rn\<^esup> = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5585
apply(case_tac rn, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5586
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5587
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5588
lemma [simp]: "\<exists>lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5589
apply(rule_tac x = "Suc ln" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5590
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5591
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5592
lemma mopup_bef_erase_a_2_jump_over[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5593
 "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5594
   mopup_bef_erase_a (s, l, Bk # xs) lm n ires; Suc s = 2 * n\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5595
\<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5596
apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5597
apply(case_tac m, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5598
apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5599
      simp add: tape_of_nl_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5600
apply(case_tac "drop (Suc n) lm", auto simp: tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5601
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5602
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5603
lemma Suc_Suc_div:  "\<lbrakk>0 < s; s mod 2 = Suc 0; Suc (Suc s) \<le> 2 * n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5604
           \<Longrightarrow> (Suc (Suc (s div 2))) \<le> n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5605
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5606
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5607
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5608
lemma mopup_bef_erase_a_2_a[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5609
 "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5610
   mopup_bef_erase_a (s, l, Bk # xs) lm n ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5611
   Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5612
   mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5613
apply(auto simp: mopup_bef_erase_a.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5614
apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \<noteq> []")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5615
apply(case_tac m, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5616
apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5617
      rule_tac x = rn in exI, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5618
apply(subgoal_tac "(Suc (Suc (s div 2))) \<le> n", simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5619
      rule_tac Suc_Suc_div, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5620
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5621
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5622
lemma mopup_false2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5623
 "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5624
   s mod 2 = Suc 0; Suc s \<noteq> 2 * n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5625
   \<not> Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> RR"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5626
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5627
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5628
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5629
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5630
  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5631
   s mod 2 = Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5632
   mopup_bef_erase_a (s, l, Bk # xs) lm n ires; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5633
   r = Bk # xs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5634
 \<Longrightarrow> (Suc s = 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5635
             mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5636
     (Suc s \<noteq> 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5637
       (Suc (Suc s) \<le> 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5638
          mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5639
       (\<not> Suc (Suc s) \<le> 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5640
          mopup_aft_erase_a (Suc (Suc s), Bk # l, xs) lm n ires))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5641
apply(auto elim: mopup_false2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5642
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5643
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5644
lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5645
                        mopup_bef_erase_a (s, l, [Bk]) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5646
apply(auto simp: mopup_bef_erase_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5647
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5648
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5649
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5650
   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5651
     mopup_bef_erase_a (s, l, []) lm n ires; r = []\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5652
    \<Longrightarrow> (Suc s = 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5653
              mopup_jump_over1 (Suc (2 * n), Bk # l, []) lm n ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5654
        (Suc s \<noteq> 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5655
             (Suc (Suc s) \<le> 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5656
                 mopup_bef_erase_a (Suc (Suc s), Bk # l, []) lm n ires) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5657
             (\<not> Suc (Suc s) \<le> 2 * n \<longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5658
                 mopup_aft_erase_a (Suc (Suc s), Bk # l, []) lm n ires))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5659
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5660
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5661
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5662
lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5663
apply(auto simp: mopup_bef_erase_b.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5664
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5665
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5666
lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5667
apply(auto simp: mopup_bef_erase_b.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5668
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5669
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5670
lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5671
                                      (s - Suc 0) mod 2 = Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5672
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5673
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5674
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5675
lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5676
                                       s - Suc 0 \<le> 2 * n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5677
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5678
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5679
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5680
lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> \<not> s \<le> Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5681
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5682
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5683
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5684
lemma [simp]: "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5685
               s mod 2 \<noteq> Suc 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5686
               mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5687
           \<Longrightarrow> mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5688
apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5689
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5690
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5691
lemma [simp]: "\<lbrakk>mopup_bef_erase_b (s, l, []) lm n ires\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5692
                   mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5693
apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5694
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5695
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5696
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5697
   "\<lbrakk>n < length lm;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5698
    mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5699
    r = Oc # xs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5700
  \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5701
apply(auto simp: mopup_jump_over1.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5702
apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5703
       rule_tac x = "m2 - 1" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5704
apply(case_tac "m2", simp, simp, rule_tac x = rn in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5705
apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5706
      rule_tac x = "m2 - 1" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5707
apply(case_tac m2, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5708
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5709
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5710
lemma mopup_jump_over1_2_aft_erase_a[simp]:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5711
 "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5712
  \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5713
apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5714
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5715
apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5716
apply(case_tac m2, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5717
apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5718
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5719
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5720
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5721
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5722
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5723
 "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5724
    mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5725
apply(rule mopup_jump_over1_2_aft_erase_a, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5726
apply(auto simp: mopup_jump_over1.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5727
apply(rule_tac x = ln in exI, rule_tac x = m1 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5728
      rule_tac x = m2 in exI, simp add: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5729
apply(rule_tac x = 0 in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5730
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5731
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5732
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5733
 "\<lbrakk>n < length lm; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5734
   mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5735
 \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5736
apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5737
apply(case_tac ml, simp, case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5738
apply(case_tac list, auto simp: tape_of_nl_abv 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5739
                                tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5740
apply(case_tac a, simp, rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5741
      rule_tac x = "[]" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5742
       simp add: tape_of_nat_list.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5743
apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5744
      simp add: tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5745
apply(case_tac a, simp, rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5746
       rule_tac x = "aa # lista" in exI, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5747
apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5748
       simp add: tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5749
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5750
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5751
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5752
  "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5753
apply(auto simp: mopup_aft_erase_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5754
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5755
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5756
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5757
  "\<lbrakk>n < length lm;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5758
    mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5759
  \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5760
apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5761
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5762
apply(case_tac lnr, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5763
apply(rule_tac x = lnl in exI, simp, rule_tac x = rn in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5764
apply(subgoal_tac "ml = []", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5765
apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5766
apply(subgoal_tac "ml = []", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5767
apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5768
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5769
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5770
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5771
  "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5772
apply(simp only: mopup_aft_erase_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5773
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5774
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5775
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5776
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5777
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5778
  "\<lbrakk>n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5779
  \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5780
apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5781
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5782
apply(subgoal_tac "ml = [] \<and> rn = 0", erule conjE, erule conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5783
apply(case_tac lnr, simp, rule_tac x = lnl in exI, simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5784
      rule_tac x = 0 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5785
apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5786
      rule_tac x = "Suc 0" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5787
apply(case_tac ml, simp, case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5788
apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5789
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5790
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5791
lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5792
apply(auto simp: mopup_aft_erase_b.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5793
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5794
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5795
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5796
 "\<lbrakk>n < length lm; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5797
   mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5798
  \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5799
apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5800
apply(case_tac ml, simp, case_tac rn, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5801
apply(case_tac list, auto simp: tape_of_nl_abv 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5802
                        tape_of_nat_list.simps tape_of_nat_abv )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5803
apply(case_tac a, rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5804
      rule_tac x = "[]" in exI, simp add: tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5805
apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5806
      simp add: tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5807
apply(case_tac a, simp, rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5808
      rule_tac x = "aa # lista" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5809
apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5810
      simp add: tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5811
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5812
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5813
lemma mopup_aft_erase_c_aft_erase_a[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5814
 "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5815
 \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5816
apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5817
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5818
apply(erule conjE, erule conjE, erule disjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5819
apply(subgoal_tac "ml = []", simp, case_tac rn, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5820
      simp, simp, rule conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5821
apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5822
apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5823
apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5824
      rule conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5825
apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5826
apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5827
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5828
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5829
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5830
 "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5831
 \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5832
apply(rule mopup_aft_erase_c_aft_erase_a, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5833
apply(simp only: mopup_aft_erase_c.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5834
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5835
apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5836
apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5837
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5838
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5839
lemma mopup_aft_erase_b_2_aft_erase_c[simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5840
  "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\<rbrakk>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5841
 \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5842
apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5843
apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5844
apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5845
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5846
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5847
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5848
 "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5849
 \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5850
apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5851
apply(simp add: mopup_aft_erase_b.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5852
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5853
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5854
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5855
    "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5856
apply(auto simp: mopup_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5857
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5858
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5859
lemma [simp]:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5860
 "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5861
  \<Longrightarrow> mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5862
apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5863
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5864
apply(erule conjE, erule disjE, erule conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5865
apply(case_tac rn, simp, simp add: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5866
apply(case_tac "hd l", simp add:  )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5867
apply(case_tac "abc_lm_v lm n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5868
apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5869
      rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5870
apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5871
apply(case_tac "abc_lm_v lm n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5872
apply(case_tac lnl, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5873
apply(rule_tac x = lnl in exI, rule_tac x = rn in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5874
apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5875
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5876
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5877
lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5878
apply(auto simp: mopup_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5879
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5880
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5881
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5882
  "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5883
 \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5884
apply(simp only: mopup_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5885
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5886
apply(case_tac lnr, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5887
apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5888
      rule_tac x = rn in exI, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5889
apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5890
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5891
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5892
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5893
"\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5894
    \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5895
apply(simp only: mopup_left_moving.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5896
apply(erule exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5897
apply(case_tac lnr, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5898
apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5899
      rule_tac x = 0 in exI, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5900
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5901
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5902
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5903
 "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5904
apply(auto simp: mopup_jump_over2.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5905
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5906
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5907
lemma [intro]: "\<exists>lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup> @ [Bk]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5908
apply(simp only: exp_ind[THEN sym], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5909
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5910
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5911
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5912
"\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5913
 \<Longrightarrow>  mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5914
apply(simp only: mopup_jump_over2.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5915
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5916
apply(simp add:  , erule conjE, erule_tac conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5917
apply(case_tac m1, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5918
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5919
      rule_tac x = 0 in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5920
apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5921
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5922
      rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5923
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5924
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5925
lemma [simp]: "\<exists>rna. Oc # Oc\<^bsup>a\<^esup> @ Bk\<^bsup>rn\<^esup> = <a> @ Bk\<^bsup>rna\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5926
apply(case_tac a, auto simp: tape_of_nat_abv )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5927
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5928
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5929
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5930
 "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5931
  \<Longrightarrow> mopup_stop (0, Bk # l, xs) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5932
apply(auto simp: mopup_jump_over2.simps mopup_stop.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5933
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5934
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5935
lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5936
apply(simp only: mopup_jump_over2.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5937
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5938
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5939
lemma mopup_inv_step:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5940
  "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5941
  \<Longrightarrow> mopup_inv (t_step (s, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5942
       ((mop_bef n @ tshift mp_up (2 * n)), 0)) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5943
apply(auto split:if_splits simp add:t_step.simps,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5944
      tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *})
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5945
apply(simp_all add: mopupfetchs new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5946
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5947
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5948
declare mopup_inv.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5949
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5950
lemma mopup_inv_steps: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5951
"\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5952
     mopup_inv (t_steps (s, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5953
                   ((mop_bef n @ tshift mp_up (2 * n)), 0) stp) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5954
apply(induct stp, simp add: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5955
apply(simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5956
apply(case_tac "(t_steps (s, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5957
                (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5958
apply(rule_tac mopup_inv_step, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5959
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5960
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5961
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5962
 "\<lbrakk>n < length lm; Suc 0 \<le> n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5963
            mopup_bef_erase_a (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5964
apply(auto simp: mopup_bef_erase_a.simps  abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5965
apply(case_tac lm, simp, case_tac list, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5966
apply(rule_tac x = "Suc a" in exI, rule_tac x = rn in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5967
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5968
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5969
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5970
  "lm \<noteq> [] \<Longrightarrow> mopup_jump_over1 (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) lm 0  ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5971
apply(auto simp: mopup_jump_over1.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5972
apply(rule_tac x = ln in exI, rule_tac x = 0 in exI, simp add: )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5973
apply(case_tac lm, simp, simp add: abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5974
apply(case_tac rn, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5975
apply(case_tac list, rule_tac disjI2, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5976
      simp add: tape_of_nl_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5977
apply(rule_tac disjI1,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5978
      simp add: tape_of_nl_abv tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5979
apply(rule_tac disjI1, case_tac list, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5980
      simp add: tape_of_nl_abv tape_of_nat_list.simps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5981
      rule_tac x = nat in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5982
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5983
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5984
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5985
lemma mopup_init: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5986
 "\<lbrakk>n < length lm; crsp_l ly (as, lm) (ac, l, r) ires\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5987
                               mopup_inv (Suc 0, l, r) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5988
apply(auto simp: crsp_l.simps mopup_inv.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5989
apply(case_tac n, simp, auto simp: mopup_bef_erase_a.simps )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5990
apply(rule_tac x = "Suc (hd lm)" in exI, rule_tac x = rn in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5991
apply(case_tac lm, simp, case_tac list, simp, case_tac lista, simp add: abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5992
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5993
apply(simp add: mopup_jump_over1.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5994
apply(rule_tac x = 0 in exI, rule_tac x = 0 in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5995
apply(case_tac [!] n, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5996
apply(case_tac [!] lm, simp, case_tac list, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5997
apply(case_tac rn, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5998
apply(erule_tac x = nat in allE, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5999
apply(simp add: abc_lm_v.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6000
apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6001
apply(erule_tac x = rn in allE, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6002
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6003
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6004
(***Begin: mopup stop***)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6005
fun abc_mopup_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6006
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6007
  "abc_mopup_stage1 (s, l, r) n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6008
           (if s > 0 \<and> s \<le> 2*n then 6
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6009
            else if s = 2*n + 1 then 4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6010
            else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6011
            else if s = 2*n + 5 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6012
            else if s = 2*n + 6 then 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6013
            else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6014
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6015
fun abc_mopup_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6016
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6017
  "abc_mopup_stage2 (s, l, r) n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6018
           (if s > 0 \<and> s \<le> 2*n then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6019
            else if s = 2*n + 1 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6020
            else if s = 2*n + 5 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6021
            else if s = 2*n + 6 then length l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6022
            else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then length r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6023
            else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6024
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6025
fun abc_mopup_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6026
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6027
  "abc_mopup_stage3 (s, l, r) n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6028
          (if s > 0 \<and> s \<le> 2*n then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6029
              if hd r = Bk then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6030
              else 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6031
           else if s = 2*n + 2 then 1 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6032
           else if s = 2*n + 3 then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6033
           else if s = 2*n + 4 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6034
           else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6035
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6036
fun abc_mopup_measure :: "(t_conf \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6037
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6038
  "abc_mopup_measure (c, n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6039
    (abc_mopup_stage1 c n, abc_mopup_stage2 c n, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6040
                                       abc_mopup_stage3 c n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6041
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6042
definition abc_mopup_LE ::
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6043
   "(((nat \<times> block list \<times> block list) \<times> nat) \<times> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6044
    ((nat \<times> block list \<times> block list) \<times> nat)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6045
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6046
  "abc_mopup_LE \<equiv> (inv_image lex_triple abc_mopup_measure)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6047
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6048
lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6049
by(auto intro:wf_inv_image wf_lex_triple simp:abc_mopup_LE_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6050
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6051
lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6052
apply(auto simp: mopup_bef_erase_a.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6053
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6054
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6055
lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6056
apply(auto simp: mopup_bef_erase_b.simps) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6057
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6058
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6059
lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6060
apply(auto simp: mopup_aft_erase_b.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6061
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6062
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6063
lemma mopup_halt_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6064
 "\<lbrakk>n < length lm; mopup_inv (Suc 0, l, r) lm n ires; wf abc_mopup_LE\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6065
 \<Longrightarrow>  \<forall>na. \<not> (\<lambda>(s, l, r) n. s = 0) (t_steps (Suc 0, l, r)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6066
      (mop_bef n @ tshift mp_up (2 * n), 0) na) n \<longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6067
       ((t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6068
        (Suc na), n),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6069
       t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6070
         na, n) \<in> abc_mopup_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6071
apply(rule allI, rule impI, simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6072
apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6073
                     (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6074
apply(case_tac "(t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6075
               (mop_bef n @ tshift mp_up (2 * n), 0) na)",  simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6076
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6077
  fix na a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6078
  assume  "n < length lm" "mopup_inv (a, b, c) lm n ires" "0 < a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6079
  thus "((t_step (a, b, c) (mop_bef n @ tshift mp_up (2 * n), 0), n),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6080
         (a, b, c), n) \<in> abc_mopup_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6081
    apply(auto split:if_splits simp add:t_step.simps mopup_inv.simps,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6082
      tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *})
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6083
    apply(simp_all add: mopupfetchs new_tape.simps abc_mopup_LE_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6084
                   lex_triple_def lex_pair_def )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6085
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6086
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6087
  fix na
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6088
  assume "n < length lm" "mopup_inv (Suc 0, l, r) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6089
  thus "mopup_inv (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6090
       (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6091
    apply(rule mopup_inv_steps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6092
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6093
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6094
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6095
lemma mopup_halt: "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6096
  \<exists> stp. (\<lambda> (s, l, r). s = 0) (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6097
        ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6098
apply(subgoal_tac "mopup_inv (Suc 0, l, r) lm n ires")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6099
apply(insert wf_abc_mopup_le)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6100
apply(insert halt_lemma[of abc_mopup_LE 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6101
    "\<lambda> ((s, l, r), n). s = 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6102
    "\<lambda> stp. (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6103
           , 0) stp, n)"], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6104
apply(insert mopup_halt_pre[of n lm l r], simp, erule exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6105
apply(rule_tac x = na in exI, case_tac "(t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6106
          (mop_bef n @ tshift mp_up (2 * n), 0) na)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6107
apply(rule_tac mopup_init, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6108
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6109
(***End: mopup stop****)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6110
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6111
lemma mopup_stop_cond: "mopup_inv (0, l, r) lm n ires \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6112
                                     (\<exists>ln rn. ?l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ?ires \<and> ?r = <abc_lm_v ?lm ?n> @ Bk\<^bsup>rn\<^esup>) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6113
         t_halt_conf (0, l, r) \<and> t_result r = Suc (abc_lm_v lm n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6114
apply(simp add: mopup_inv.simps mopup_stop.simps t_halt_conf.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6115
                t_result.simps, auto simp: tape_of_nat_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6116
apply(rule_tac x = rn in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6117
      rule_tac x = "Suc (abc_lm_v lm n)" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6118
       simp add: tape_of_nat_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6119
apply(simp add: tape_of_nat_abv  exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6120
apply(subgoal_tac "takeWhile (\<lambda>a. a = Oc) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6121
             (replicate (abc_lm_v lm n) Oc @ replicate rn Bk)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6122
       = replicate (abc_lm_v lm n) Oc @ takeWhile (\<lambda>a. a = Oc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6123
                                          (replicate rn Bk)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6124
apply(case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6125
apply(rule takeWhile_append2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6126
apply(case_tac x, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6127
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6128
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6129
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6130
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6131
lemma mopup_halt_conf_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6132
 "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6133
  \<Longrightarrow> \<exists> na. (\<lambda> (s', l', r').  s' = 0 \<and> mopup_stop (s', l', r') lm n ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6134
      (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6135
            ((mop_bef n @ tshift mp_up (2 * n)), 0) na)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6136
apply(subgoal_tac "\<exists> stp. (\<lambda> (s, l, r). s = 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6137
 (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6138
       erule exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6139
apply(rule_tac x = stp in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6140
      case_tac "(t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6141
          (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6142
apply(subgoal_tac " mopup_inv (Suc 0, l, r) lm n ires")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6143
apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6144
            (mop_bef n @ tshift mp_up (2 * n), 0) stp) lm n ires", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6145
apply(simp only: mopup_inv.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6146
apply(rule_tac mopup_inv_steps, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6147
apply(rule_tac mopup_init, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6148
apply(rule_tac mopup_halt, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6149
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6150
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6151
thm mopup_stop.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6152
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6153
lemma  mopup_halt_conf:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6154
  assumes len: "n < length lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6155
  and correspond: "crsp_l ly (as, lm) (s, l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6156
  shows 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6157
  "\<exists> na. (\<lambda> (s', l', r'). ((\<exists>ln rn. s' = 0 \<and> l' = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r' = Oc\<^bsup>Suc (abc_lm_v lm n)\<^esup> @ Bk\<^bsup>rn\<^esup>)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6158
             (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6159
                  ((mop_bef n @ tshift mp_up (2 * n)), 0) na)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6160
using len correspond mopup_halt_conf_pre[of n lm ly as s l r ires]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6161
apply(simp add: mopup_stop.simps tape_of_nat_abv tape_of_nat_list.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6162
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6163
(*********End: mop_up****************************)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6164
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6165
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6166
subsubsection {* Final results about Abacus machine *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6167
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6168
thm mopup_halt
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6169
lemma mopup_halt_bef: "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6170
    \<Longrightarrow> \<exists>stp. (\<lambda>(s, l, r). s \<noteq> 0 \<and> ((\<lambda> (s', l', r'). s' = 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6171
   (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6172
    (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6173
apply(insert mopup_halt[of n lm ly as s l r ires], simp, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6174
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6175
  fix stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6176
  assume "n < length lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6177
         "crsp_l ly (as, lm) (s, l, r) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6178
         "(\<lambda>(s, l, r). s = 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6179
            (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6180
              (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6181
  thus "\<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6182
   (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6183
    (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6184
  proof(induct stp, simp add: t_steps.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6185
    fix stpa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6186
    assume h1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6187
      "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6188
           (mop_bef n @ tshift mp_up (2 * n), 0) stpa) \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6189
       \<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6190
         (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6191
            (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6192
              (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6193
      and h2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6194
        "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6195
                    (mop_bef n @ tshift mp_up (2 * n), 0) (Suc stpa))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6196
         "n < length lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6197
         "crsp_l ly (as, lm) (s, l, r) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6198
    thus "\<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6199
             (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) (
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6200
                t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6201
                  (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6202
      apply(case_tac "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6203
                     (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6204
            simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6205
      apply(rule_tac x = "stpa" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6206
      apply(case_tac "(t_steps (Suc 0, l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6207
                         (mop_bef n @ tshift mp_up (2 * n), 0) stpa)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6208
            simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6209
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6210
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6211
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6212
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6213
lemma tshift_nth_state0: "\<lbrakk>n < length tp; tp ! n = (a, 0)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6214
    \<Longrightarrow> tshift tp off ! n = (a, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6215
apply(induct n, case_tac tp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6216
apply(auto simp: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6217
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6218
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6219
lemma shift_length: "length (tshift tp n) = length tp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6220
apply(auto simp: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6221
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6222
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6223
lemma even_Suc_le: "\<lbrakk>y mod 2 = 0; 2 * x < y\<rbrakk> \<Longrightarrow> Suc (2 * x) < y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6224
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6225
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6226
lemma [simp]: "(4::nat) * n mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6227
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6228
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6229
lemma tm_append_fetch_equal: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6230
  "\<lbrakk>t_ncorrect (tm_of aprog); s'> 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6231
    fetch (mop_bef n @ tshift mp_up (2 * n)) s' b = (a, 0)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6232
\<Longrightarrow> fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6233
    (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) b 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6234
   = (a, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6235
apply(case_tac s', simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6236
apply(auto simp: fetch.simps nth_of.simps t_ncorrect.simps shift_length nth_append
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6237
                 tshift.simps split: list.splits block.splits split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6238
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6239
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6240
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6241
  "\<lbrakk>t_ncorrect (tm_of aprog);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6242
    t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6243
                                               (0, l'', r''); s' > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6244
  \<Longrightarrow> t_step (s' + length (tm_of aprog) div 2, l', r') 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6245
        (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6246
           (length (tm_of aprog) div 2), 0) = (0, l'', r'')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6247
apply(simp add: t_step.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6248
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6249
   "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6250
              (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6251
  = (fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6252
       (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6253
    (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6254
apply(case_tac "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6255
       (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6256
apply(drule_tac tm_append_fetch_equal, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6257
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6258
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6259
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6260
  "start_of (layout_of aprog) (length aprog) - Suc 0 = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6261
                                      length (tm_of aprog) div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6262
apply(subgoal_tac  "abc2t_correct aprog")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6263
apply(insert pre_lheq[of "concat (take (length aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6264
       (tms_of aprog))" "length aprog" aprog], simp add: tm_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6265
by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6266
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6267
lemma tm_append_stop_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6268
  "\<lbrakk>t_ncorrect (tm_of aprog); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6269
    t_ncorrect (mop_bef n @ tshift mp_up (2 * n)); n < length lm; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6270
   (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6271
                         (s', l', r');
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6272
    s' \<noteq> 0;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6273
    t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6274
                                                     = (0, l'', r'')\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6275
     \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6276
(t_steps ((start_of (layout_of aprog) (length aprog), l, r)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6277
  (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6278
   (start_of (layout_of aprog) (length aprog) - Suc 0), 0) (Suc stp))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6279
  = (0, l'', r'')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6280
apply(insert tm_append_steps_equal[of "tm_of aprog" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6281
      "(mop_bef n @ tshift mp_up (2 * n))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6282
      "(start_of (layout_of aprog) (length aprog) - Suc 0)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6283
      "Suc 0" l r stp], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6284
apply(subgoal_tac "(start_of (layout_of aprog) (length aprog) - Suc 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6285
              = (length (tm_of aprog) div 2)", simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6286
apply(case_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6287
 "(t_steps (start_of (layout_of aprog) (length aprog), l, r) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6288
      (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6289
           (length (tm_of aprog) div 2), 0) stp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6290
apply(subgoal_tac "start_of (layout_of aprog) (length aprog) > 0", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6291
      case_tac "start_of (layout_of aprog) (length aprog)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6292
      simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6293
apply(rule startof_not0, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6294
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6295
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6296
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6297
lemma stop_conf: "mopup_inv (0, aca, bc) am n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6298
    \<Longrightarrow> t_halt_conf (0, aca, bc) \<and> t_result bc = Suc (abc_lm_v am n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6299
apply(case_tac n, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6300
      auto simp: mopup_inv.simps mopup_stop.simps t_halt_conf.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6301
                 t_result.simps tape_of_nl_abv tape_of_nat_abv )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6302
apply(rule_tac x = "rn" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6303
      rule_tac x = "Suc (abc_lm_v am 0)" in exI, simp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6304
apply(subgoal_tac "takeWhile (\<lambda>a. a = Oc) (Oc\<^bsup>abc_lm_v am 0\<^esup> @ Bk\<^bsup>rn\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6305
              = Oc\<^bsup>abc_lm_v am 0\<^esup> @ takeWhile (\<lambda>a. a = Oc) (Bk\<^bsup>rn\<^esup>)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6306
apply(simp add: exponent_def, case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6307
apply(rule_tac takeWhile_append2, simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6308
apply(rule_tac x = rn in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6309
      rule_tac x = "Suc (abc_lm_v am (Suc nat))" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6310
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6311
 "takeWhile (\<lambda>a. a = Oc) (Oc\<^bsup>abc_lm_v am (Suc nat)\<^esup> @ Bk\<^bsup>rn\<^esup>) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6312
       Oc\<^bsup>abc_lm_v am (Suc nat)\<^esup> @ takeWhile (\<lambda>a. a = Oc) (Bk\<^bsup>rn\<^esup>)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6313
apply(simp add: exponent_def, case_tac rn, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6314
apply(rule_tac takeWhile_append2, simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6315
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6316
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6317
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6318
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6319
lemma start_of_out_range: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6320
"as \<ge> length aprog \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6321
   start_of (layout_of aprog) as = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6322
             start_of (layout_of aprog) (length aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6323
apply(induct as, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6324
apply(case_tac "length aprog = Suc as", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6325
apply(simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6326
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6327
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6328
lemma [intro]: "t_ncorrect (tm_of aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6329
apply(simp add: tm_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6330
apply(insert tms_mod2[of "length aprog" aprog], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6331
                                simp add: t_ncorrect.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6332
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6333
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6334
lemma abacus_turing_eq_halt_case_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6335
   "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6336
     tprog = tm_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6337
     crsp_l ly ac tc ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6338
     n < length am;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6339
     abc_steps_l ac aprog stp = (as, am); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6340
     mop_ss = start_of ly (length aprog);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6341
     as \<ge> length aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6342
     \<Longrightarrow> \<exists> stp. (\<lambda> (s, l, r). s = 0 \<and> mopup_inv (0, l, r) am n ires)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6343
                (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6344
apply(insert steps_crsp[of ly aprog tprog ac tc ires "(as, am)" stp], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6345
apply(case_tac "(t_steps tc (tm_of aprog, 0) n')",  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6346
      simp add: tMp.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6347
apply(subgoal_tac "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6348
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6349
  fix n' a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6350
  assume h1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6351
    "crsp_l (layout_of aprog) ac tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6352
    "abc_steps_l ac aprog stp = (as, am)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6353
    "length aprog \<le> as"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6354
    "crsp_l (layout_of aprog) (as, am) (a, b, c) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6355
    "t_steps tc (tm_of aprog, 0) n' = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6356
    "n < length am"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6357
    "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6358
  hence h2:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6359
  "t_steps tc (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6360
    (start_of (layout_of aprog) (length aprog) - Suc 0), 0) n' 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6361
                                    = (a, b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6362
    apply(rule_tac tm_append_steps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6363
    apply(simp add: crsp_l.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6364
    apply(simp add: crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6365
    apply(rule startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6366
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6367
  from h1 have h3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6368
  "\<exists>stp. (\<lambda>(s, l, r). s \<noteq> 0 \<and> ((\<lambda> (s', l', r'). s' = 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6369
           (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6370
         (t_steps (Suc 0, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6371
               (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6372
    apply(rule_tac mopup_halt_bef, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6373
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6374
  from h1 and h2 and h3 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6375
    "\<exists>stp. case t_steps tc (tm_of aprog @ abacus.tshift (mop_bef n @ abacus.tshift mp_up (2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6376
    (start_of (layout_of aprog) (length aprog) - Suc 0), 0) stp of (s, ab)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6377
    \<Rightarrow> s = 0 \<and> mopup_inv (0, ab) am n ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6378
  proof(erule_tac exE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6379
    case_tac "(t_steps (Suc 0, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6380
              (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", simp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6381
    case_tac "(t_step (aa, ba, ca) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6382
              (mop_bef n @ tshift mp_up (2 * n), 0))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6383
    fix stpa aa ba ca aaa baa caa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6384
    assume g1: "0 < aa \<and> aaa = 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6385
      "t_steps (Suc 0, b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6386
      (mop_bef n @ tshift mp_up (2 * n), 0) stpa = (aa, ba,ca)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6387
      "t_step (aa, ba, ca) (mop_bef n @ tshift mp_up (2 * n), 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6388
      = (0, baa, caa)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6389
    from h1 and this have g2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6390
      "t_steps (start_of (layout_of aprog) (length aprog), b, c) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6391
         (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6392
           (start_of (layout_of aprog) (length aprog) - Suc 0), 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6393
                (Suc stpa) = (0, baa, caa)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6394
      apply(rule_tac tm_append_stop_step, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6395
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6396
    from h1 and h2 and g1 and this show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6397
      apply(rule_tac x = "n' + Suc stpa" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6398
      apply(simp add: t_steps_ind del: t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6399
      apply(subgoal_tac "a = start_of (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6400
                                          (length aprog)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6401
      apply(insert mopup_inv_steps[of n am "Suc 0" b c ires "Suc stpa"],
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6402
            simp add: t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6403
      apply(subgoal_tac "mopup_inv (Suc 0, b, c) am n ires", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6404
      apply(rule_tac mopup_init, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6405
      apply(simp add: crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6406
      apply(erule_tac start_of_out_range)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6407
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6408
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6409
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6410
  show " t_ncorrect (mop_bef n @ tshift mp_up (2 * n))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6411
    apply(auto simp: t_ncorrect.simps tshift.simps mp_up_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6412
    done   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6413
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6414
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6415
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6416
  One of the main theorems about Abacus compilation.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6417
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6418
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6419
lemma abacus_turing_eq_halt_case: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6420
  assumes layout: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6421
  -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6422
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6423
  and complied: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6424
  -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6425
  "tprog = tm_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6426
  and correspond: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6427
  -- {* TM configuration @{text "tc"} and Abacus configuration @{text "ac"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6428
  are in correspondence: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6429
  "crsp_l ly ac tc ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6430
  and halt_state: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6431
  -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6432
  if Abacus is in such a state, it is in halt state: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6433
  "as \<ge> length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6434
  and abc_exec: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6435
  -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6436
  reaches such a halt state: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6437
  "abc_steps_l ac aprog stp = (as, am)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6438
  and rs_len: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6439
  -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6440
  "n < length am"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6441
  and mopup_start:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6442
  -- {* The startling label for mopup mahines, according to the layout and Abacus program 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6443
   should be @{text "mop_ss"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6444
  "mop_ss = start_of ly (length aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6445
  shows 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6446
  -- {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6447
  After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6448
  TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6449
  only hold the content of memory cell @{text "n"}:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6450
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6451
  "\<exists> stp. (\<lambda> (s, l, r). \<exists> ln rn. s = 0 \<and>  l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>rn\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6452
           (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6453
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6454
  from layout complied correspond halt_state abc_exec rs_len mopup_start
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6455
       and abacus_turing_eq_halt_case_pre [of ly aprog tprog ac tc ires n am stp as mop_ss]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6456
  show "?thesis" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6457
    apply(simp add: mopup_inv.simps mopup_stop.simps tape_of_nat_abv)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6458
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6459
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6460
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6461
lemma abc_unhalt_case_zero: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6462
"\<lbrakk>crsp_l (layout_of aprog) ac tc ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6463
  n < length am; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6464
  \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6465
 \<Longrightarrow> \<exists>astp bstp. 0 \<le> bstp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6466
          crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6467
                (t_steps tc (tm_of aprog, 0) bstp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6468
apply(rule_tac x = "Suc 0" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6469
apply(case_tac " abc_steps_l ac aprog (Suc 0)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6470
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6471
  fix a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6472
  assume "crsp_l (layout_of aprog) ac tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6473
         "abc_steps_l ac aprog (Suc 0) = (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6474
  thus "\<exists>bstp. crsp_l (layout_of aprog) (a, b) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6475
               (t_steps tc (tm_of aprog, 0) bstp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6476
    apply(insert steps_crsp[of "layout_of aprog" aprog 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6477
                  "tm_of aprog" ac tc ires "(a, b)" "Suc 0"], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6478
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6479
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6480
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6481
declare abc_steps_l.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6482
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6483
lemma abc_steps_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6484
 "let (as, am) = abc_steps_l ac aprog stp in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6485
   abc_steps_l ac aprog (Suc stp) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6486
              abc_step_l (as, am) (abc_fetch as aprog) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6487
proof(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6488
  show "(\<lambda>(as, am). abc_steps_l ac aprog (Suc stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6489
        abc_step_l (as, am) (abc_fetch as aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6490
              (abc_steps_l ac aprog stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6491
  proof(induct stp arbitrary: ac)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6492
    fix ac
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6493
    show "(\<lambda>(as, am). abc_steps_l ac aprog (Suc 0) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6494
            abc_step_l (as, am) (abc_fetch as aprog))  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6495
                    (abc_steps_l ac aprog 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6496
      apply(case_tac "ac:: nat \<times> nat list", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6497
            simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6498
      apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6499
            simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6500
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6501
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6502
    fix stp ac
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6503
    assume h1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6504
      "(\<And>ac. (\<lambda>(as, am). abc_steps_l ac aprog (Suc stp) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6505
                            abc_step_l (as, am) (abc_fetch as aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6506
             (abc_steps_l ac aprog stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6507
    thus 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6508
      "(\<lambda>(as, am). abc_steps_l ac aprog (Suc (Suc stp)) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6509
              abc_step_l (as, am) (abc_fetch as aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6510
                             (abc_steps_l ac aprog (Suc stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6511
      apply(case_tac "ac::nat \<times> nat list", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6512
      apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6513
           "abc_steps_l (a, b) aprog (Suc (Suc stp)) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6514
            abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6515
                                              aprog (Suc stp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6516
      apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6517
    proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6518
      fix a b aa ba
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6519
      assume h2: "abc_step_l (a, b) (abc_fetch a aprog) = (aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6520
      from h1 and h2  show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6521
      "(\<lambda>(as, am). abc_steps_l (aa, ba) aprog (Suc stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6522
          abc_step_l (as, am) (abc_fetch as aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6523
                    (abc_steps_l (a, b) aprog (Suc stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6524
	apply(insert h1[of "(aa, ba)"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6525
	apply(simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6526
	apply(insert h2, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6527
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6528
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6529
      fix a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6530
      show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6531
        "abc_steps_l (a, b) aprog (Suc (Suc stp)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6532
         abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6533
                                                   aprog (Suc stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6534
	apply(simp only: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6535
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6536
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6537
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6538
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6539
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6540
lemma abc_unhalt_case_induct: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6541
  "\<lbrakk>crsp_l (layout_of aprog) ac tc ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6542
    n < length am; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6543
    \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6544
    stp \<le> bstp;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6545
    crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6546
                           (t_steps tc (tm_of aprog, 0) bstp) ires\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6547
 \<Longrightarrow> \<exists>astp bstp. Suc stp \<le> bstp \<and> crsp_l (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6548
       (abc_steps_l ac aprog astp) (t_steps tc (tm_of aprog, 0) bstp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6549
apply(rule_tac x = "Suc astp" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6550
apply(case_tac "abc_steps_l ac aprog astp")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6551
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6552
  fix a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6553
  assume 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6554
    "\<forall>stp. (\<lambda>(as, am). as < length aprog)  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6555
                 (abc_steps_l ac aprog stp)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6556
    "stp \<le> bstp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6557
    "crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6558
      (t_steps tc (tm_of aprog, 0) bstp) ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6559
    "abc_steps_l ac aprog astp = (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6560
  thus 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6561
 "\<exists>bstp\<ge>Suc stp. crsp_l (layout_of aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6562
       (abc_steps_l ac aprog (Suc astp)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6563
   (t_steps tc (tm_of aprog, 0) bstp) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6564
    apply(insert crsp_inside[of "layout_of aprog" aprog 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6565
      "tm_of aprog" a b "(t_steps tc (tm_of aprog, 0) bstp)" "ires"], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6566
    apply(erule_tac x = astp in allE, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6567
    apply(rule_tac x = "bstp + stpa" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6568
    apply(insert abc_steps_ind[of ac aprog "astp"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6569
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6570
qed   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6571
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6572
lemma abc_unhalt_case: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6573
  "\<lbrakk>crsp_l (layout_of aprog) ac tc ires;  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6574
    \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6575
 \<Longrightarrow> (\<exists> astp bstp. bstp \<ge> stp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6576
         crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6577
                                (t_steps tc (tm_of aprog, 0) bstp) ires)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6578
apply(induct stp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6579
apply(rule_tac abc_unhalt_case_zero, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6580
apply(rule_tac abc_unhalt_case_induct, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6581
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6582
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6583
lemma abacus_turing_eq_unhalt_case_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6584
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6585
    tprog = tm_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6586
    crsp_l ly ac tc ires;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6587
    \<forall> stp. ((\<lambda> (as, am). as < length aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6588
                       (abc_steps_l ac aprog stp));
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6589
    mop_ss = start_of ly (length aprog)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6590
  \<Longrightarrow> (\<not> (\<exists> stp. (\<lambda> (s, l, r). s = 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6591
              (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6592
  apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6593
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6594
  fix stp a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6595
  assume h1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6596
    "crsp_l (layout_of aprog) ac tc ires" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6597
    "\<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6598
    "t_steps tc (tm_of aprog @ tMp n (start_of (layout_of aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6599
    (length aprog) - Suc 0), 0) stp = (0, a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6600
  thus "False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6601
  proof(insert abc_unhalt_case[of aprog ac tc ires stp], auto, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6602
        case_tac "(abc_steps_l ac aprog astp)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6603
        case_tac "(t_steps tc (tm_of aprog, 0) bstp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6604
    fix astp bstp aa ba aaa baa c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6605
    assume h2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6606
      "abc_steps_l ac aprog astp = (aa, ba)" "stp \<le> bstp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6607
      "t_steps tc (tm_of aprog, 0) bstp = (aaa, baa, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6608
      "crsp_l (layout_of aprog) (aa, ba) (aaa, baa, c) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6609
    hence h3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6610
      "t_steps tc (tm_of aprog @ tMp n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6611
       (start_of (layout_of aprog) (length aprog) - Suc 0), 0) bstp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6612
                    = (aaa, baa, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6613
      apply(intro tm_append_steps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6614
      apply(simp add: crsp_l.simps, rule startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6615
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6616
    from h2 have h4: "\<exists> diff. bstp = stp + diff"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6617
      apply(rule_tac x = "bstp - stp" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6618
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6619
    from h4 and h3 and h2  and h1 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6620
      apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6621
      apply(simp add: state0_ind crsp_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6622
      apply(subgoal_tac "start_of (layout_of aprog) aa > 0", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6623
      apply(rule startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6624
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6625
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6626
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6627
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6628
lemma abacus_turing_eq_unhalt_case:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6629
  assumes layout: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6630
  -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6631
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6632
  and compiled: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6633
  -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6634
  "tprog = tm_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6635
  and correspond: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6636
  -- {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6637
  TM configuration @{text "tc"} and Abacus configuration @{text "ac"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6638
  are in correspondence: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6639
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6640
  "crsp_l ly ac tc ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6641
  and abc_unhalt: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6642
  -- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6643
  If, no matter how many steps the Abacus program @{text "aprog"} executes, it
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6644
  may never reach a halt state. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6645
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6646
  "\<forall> stp. ((\<lambda> (as, am). as < length aprog)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6647
                       (abc_steps_l ac aprog stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6648
  and mopup_start: "mop_ss = start_of ly (length aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6649
  shows
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6650
  -- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6651
  The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6652
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6653
  "\<not> (\<exists> stp. (\<lambda> (s, l, r). s = 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6654
              (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6655
  using layout compiled correspond abc_unhalt mopup_start
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6656
  apply(rule_tac abacus_turing_eq_unhalt_case_pre, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6657
  done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6658
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6659
definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6660
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6661
  "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<^bsup>n\<^esup> \<or> ys = xs @ 0\<^bsup>n\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6662
lemma [intro]: "abc_list_crsp (lm @ 0\<^bsup>m\<^esup>) lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6663
apply(auto simp: abc_list_crsp_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6664
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6665
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6666
thm abc_lm_v.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6667
lemma abc_list_crsp_lm_v: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6668
  "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6669
apply(auto simp: abc_list_crsp_def abc_lm_v.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6670
                 nth_append exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6671
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6672
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6673
lemma  rep_app_cons_iff: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6674
  "k < n \<Longrightarrow> replicate n a[k:=b] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6675
          replicate k a @ b # replicate (n - k - 1) a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6676
apply(induct n arbitrary: k, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6677
apply(simp split:nat.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6678
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6679
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6680
lemma abc_list_crsp_lm_s: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6681
  "abc_list_crsp lma lmb \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6682
      abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6683
apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6684
apply(simp_all add: list_update_append, auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6685
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6686
  fix na
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6687
  assume h: "m < length lmb + na" " \<not> m < length lmb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6688
  hence "m - length lmb < na" by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6689
  hence "replicate na 0[(m- length lmb):= n] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6690
           replicate (m - length lmb) 0 @ n # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6691
              replicate (na - (m - length lmb) - 1) 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6692
    apply(erule_tac rep_app_cons_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6693
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6694
  thus "\<exists>nb. replicate na 0[m - length lmb := n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6695
                 replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6696
                 replicate (m - length lmb) 0 @ [n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6697
                 replicate na 0[m - length lmb := n] @ replicate nb 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6698
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6699
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6700
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6701
  fix na
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6702
  assume h: "\<not> m < length lmb + na"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6703
  show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6704
    "\<exists>nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6705
           replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6706
          replicate (m - length lmb) 0 @ [n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6707
            replicate na 0 @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6708
            replicate (m - (length lmb + na)) 0 @ n # replicate nb 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6709
    apply(rule_tac x = 0 in exI, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6710
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6711
    apply(simp add: replicate_add[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6712
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6713
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6714
  fix na
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6715
  assume h: "\<not> m < length lma" "m < length lma + na"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6716
  hence "m - length lma < na" by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6717
  hence 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6718
    "replicate na 0[(m- length lma):= n] = replicate (m - length lma) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6719
                  0 @ n # replicate (na - (m - length lma) - 1) 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6720
    apply(erule_tac rep_app_cons_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6721
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6722
  thus "\<exists>nb. replicate (m - length lma) 0 @ [n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6723
                 replicate na 0[m - length lma := n] @ replicate nb 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6724
           \<or> replicate na 0[m - length lma := n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6725
                 replicate (m - length lma) 0 @ n # replicate nb 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6726
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6727
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6728
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6729
  fix na
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6730
  assume "\<not> m < length lma + na"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6731
  thus " \<exists>nb. replicate (m - length lma) 0 @ [n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6732
            replicate na 0 @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6733
            replicate (m - (length lma + na)) 0 @ n # replicate nb 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6734
        \<or>   replicate na 0 @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6735
               replicate (m - (length lma + na)) 0 @ [n] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6736
            replicate (m - length lma) 0 @ n # replicate nb 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6737
    apply(rule_tac x = 0 in exI, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6738
    apply(simp add: replicate_add[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6739
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6740
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6741
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6742
lemma abc_list_crsp_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6743
  "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6744
    abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6745
    \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6746
apply(case_tac i, auto simp: abc_step_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6747
       abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6748
                       split: abc_inst.splits if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6749
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6750
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6751
thm abc_step_l.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6752
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6753
lemma abc_steps_red: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6754
  "abc_steps_l ac aprog stp = (as, am) \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6755
     abc_steps_l ac aprog (Suc stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6756
           abc_step_l (as, am) (abc_fetch as aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6757
using abc_steps_ind[of ac aprog stp]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6758
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6759
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6760
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6761
lemma abc_list_crsp_steps: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6762
  "\<lbrakk>abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (a, lm'); aprog \<noteq> []\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6763
      \<Longrightarrow> \<exists> lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6764
                                          abc_list_crsp lm' lma"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6765
apply(induct stp arbitrary: a lm', simp add: abc_steps_l.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6766
apply(case_tac "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6767
      simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6768
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6769
  fix stp a lm' aa b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6770
  assume ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6771
    "\<And>a lm'. aa = a \<and> b = lm' \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6772
     \<exists>lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6773
                                          abc_list_crsp lm' lma"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6774
    and h: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) = (a, lm')" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6775
           "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (aa, b)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6776
           "aprog \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6777
  hence g1: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6778
          = abc_step_l (aa, b) (abc_fetch aa aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6779
    apply(rule_tac abc_steps_red, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6780
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6781
  have "\<exists>lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6782
              abc_list_crsp b lma"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6783
    apply(rule_tac ind, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6784
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6785
  from this obtain lma where g2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6786
    "abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6787
     abc_list_crsp b lma"   ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6788
  hence g3: "abc_steps_l (0, lm) aprog (Suc stp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6789
          = abc_step_l (aa, lma) (abc_fetch aa aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6790
    apply(rule_tac abc_steps_red, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6791
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6792
  show "\<exists>lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \<and>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6793
              abc_list_crsp lm' lma"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6794
    using g1 g2 g3 h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6795
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6796
    apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6797
          case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6798
    apply(rule_tac abc_list_crsp_step, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6799
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6800
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6801
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6802
text {* Begin: equvilence between steps and t_steps*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6803
lemma [simp]: "(case ca of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6804
                (case ca of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6805
by(case_tac ca, simp_all, case_tac a, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6806
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6807
text {* needed to interpret*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6808
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6809
lemma steps_eq: "length t mod 2 = 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6810
                    t_steps c (t, 0) stp = steps c t stp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6811
apply(induct stp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6812
apply(simp add: steps.simps t_steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6813
apply(simp add:tstep_red t_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6814
apply(case_tac "steps c t stp", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6815
apply(auto simp: t_step.simps tstep.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6816
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6817
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6818
text{* end: equvilence between steps and t_steps*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6819
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6820
lemma crsp_l_start: "crsp_l ly (0, lm) (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) ires"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6821
apply(simp add: crsp_l.simps, auto simp: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6822
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6823
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6824
lemma t_ncorrect_app: "\<lbrakk>t_ncorrect t1; t_ncorrect t2\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6825
                                          t_ncorrect (t1 @ t2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6826
apply(simp add: t_ncorrect.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6827
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6828
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6829
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6830
  "(length (tm_of aprog) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6831
    length (tMp n (start_of ly (length aprog) - Suc 0))) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6832
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6833
 "t_ncorrect (tm_of aprog @ tMp n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6834
             (start_of ly (length aprog) - Suc 0))")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6835
apply(simp add: t_ncorrect.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6836
apply(rule_tac t_ncorrect_app, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6837
      auto simp: tMp.simps t_ncorrect.simps tshift.simps mp_up_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6838
apply(subgoal_tac
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6839
       "t_ncorrect (tm_of aprog)", simp add: t_ncorrect.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6840
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6841
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6842
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6843
lemma [simp]: "takeWhile (\<lambda>a. a = Oc) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6844
              (replicate rs Oc @ replicate rn Bk) = replicate rs Oc"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6845
apply(induct rs, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6846
apply(induct rn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6847
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6848
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6849
lemma abacus_turing_eq_halt': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6850
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6851
    tprog = tm_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6852
    n < length am;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6853
    abc_steps_l (0, lm) aprog stp = (as, am); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6854
    mop_ss = start_of ly (length aprog);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6855
    as \<ge> length aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6856
    \<Longrightarrow> \<exists> stp m l. steps (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6857
                (tprog @ (tMp n (mop_ss - 1))) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6858
                  = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6859
apply(drule_tac tc = "(Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>)" in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6860
               abacus_turing_eq_halt_case, auto intro: crsp_l_start)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6861
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6862
         "length (tm_of aprog @ tMp n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6863
                  (start_of ly (length aprog) - Suc 0)) mod 2 = 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6864
apply(simp add: steps_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6865
apply(rule_tac x = stpa in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6866
       simp add:  exponent_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6867
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6868
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6869
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6870
thm tinres_steps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6871
lemma list_length: "xs = ys \<Longrightarrow> length xs = length ys"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6872
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6873
lemma [elim]: "tinres (Bk\<^bsup>m\<^esup>) b \<Longrightarrow> \<exists>m. b = Bk\<^bsup>m\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6874
apply(auto simp: tinres_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6875
apply(rule_tac x = "m-n" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6876
             auto simp: exponent_def replicate_add[THEN sym]) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6877
apply(case_tac "m < n", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6878
apply(drule_tac list_length, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6879
apply(subgoal_tac "\<exists> d. m = d + n", auto simp: replicate_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6880
apply(rule_tac x = "m - n" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6881
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6882
lemma [intro]: "tinres [Bk] (Bk\<^bsup>k\<^esup>) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6883
apply(auto simp: tinres_def exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6884
apply(case_tac k, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6885
apply(rule_tac x = "Suc 0" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6886
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6887
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6888
lemma abacus_turing_eq_halt_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6889
 "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6890
   tprog = tm_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6891
   n < length am;     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6892
   abc_steps_l (0, lm) aprog stp = (as, am);  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6893
   mop_ss = start_of ly (length aprog);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6894
   as \<ge> length aprog\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6895
  \<Longrightarrow> \<exists> stp m l. steps  (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6896
               (tprog @ (tMp n (mop_ss - 1))) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6897
                 = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6898
using abacus_turing_eq_halt'
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6899
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6900
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6901
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6902
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6903
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6904
  Main theorem for the case when the original Abacus program does halt.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6905
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6906
lemma abacus_turing_eq_halt: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6907
  assumes layout:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6908
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6909
  -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6910
  and compiled: "tprog = tm_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6911
  -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6912
  and halt_state: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6913
   -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6914
  if Abacus is in such a state, it is in halt state: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6915
  "as \<ge> length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6916
  and abc_exec: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6917
  -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6918
  reaches such a halt state: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6919
  "abc_steps_l (0, lm) aprog stp = (as, am)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6920
  and rs_locate: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6921
   -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6922
  "n < length am"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6923
  and mopup_start: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6924
   -- {* The startling label for mopup mahines, according to the layout and Abacus program 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6925
   should be @{text "mop_ss"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6926
  "mop_ss = start_of ly (length aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6927
  shows 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6928
  -- {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6929
  After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6930
  TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6931
  only hold the content of memory cell @{text "n"}:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6932
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6933
  "\<exists> stp m l. steps (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) (tprog @ (tMp n (mop_ss - 1))) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6934
                      = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6935
  using layout compiled halt_state abc_exec rs_locate mopup_start
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6936
  by(rule_tac abacus_turing_eq_halt_pre, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6937
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6938
lemma abacus_turing_eq_uhalt': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6939
 "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6940
   tprog = tm_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6941
   \<forall> stp. ((\<lambda> (as, am). as < length aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6942
                   (abc_steps_l (0, lm) aprog stp));
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6943
   mop_ss = start_of ly (length aprog)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6944
  \<Longrightarrow> (\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6945
                      (tprog @ (tMp n (mop_ss - 1))) stp)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6946
apply(drule_tac tc = "(Suc 0, [Bk, Bk], <lm>)" and n = n and ires = "[]" in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6947
         abacus_turing_eq_unhalt_case, auto intro: crsp_l_start)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6948
apply(simp add: crsp_l.simps start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6949
apply(erule_tac x = stp in allE, erule_tac x = stp in allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6950
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6951
   "length (tm_of aprog @ tMp n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6952
         (start_of ly (length aprog) - Suc 0)) mod 2 = 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6953
apply(simp add: steps_eq, auto simp: isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6954
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6955
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6956
lemma abacus_turing_eq_uhalt_pre: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6957
  "\<lbrakk>ly = layout_of aprog; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6958
    tprog = tm_of aprog;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6959
    \<forall> stp. ((\<lambda> (as, am). as < length aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6960
                      (abc_steps_l (0, lm) aprog stp));
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6961
    mop_ss = start_of ly (length aprog)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6962
  \<Longrightarrow> (\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6963
                    (tprog @ (tMp n (mop_ss - 1))) stp)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6964
apply(drule_tac k = 0 and n = n  in abacus_turing_eq_uhalt', auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6965
apply(erule_tac x = stp in allE, erule_tac x = stp in allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6966
apply(subgoal_tac "tinres ([Bk]) (Bk\<^bsup>k\<^esup>)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6967
apply(case_tac "steps (Suc 0, Bk\<^bsup>k\<^esup>, <lm>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6968
      (tm_of aprog @ tMp n (start_of ly (length aprog) - Suc 0)) stp")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6969
apply(case_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6970
  "steps (Suc 0, [Bk], <lm>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6971
    (tm_of aprog @ tMp n (start_of ly (length aprog) - Suc 0)) stp")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6972
apply(drule_tac tinres_steps, auto simp: isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6973
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6974
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6975
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6976
  Main theorem for the case when the original Abacus program does not halt.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6977
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6978
lemma abacus_turing_eq_uhalt:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6979
  assumes layout: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6980
  -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6981
  "ly = layout_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6982
  and compiled:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6983
   -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6984
  "tprog = tm_of aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6985
  and abc_unhalt:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6986
  -- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6987
  If, no matter how many steps the Abacus program @{text "aprog"} executes, it
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6988
  may never reach a halt state. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6989
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6990
  "\<forall> stp. ((\<lambda> (as, am). as < length aprog) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6991
                      (abc_steps_l (0, lm) aprog stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6992
  and mop_start: "mop_ss = start_of ly (length aprog)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6993
  shows 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6994
   -- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6995
  The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6996
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6997
  "\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6998
                    (tprog @ (tMp n (mop_ss - 1))) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  6999
  using abacus_turing_eq_uhalt'
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  7000
        layout compiled abc_unhalt mop_start
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  7001
  by(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  7002
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  7003
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  7004
end
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  7005