utm/recursive.thy
author zhang
Mon, 15 Oct 2012 13:23:52 +0000
changeset 371 48b231495281
parent 370 1ce04eb1c8ad
permissions -rw-r--r--
Some illustration added together with more explanations.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
370
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     1
theory recursive
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     2
imports Main rec_def abacus
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     3
begin
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     5
section {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     6
  Compiling from recursive functions to Abacus machines
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     7
  *}
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
  Some auxilliary Abacus machines used to construct the result Abacus machines.
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
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    14
  @{text "get_paras_num recf"} returns the arity of recursive function @{text "recf"}.
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
fun get_paras_num :: "recf \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    17
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    18
  "get_paras_num z = 1" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    19
  "get_paras_num s = 1" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    20
  "get_paras_num (id m n) = m" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    21
  "get_paras_num (Cn n f gs) = n" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    22
  "get_paras_num (Pr n f g) = Suc n"  |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    23
  "get_paras_num (Mn n f) = n"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    24
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    25
fun addition :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    26
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    27
  "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    28
                       Inc m, Goto 4]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    29
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    30
fun empty :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    31
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    32
  "empty m n = [Dec m 3, Inc n, Goto 0]"
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
fun abc_inst_shift :: "abc_inst \<Rightarrow> nat \<Rightarrow> abc_inst"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    35
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    36
  "abc_inst_shift (Inc m) n = Inc m" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    37
  "abc_inst_shift (Dec m e) n = Dec m (e + n)" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    38
  "abc_inst_shift (Goto m) n = Goto (m + n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    39
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    40
fun abc_shift :: "abc_inst list \<Rightarrow> nat \<Rightarrow> abc_inst list" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    41
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    42
  "abc_shift xs n = map (\<lambda> x. abc_inst_shift x n) xs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    43
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    44
fun abc_append :: "abc_inst list \<Rightarrow> abc_inst list \<Rightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    45
                           abc_inst list" (infixl "[+]" 60)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    46
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    47
  "abc_append al bl = (let al_len = length al in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    48
                           al @ abc_shift bl al_len)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    49
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    50
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    51
  The compilation of @{text "z"}-operator.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    52
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    53
definition rec_ci_z :: "abc_inst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    54
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    55
  "rec_ci_z \<equiv> [Goto 1]"
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
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    58
  The compilation of @{text "s"}-operator.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    59
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    60
definition rec_ci_s :: "abc_inst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    61
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    62
  "rec_ci_s \<equiv> (addition 0 1 2 [+] [Inc 1])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    63
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
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    66
  The compilation of @{text "id i j"}-operator
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    67
*}
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
fun rec_ci_id :: "nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    70
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    71
  "rec_ci_id i j = addition j i (i + 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    72
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
fun mv_boxes :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    75
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    76
  "mv_boxes ab bb 0 = []" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    77
  "mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] empty (ab + n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    78
  (bb + n)"
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
fun empty_boxes :: "nat \<Rightarrow> abc_inst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    81
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    82
  "empty_boxes 0 = []" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    83
  "empty_boxes (Suc n) = empty_boxes n [+] [Dec n 2, Goto 0]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    84
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    85
fun cn_merge_gs ::
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    86
  "(abc_inst list \<times> nat \<times> nat) list \<Rightarrow> nat \<Rightarrow> abc_inst list"
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
  "cn_merge_gs [] p = []" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    89
  "cn_merge_gs (g # gs) p = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    90
      (let (gprog, gpara, gn) = g in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    91
         gprog [+] empty gpara p [+] cn_merge_gs gs (Suc p))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    92
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    93
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    94
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    95
  The compiler of recursive functions, where @{text "rec_ci recf"} return 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    96
  @{text "(ap, arity, fp)"}, where @{text "ap"} is the Abacus program, @{text "arity"} is the 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    97
  arity of the recursive function @{text "recf"}, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    98
@{text "fp"} is the amount of memory which is going to be
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    99
  used by @{text "ap"} for its execution. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   100
*}
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
function rec_ci :: "recf \<Rightarrow> abc_inst list \<times> nat \<times> nat"
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
  "rec_ci z = (rec_ci_z, 1, 2)" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   105
  "rec_ci s = (rec_ci_s, 1, 3)" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   106
  "rec_ci (id m n) = (rec_ci_id m n, m, m + 2)" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   107
  "rec_ci (Cn n f gs) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   108
      (let cied_gs = map (\<lambda> g. rec_ci g) (f # gs) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   109
       let (fprog, fpara, fn) = hd cied_gs in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   110
       let pstr = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   111
        Max (set (Suc n # fn # (map (\<lambda> (aprog, p, n). n) cied_gs))) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   112
       let qstr = pstr + Suc (length gs) in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   113
       (cn_merge_gs (tl cied_gs) pstr [+] mv_boxes 0 qstr n [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   114
          mv_boxes pstr 0 (length gs) [+] fprog [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   115
            empty fpara pstr [+] empty_boxes (length gs) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   116
             empty pstr n [+] mv_boxes qstr 0 n, n,  qstr + n))" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   117
  "rec_ci (Pr n f g) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   118
         (let (fprog, fpara, fn) = rec_ci f in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   119
          let (gprog, gpara, gn) = rec_ci g in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   120
          let p = Max (set ([n + 3, fn, gn])) in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   121
          let e = length gprog + 7 in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   122
           (empty n p [+] fprog [+] empty n (Suc n) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   123
               (([Dec p e] [+] gprog [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   124
                 [Inc n, Dec (Suc n) 3, Goto 1]) @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   125
                     [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gprog + 4)]),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   126
             Suc n, p + 1))" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   127
  "rec_ci (Mn n f) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   128
         (let (fprog, fpara, fn) = rec_ci f in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   129
          let len = length (fprog) in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   130
            (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   131
             Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn) )"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   132
  by pat_completeness auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   133
termination 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   134
proof
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   135
term size
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   136
  show "wf (measure size)" by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   137
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   138
  fix n f gs x
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   139
  assume "(x::recf) \<in> set (f # gs)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   140
  thus "(x, Cn n f gs) \<in> measure size"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   141
    by(induct gs, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   142
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   143
  fix n f g
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   144
  show "(f, Pr n f g) \<in> measure size" by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   145
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   146
  fix n f g x xa y xb ya
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   147
  show "(g, Pr n f g) \<in> measure size" by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   148
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   149
  fix n f
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   150
  show "(f, Mn n f) \<in> measure size" by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   151
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   152
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   153
declare rec_ci.simps [simp del] rec_ci_s_def[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   154
        rec_ci_z_def[simp del] rec_ci_id.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   155
        mv_boxes.simps[simp del] abc_append.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   156
        empty.simps[simp del] addition.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   157
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   158
thm rec_calc_rel.induct
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   159
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   160
declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   161
        abc_step_l.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   162
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   163
lemma abc_steps_add: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   164
  "abc_steps_l (as, lm) ap (m + n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   165
         abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   166
apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   167
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   168
  fix m n as lm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   169
  assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   170
    "\<And>n as lm. abc_steps_l (as, lm) ap (m + n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   171
                   abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   172
  show "abc_steps_l (as, lm) ap (Suc m + n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   173
             abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   174
    apply(insert ind[of as lm "Suc n"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   175
    apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   176
    apply(case_tac "(abc_steps_l (as, lm) ap m)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   177
    apply(simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   178
    apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   179
          simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   180
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   181
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   182
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   183
(*lemmas: rec_ci and rec_calc_rel*)
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
lemma rec_calc_inj_case_z: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   186
  "\<lbrakk>rec_calc_rel z l x; rec_calc_rel z l y\<rbrakk> \<Longrightarrow> x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   187
apply(auto elim: calc_z_reverse)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   188
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   189
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   190
lemma  rec_calc_inj_case_s: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   191
  "\<lbrakk>rec_calc_rel s l x; rec_calc_rel s l y\<rbrakk> \<Longrightarrow> x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   192
apply(auto elim: calc_s_reverse)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   193
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   194
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   195
lemma rec_calc_inj_case_id:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   196
  "\<lbrakk>rec_calc_rel (recf.id nat1 nat2) l x;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   197
    rec_calc_rel (recf.id nat1 nat2) l y\<rbrakk> \<Longrightarrow> x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   198
apply(auto elim: calc_id_reverse)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   199
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   200
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   201
lemma rec_calc_inj_case_mn:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   202
  assumes ind: "\<And> l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   203
           \<Longrightarrow> x = y" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   204
  and h: "rec_calc_rel (Mn n f) l x" "rec_calc_rel (Mn n f) l y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   205
  shows "x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   206
  apply(insert h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   207
  apply(elim  calc_mn_reverse)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   208
  apply(case_tac "x > y", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   209
  apply(erule_tac x = "y" in allE, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   210
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   211
  fix v va
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   212
  assume "rec_calc_rel f (l @ [y]) 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   213
    "rec_calc_rel f (l @ [y]) v"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   214
    "0 < v"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   215
  thus "False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   216
    apply(insert ind[of "l @ [y]" 0 v], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   217
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   218
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   219
  fix v va
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   220
  assume 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   221
    "rec_calc_rel f (l @ [x]) 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   222
    "\<forall>x<y. \<exists>v. rec_calc_rel f (l @ [x]) v \<and> 0 < v" "\<not> y < x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   223
  thus "x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   224
    apply(erule_tac x = "x" in allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   225
    apply(case_tac "x = y", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   226
    apply(drule_tac y = v in ind, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   227
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   228
qed 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   229
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   230
lemma rec_calc_inj_case_pr: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   231
  assumes f_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   232
  "\<And>l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   233
  and g_ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   234
  "\<And>x xa y xb ya l xc yb. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   235
  \<lbrakk>x = rec_ci f; (xa, y) = x; (xb, ya) = y; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   236
  rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk> \<Longrightarrow> xc = yb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   237
  and h: "rec_calc_rel (Pr n f g) l x" "rec_calc_rel (Pr n f g) l y"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   238
  shows "x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   239
  apply(case_tac "rec_ci f")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   240
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   241
  fix a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   242
  assume "rec_ci f = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   243
  hence ng_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   244
    "\<And> l xc yb. \<lbrakk>rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   245
    \<Longrightarrow> xc = yb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   246
    apply(insert g_ind[of "(a, b, c)" "a" "(b, c)" b c], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   247
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   248
  from h show "x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   249
    apply(erule_tac calc_pr_reverse, erule_tac calc_pr_reverse)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   250
    apply(erule f_ind, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   251
    apply(erule_tac calc_pr_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   252
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   253
    fix la ya ry laa yaa rya
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   254
    assume k1:  "rec_calc_rel g (la @ [ya, ry]) x" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   255
      "rec_calc_rel g (la @ [ya, rya]) y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   256
      and k2: "rec_calc_rel (Pr (length la) f g) (la @ [ya]) ry"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   257
              "rec_calc_rel (Pr (length la) f g) (la @ [ya]) rya"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   258
    from k2 have "ry = rya"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   259
      apply(induct ya arbitrary: ry rya)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   260
      apply(erule_tac calc_pr_reverse, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   261
        erule_tac calc_pr_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   262
      apply(erule f_ind, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   263
      apply(erule_tac calc_pr_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   264
      apply(erule_tac rSucy = rya in calc_pr_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   265
    proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   266
      fix ya ry rya l y ryb laa yb ryc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   267
      assume ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   268
        "\<And>ry rya. \<lbrakk>rec_calc_rel (Pr (length l) f g) (l @ [y]) ry; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   269
                   rec_calc_rel (Pr (length l) f g) (l @ [y]) rya\<rbrakk> \<Longrightarrow> ry = rya"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   270
        and j: "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   271
        "rec_calc_rel g (l @ [y, ryb]) ry" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   272
        "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryc" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   273
        "rec_calc_rel g (l @ [y, ryc]) rya"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   274
      from j show "ry = rya"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   275
	apply(insert ind[of ryb ryc], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   276
	apply(insert ng_ind[of "l @ [y, ryc]" ry rya], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   277
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   278
    qed 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   279
    from k1 and this show "x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   280
      apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   281
      apply(insert ng_ind[of "la @ [ya, rya]" x y], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   282
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   283
  qed  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   284
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   285
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   286
lemma Suc_nth_part_eq:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   287
  "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   288
       \<Longrightarrow> \<forall>k<(length list). (xs) ! k = (list) ! k"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   289
apply(rule allI, rule impI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   290
apply(erule_tac x = "Suc k" in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   291
done
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
lemma list_eq_intro:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   295
  "\<lbrakk>length xs = length ys; \<forall> k < length xs. xs ! k = ys ! k\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   296
  \<Longrightarrow> xs = ys"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   297
apply(induct xs arbitrary: ys, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   298
apply(case_tac ys, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   299
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   300
  fix a xs ys aa list
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   301
  assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   302
    "\<And>ys. \<lbrakk>length list = length ys; \<forall>k<length ys. xs ! k = ys ! k\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   303
    \<Longrightarrow> xs = ys"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   304
    and h: "length xs = length list" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   305
    "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   306
  from h show "a = aa \<and> xs = list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   307
    apply(insert ind[of list], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   308
    apply(frule Suc_nth_part_eq, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   309
    apply(erule_tac x = "0" in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   310
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   311
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   312
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   313
lemma rec_calc_inj_case_cn: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   314
  assumes ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   315
  "\<And>x l xa y.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   316
  \<lbrakk>x = f \<or> x \<in> set gs; rec_calc_rel x l xa; rec_calc_rel x l y\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   317
  \<Longrightarrow> xa = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   318
  and h: "rec_calc_rel (Cn n f gs) l x" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   319
         "rec_calc_rel (Cn n f gs) l y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   320
  shows "x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   321
  apply(insert h, elim  calc_cn_reverse)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   322
  apply(subgoal_tac "rs = rsa")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   323
  apply(rule_tac x = f and l = rsa and xa = x and y = y in ind, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   324
        simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   325
  apply(intro list_eq_intro, simp, rule allI, rule impI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   326
  apply(erule_tac x = k in allE, rule_tac x = k in allE, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   327
  apply(rule_tac x = "gs ! k" in ind, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   328
  done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   329
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   330
lemma rec_calc_inj:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   331
  "\<lbrakk>rec_calc_rel f l x; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   332
    rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   333
apply(induct f arbitrary: l x y rule: rec_ci.induct)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   334
apply(simp add: rec_calc_inj_case_z)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   335
apply(simp add: rec_calc_inj_case_s)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   336
apply(simp add: rec_calc_inj_case_id, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   337
apply(erule rec_calc_inj_case_cn,simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   338
apply(erule rec_calc_inj_case_pr, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   339
apply(erule rec_calc_inj_case_mn, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   340
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   341
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   342
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   343
lemma calc_rel_reverse_ind_step_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   344
  "\<lbrakk>rec_calc_rel (Pr n f g) (lm @ [Suc x]) rs\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   345
  \<Longrightarrow> \<exists> rs. rec_calc_rel (Pr n f g) (lm @ [x]) rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   346
apply(erule calc_pr_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   347
apply(rule_tac x = rk in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   348
done
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
lemma [simp]: "Suc x \<le> y \<Longrightarrow> Suc (y - Suc x) = y - x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   351
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   352
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   353
lemma calc_pr_para_not_null: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   354
  "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> lm \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   355
apply(erule calc_pr_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   356
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   357
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   358
lemma calc_pr_less_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   359
 "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; x \<le> last lm\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   360
 \<exists>rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   361
apply(subgoal_tac "lm \<noteq> []")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   362
apply(induct x, rule_tac x = rs in exI, simp, simp, erule exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   363
apply(rule_tac rs = xa in calc_rel_reverse_ind_step_ex, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   364
apply(simp add: calc_pr_para_not_null)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   365
done
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
lemma calc_pr_zero_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   368
  "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   369
             \<exists>rs. rec_calc_rel f (butlast lm) rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   370
apply(drule_tac x = "last lm" in calc_pr_less_ex, simp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   371
      erule_tac exE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   372
apply(erule_tac calc_pr_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   373
apply(rule_tac x = rs in exI, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   374
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   375
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   376
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   377
lemma abc_steps_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   378
  "abc_steps_l (as, am) ap (Suc stp) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   379
          abc_steps_l (abc_steps_l (as, am) ap stp) ap (Suc 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   380
apply(insert abc_steps_add[of as am ap stp "Suc 0"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   381
done
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
lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   384
apply(case_tac asm, simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   385
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   386
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   387
lemma abc_append_nth: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   388
  "n < length ap + length bp \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   389
       (ap [+] bp) ! n =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   390
         (if n < length ap then ap ! n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   391
          else abc_inst_shift (bp ! (n - length ap)) (length ap))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   392
apply(simp add: abc_append.simps nth_append map_nth split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   393
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   394
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   395
lemma abc_state_keep:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   396
  "as \<ge> length bp \<Longrightarrow> abc_steps_l (as, lm) bp stp = (as, lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   397
apply(induct stp, simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   398
apply(simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   399
apply(simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   400
apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   401
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   402
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   403
lemma abc_halt_equal: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   404
  "\<lbrakk>abc_steps_l (0, lm) bp stpa = (length bp, lm1); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   405
    abc_steps_l (0, lm) bp stpb = (length bp, lm2)\<rbrakk> \<Longrightarrow> lm1 = lm2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   406
apply(case_tac "stpa - stpb > 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   407
apply(insert abc_steps_add[of 0 lm bp stpb "stpa - stpb"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   408
apply(insert abc_state_keep[of bp "length bp" lm2 "stpa - stpb"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   409
      simp, simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   410
apply(insert abc_steps_add[of 0 lm bp stpa "stpb - stpa"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   411
apply(insert abc_state_keep[of bp "length bp" lm1 "stpb - stpa"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   412
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   413
done  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   414
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   415
lemma abc_halt_point_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   416
  "\<lbrakk>\<exists>stp. abc_steps_l (0, lm) bp stp = (bs, lm');
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   417
    bs = length bp; bp \<noteq> []\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   418
  \<Longrightarrow> \<exists> stp. (\<lambda> (s, l). s < bs \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   419
              (abc_steps_l (s, l) bp (Suc 0)) = (bs, lm')) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   420
      (abc_steps_l (0, lm) bp stp) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   421
apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   422
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   423
  fix stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   424
  assume "bs = length bp" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   425
         "abc_steps_l (0, lm) bp stp = (bs, lm')" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   426
         "bp \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   427
  thus 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   428
    "\<exists>stp. (\<lambda>(s, l). s < bs \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   429
      abc_steps_l (s, l) bp (Suc 0) = (bs, lm')) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   430
                       (abc_steps_l (0, lm) bp stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   431
    apply(induct stp, simp add: abc_steps_zero, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   432
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   433
    fix stpa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   434
    assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   435
     "abc_steps_l (0, lm) bp stpa = (length bp, lm')
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   436
       \<Longrightarrow> \<exists>stp. (\<lambda>(s, l). s < length bp  \<and> abc_steps_l (s, l) bp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   437
             (Suc 0) = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   438
    and h: "abc_steps_l (0, lm) bp (Suc stpa) = (length bp, lm')" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   439
           "abc_steps_l (0, lm) bp stp = (length bp, lm')" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   440
           "bp \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   441
    from h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   442
      "\<exists>stp. (\<lambda>(s, l). s < length bp \<and> abc_steps_l (s, l) bp (Suc 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   443
                    = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   444
      apply(case_tac "abc_steps_l (0, lm) bp stpa", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   445
            case_tac "a = length bp")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   446
      apply(insert ind, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   447
      apply(subgoal_tac "b = lm'", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   448
      apply(rule_tac abc_halt_equal, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   449
      apply(rule_tac x = stpa in exI, simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   450
      apply(simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   451
      apply(rule classical, simp add: abc_steps_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   452
                             abc_fetch.simps abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   453
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   454
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   455
qed  
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   458
lemma abc_append_empty_r[simp]: "[] [+] ab = ab"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   459
apply(simp add: abc_append.simps abc_inst_shift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   460
apply(induct ab, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   461
apply(case_tac a, simp_all add: abc_inst_shift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   462
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   463
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   464
lemma abc_append_empty_l[simp]:  "ab [+] [] = ab"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   465
apply(simp add: abc_append.simps abc_inst_shift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   466
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   467
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   468
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   469
lemma abc_append_length[simp]:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   470
  "length (ap [+] bp) = length ap + length bp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   471
apply(simp add: abc_append.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   472
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   473
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   474
lemma abc_append_commute: "as [+] bs [+] cs = as [+] (bs [+] cs)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   475
apply(simp add: abc_append.simps abc_shift.simps abc_inst_shift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   476
apply(induct cs, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   477
apply(case_tac a, auto simp: abc_inst_shift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   478
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   479
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   480
lemma abc_halt_point_step[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   481
  "\<lbrakk>a < length bp; abc_steps_l (a, b) bp (Suc 0) = (length bp, lm')\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   482
  \<Longrightarrow> abc_steps_l (length ap + a, b) (ap [+] bp [+] cp) (Suc 0) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   483
                                        (length ap + length bp, lm')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   484
apply(simp add: abc_steps_l.simps abc_fetch.simps abc_append_nth)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   485
apply(case_tac "bp ! a", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   486
                      auto simp: abc_steps_l.simps abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   487
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   488
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   489
lemma abc_step_state_in:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   490
  "\<lbrakk>bs < length bp;  abc_steps_l (a, b) bp (Suc 0) = (bs, l)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   491
  \<Longrightarrow> a < length bp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   492
apply(simp add: abc_steps_l.simps abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   493
apply(rule_tac classical, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   494
      simp add: abc_step_l.simps abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   495
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   496
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   497
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   498
lemma abc_append_state_in_exc: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   499
  "\<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   500
 \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   501
                                             (length ap + bs, l)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   502
apply(induct stpa arbitrary: bs l, simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   503
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   504
  fix stpa bs l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   505
  assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   506
    "\<And>bs l. \<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   507
    \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   508
                                                (length ap + bs, l)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   509
    and h: "bs < length bp" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   510
           "abc_steps_l (0, lm) bp (Suc stpa) = (bs, l)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   511
  from h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   512
    "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) (Suc stpa) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   513
                                                (length ap + bs, l)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   514
    apply(simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   515
    apply(case_tac "(abc_steps_l (0, lm) bp stpa)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   516
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   517
    fix a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   518
    assume g: "abc_steps_l (0, lm) bp stpa = (a, b)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   519
              "abc_steps_l (a, b) bp (Suc 0) = (bs, l)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   520
    from h and g have k1: "a < length bp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   521
      apply(simp add: abc_step_state_in)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   522
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   523
    from h and g and k1 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   524
   "abc_steps_l (abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   525
              (ap [+] bp [+] cp) (Suc 0) = (length ap + bs, l)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   526
      apply(insert ind[of a b], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   527
      apply(simp add: abc_steps_l.simps abc_fetch.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   528
                      abc_append_nth)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   529
      apply(case_tac "bp ! a", auto simp: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   530
                                 abc_steps_l.simps abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   531
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   532
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   533
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   534
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   535
lemma [simp]: "abc_steps_l (0, am) [] stp = (0, am)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   536
apply(induct stp, simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   537
apply(simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   538
apply(simp add: abc_steps_zero abc_steps_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   539
                abc_fetch.simps abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   540
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   541
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   542
lemma abc_append_exc1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   543
  "\<lbrakk>\<exists> stp. abc_steps_l (0, lm) bp stp = (bs, lm');
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   544
    bs = length bp; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   545
    as = length ap\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   546
    \<Longrightarrow> \<exists> stp. abc_steps_l (as, lm) (ap [+] bp [+] cp) stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   547
                                                 = (as + bs, lm')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   548
apply(case_tac "bp = []", erule_tac exE, simp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   549
      rule_tac x = 0 in exI, simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   550
apply(frule_tac abc_halt_point_ex, simp, simp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   551
      erule_tac exE, erule_tac exE) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   552
apply(rule_tac x = "stpa + Suc 0" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   553
apply(case_tac "(abc_steps_l (0, lm) bp stpa)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   554
      simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   555
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   556
  "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   557
                                   = (length ap + a, b)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   558
apply(simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   559
apply(rule_tac abc_append_state_in_exc, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   560
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   561
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   562
lemma abc_append_exc3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   563
  "\<lbrakk>\<exists> stp. abc_steps_l (0, am) bp stp = (bs, bm); ss = length ap\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   564
   \<Longrightarrow>  \<exists> stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   565
apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   566
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   567
  fix stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   568
  assume h: "abc_steps_l (0, am) bp stp = (bs, bm)" "ss = length ap"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   569
  thus " \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   570
  proof(induct stp arbitrary: bs bm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   571
    fix bs bm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   572
    assume "abc_steps_l (0, am) bp 0 = (bs, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   573
    thus "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   574
      apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   575
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   576
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   577
    fix stp bs bm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   578
    assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   579
      "\<And>bs bm. \<lbrakk>abc_steps_l (0, am) bp stp = (bs, bm);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   580
                 ss = length ap\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   581
          \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   582
    and g: "abc_steps_l (0, am) bp (Suc stp) = (bs, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   583
    from g show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   584
      "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   585
      apply(insert abc_steps_add[of 0 am bp stp "Suc 0"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   586
      apply(case_tac "(abc_steps_l (0, am) bp stp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   587
    proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   588
      fix a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   589
      assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   590
             "abc_steps_l (0, am) bp (Suc stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   591
                       abc_steps_l (a, b) bp (Suc 0)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   592
              "abc_steps_l (0, am) bp stp = (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   593
      thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   594
	apply(insert ind[of a b], simp add: h, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   595
	apply(rule_tac x = "Suc stp" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   596
	apply(simp only: abc_steps_ind, simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   597
      proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   598
	fix stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   599
	assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   600
	thus "abc_steps_l (a + length ap, b) (ap [+] bp) (Suc 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   601
                                              = (bs + length ap, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   602
	  apply(simp add: abc_steps_l.simps abc_steps_zero
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   603
                          abc_fetch.simps split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   604
	  apply(case_tac "bp ! a", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   605
                simp_all add: abc_inst_shift.simps abc_append_nth
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   606
                   abc_steps_l.simps abc_steps_zero abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   607
	  apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   608
	  done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   609
      qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   610
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   611
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   612
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   613
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   614
lemma abc_add_equal:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   615
  "\<lbrakk>ap \<noteq> []; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   616
    abc_steps_l (0, am) ap astp = (a, b);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   617
    a < length ap\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   618
     \<Longrightarrow> (abc_steps_l (0, am) (ap @ bp) astp) = (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   619
apply(induct astp arbitrary: a b, simp add: abc_steps_l.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   620
apply(simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   621
apply(case_tac "(abc_steps_l (0, am) ap astp)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   622
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   623
  fix astp a b aa ba
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   624
  assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   625
    "\<And>a b. \<lbrakk>abc_steps_l (0, am) ap astp = (a, b); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   626
             a < length ap\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   627
                  abc_steps_l (0, am) (ap @ bp) astp = (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   628
  and h: "abc_steps_l (abc_steps_l (0, am) ap astp) ap (Suc 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   629
                                                            = (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   630
        "a < length ap" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   631
        "abc_steps_l (0, am) ap astp = (aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   632
  from h show "abc_steps_l (abc_steps_l (0, am) (ap @ bp) astp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   633
                                        (ap @ bp) (Suc 0) = (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   634
    apply(insert ind[of aa ba], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   635
    apply(subgoal_tac "aa < length ap", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   636
    apply(simp add: abc_steps_l.simps abc_fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   637
                     nth_append abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   638
    apply(rule abc_step_state_in, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   639
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   640
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   641
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   642
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   643
lemma abc_add_exc1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   644
  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   645
  \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap @ bp) stp = (as, bm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   646
apply(case_tac "ap = []", simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   647
      rule_tac x = 0 in exI, simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   648
apply(drule_tac abc_halt_point_ex, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   649
apply(erule_tac exE, case_tac "(abc_steps_l (0, am) ap astp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   650
apply(rule_tac x = "Suc astp" in exI, simp add: abc_steps_ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   651
apply(frule_tac bp = bp in abc_add_equal, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   652
apply(simp add: abc_steps_l.simps abc_steps_zero 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   653
                abc_fetch.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   654
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   655
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   656
declare abc_shift.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   657
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   658
lemma abc_append_exc2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   659
  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   660
    \<exists> bstp. abc_steps_l (0, bm) bp bstp = (bs, bm'); bs = length bp;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   661
    cs = as + bs; bp \<noteq> []\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   662
  \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap [+] bp) stp = (cs, bm')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   663
apply(insert abc_append_exc1[of bm bp bs bm' as ap "[]"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   664
apply(drule_tac bp = "abc_shift bp (length ap)" in abc_add_exc1, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   665
apply(subgoal_tac "ap @ abc_shift bp (length ap) = ap [+] bp", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   666
      simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   667
apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   668
apply(simp add: abc_append.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   669
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   670
lemma exp_length[simp]: "length (a\<^bsup>b\<^esup>) = b"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   671
by(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   672
lemma exponent_add_iff: "a\<^bsup>b\<^esup> @ a\<^bsup>c \<^esup>@ xs = a\<^bsup>b + c \<^esup>@ xs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   673
apply(auto simp: exponent_def replicate_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   674
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   675
lemma exponent_cons_iff: "a # a\<^bsup>c \<^esup>@ xs = a\<^bsup>Suc c \<^esup>@ xs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   676
apply(auto simp: exponent_def replicate_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   677
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   678
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   679
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   680
lemma  [simp]: "length lm = n \<Longrightarrow>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   681
  abc_steps_l (Suc 0, lm @ Suc x # 0 # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   682
       [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   683
                                  = (3, lm @ Suc x # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   684
apply(simp add: abc_steps_l.simps abc_fetch.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   685
                abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   686
                nth_append list_update_append)
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 [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   690
  "length lm = n \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   691
  abc_steps_l (Suc 0, lm @ Suc x # Suc y # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   692
     [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   693
  = (Suc 0, lm @ Suc x # y # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   694
apply(simp add: abc_steps_l.simps abc_fetch.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   695
                abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   696
                nth_append list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   697
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   698
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   699
lemma pr_cycle_part_middle_inv: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   700
  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   701
  \<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   702
                         [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   703
  = (3, lm @ Suc x # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   704
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   705
  assume h: "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   706
  hence k1: "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   707
                           [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   708
    = (Suc 0, lm @ Suc x # y # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   709
    apply(rule_tac x = "Suc 0" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   710
    apply(simp add: abc_steps_l.simps abc_step_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   711
                    abc_lm_v.simps abc_lm_s.simps nth_append 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   712
                    list_update_append abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   713
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   714
  from h have k2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   715
    "\<exists> stp. abc_steps_l (Suc 0, lm @ Suc x # y # suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   716
                      [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   717
    = (3, lm @ Suc x # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   718
    apply(induct y)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   719
    apply(rule_tac x = "Suc (Suc 0)" in exI, simp, simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   720
          erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   721
    apply(rule_tac x = "Suc (Suc 0) + stp" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   722
          simp only: abc_steps_add, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   723
    done      
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   724
  from k1 and k2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   725
    "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   726
                       [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   727
    = (3, lm @ Suc x # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   728
    apply(erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   729
    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
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
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   732
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   733
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   734
  "length lm = Suc n \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   735
  (abc_steps_l (length ap, lm @ x # Suc y # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   736
           (ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length ap)]) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   737
                    (Suc (Suc (Suc 0))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   738
  = (length ap, lm @ Suc x # y # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   739
apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   740
         abc_lm_v.simps list_update_append nth_append abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   741
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   742
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   743
lemma switch_para_inv:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   744
  assumes bp_def:"bp =  ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto ss]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   745
  and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   746
         "ss = length ap" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   747
         "length lm = Suc n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   748
  shows " \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   749
                               (0, lm @ (x + y) # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   750
apply(induct y arbitrary: x)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   751
apply(rule_tac x = "Suc 0" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   752
  simp add: bp_def empty.simps abc_steps_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   753
            abc_fetch.simps h abc_step_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   754
            abc_lm_v.simps list_update_append nth_append
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   755
            abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   756
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   757
  fix y x
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   758
  assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   759
    "\<And>x. \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   760
                                     (0, lm @ (x + y) # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   761
  show "\<exists>stp. abc_steps_l (ss, lm @ x # Suc y # suf_lm) bp stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   762
                                  (0, lm @ (x + Suc y) # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   763
    apply(insert ind[of "Suc x"], erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   764
    apply(rule_tac x = "Suc (Suc (Suc 0)) + stp" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   765
          simp only: abc_steps_add bp_def h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   766
    apply(simp add: h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   767
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   768
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   769
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   770
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   771
  "length lm = rs_pos \<and> Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   772
      a_md - Suc 0 < Suc (Suc (Suc (a_md + length suf_lm - 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   773
                                         Suc (Suc (Suc 0)))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   774
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   775
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   776
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   777
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   778
  "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   779
                           \<not> a_md - Suc 0 < rs_pos - Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   780
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   781
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   782
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   783
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   784
  "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   785
           \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   786
apply(arith)
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 butlast_append_last: "lm \<noteq> [] \<Longrightarrow> lm = butlast lm @ [last lm]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   790
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   791
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   792
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   793
lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   794
           \<Longrightarrow> (Suc (Suc rs_pos)) < a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   795
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   796
apply(case_tac "rec_ci f", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   797
apply(case_tac "rec_ci g", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   798
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   799
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   800
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   801
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   802
lemma pr_para_ge_suc0: "rec_calc_rel (Pr n f g) lm xs \<Longrightarrow> 0 < n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   803
apply(erule calc_pr_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   804
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   805
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   806
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   807
lemma ci_pr_para_eq: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   808
                  \<Longrightarrow> rs_pos = Suc n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   809
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   810
apply(case_tac "rec_ci g",  case_tac "rec_ci f", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   811
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   812
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   813
lemma [intro]:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   814
  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm xs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   815
  \<Longrightarrow> length lm = rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   816
apply(simp add: rec_ci.simps rec_ci_z_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   817
apply(erule_tac calc_z_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   818
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   819
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   820
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   821
  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm xs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   822
  \<Longrightarrow> length lm = rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   823
apply(simp add: rec_ci.simps rec_ci_s_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   824
apply(erule_tac calc_s_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   825
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   826
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   827
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   828
  "\<lbrakk>rec_ci (recf.id nat1 nat2) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   829
    rec_calc_rel (recf.id nat1 nat2) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   830
apply(simp add: rec_ci.simps rec_ci_id.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   831
apply(erule_tac calc_id_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   832
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   833
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   834
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   835
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   836
    rec_calc_rel (Cn n f gs) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   837
apply(erule_tac calc_cn_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   838
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   839
apply(case_tac "rec_ci f",  simp)
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 [intro]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   843
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   844
    rec_calc_rel (Pr n f g) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   845
apply(erule_tac  calc_pr_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   846
apply(drule_tac ci_pr_para_eq, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   847
apply(drule_tac ci_pr_para_eq, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   848
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   849
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   850
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   851
  "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   852
    rec_calc_rel (Mn n f) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   853
apply(erule_tac calc_mn_reverse)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   854
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   855
apply(case_tac "rec_ci f",  simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   856
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   857
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   858
lemma para_pattern: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   859
  "\<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm xs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   860
  \<Longrightarrow> length lm = rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   861
apply(case_tac f, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   862
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   863
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   864
lemma ci_pr_g_paras:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   865
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   866
    rec_ci g = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   867
    rec_calc_rel (Pr n f g) (lm @ [x]) rs; x > 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   868
    aa = Suc rs_pos "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   869
apply(erule calc_pr_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   870
apply(subgoal_tac "length (args @ [k, rk]) = aa", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   871
apply(subgoal_tac "rs_pos = Suc n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   872
apply(simp add: ci_pr_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   873
apply(erule para_pattern, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   874
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   875
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   876
lemma ci_pr_g_md_less: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   877
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   878
    rec_ci g = (a, aa, ba)\<rbrakk> \<Longrightarrow> ba < a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   879
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   880
apply(case_tac "rec_ci f",  auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   881
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   882
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   883
lemma [intro]: "rec_ci z = (ap, rp, ad) \<Longrightarrow> rp < ad"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   884
  by(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   885
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   886
lemma [intro]: "rec_ci s = (ap, rp, ad) \<Longrightarrow> rp < ad"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   887
  by(simp add: rec_ci.simps)
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 [intro]: "rec_ci (recf.id nat1 nat2) = (ap, rp, ad) \<Longrightarrow> rp < ad"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   890
  by(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   891
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   892
lemma [intro]: "rec_ci (Cn n f gs) = (ap, rp, ad) \<Longrightarrow> rp < ad"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   893
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   894
apply(case_tac "rec_ci f",  simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   895
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   896
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   897
lemma [intro]: "rec_ci (Pr n f g) = (ap, rp, ad) \<Longrightarrow> rp < ad"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   898
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   899
by(case_tac "rec_ci f", case_tac "rec_ci g",  auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   900
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   901
lemma [intro]: "rec_ci (Mn n f) = (ap, rp, ad) \<Longrightarrow> rp < ad"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   902
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   903
apply(case_tac "rec_ci f", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   904
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   905
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   906
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   907
lemma ci_ad_ge_paras: "rec_ci f = (ap, rp, ad) \<Longrightarrow> ad > rp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   908
apply(case_tac f, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   909
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   910
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   911
lemma [elim]: "\<lbrakk>a [+] b = []; a \<noteq> [] \<or> b \<noteq> []\<rbrakk> \<Longrightarrow> RR"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   912
apply(auto simp: abc_append.simps abc_shift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   913
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   914
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   915
lemma [intro]: "rec_ci z = ([], aa, ba) \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   916
by(simp add: rec_ci.simps rec_ci_z_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   917
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   918
lemma [intro]: "rec_ci s = ([], aa, ba) \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   919
by(auto simp: rec_ci.simps rec_ci_s_def addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   920
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   921
lemma [intro]: "rec_ci (id m n) = ([], aa, ba) \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   922
by(auto simp: rec_ci.simps rec_ci_id.simps addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   923
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   924
lemma [intro]: "rec_ci (Cn n f gs) = ([], aa, ba) \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   925
apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_append.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   926
apply(simp add: abc_shift.simps empty.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   927
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   928
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   929
lemma [intro]: "rec_ci (Pr n f g) = ([], aa, ba) \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   930
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   931
apply(case_tac "rec_ci f", case_tac "rec_ci g")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   932
by(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   933
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   934
lemma [intro]: "rec_ci (Mn n f) = ([], aa, ba) \<Longrightarrow> False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   935
apply(case_tac "rec_ci f", auto simp: rec_ci.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
lemma rec_ci_not_null:  "rec_ci g = (a, aa, ba) \<Longrightarrow> a \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   939
by(case_tac g, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   940
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   941
lemma calc_pr_g_def:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   942
 "\<lbrakk>rec_calc_rel (Pr rs_pos f g) (lm @ [Suc x]) rsa;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   943
   rec_calc_rel (Pr rs_pos f g) (lm @ [x]) rsxa\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   944
 \<Longrightarrow> rec_calc_rel g (lm @ [x, rsxa]) rsa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   945
apply(erule_tac calc_pr_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   946
apply(subgoal_tac "rsxa = rk", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   947
apply(erule_tac rec_calc_inj, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   948
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   949
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   950
lemma ci_pr_md_def: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   951
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   952
    rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   953
  \<Longrightarrow> a_md = Suc (max (n + 3) (max bc ba))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   954
by(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   955
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   956
lemma  ci_pr_f_paras: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   957
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   958
    rec_calc_rel (Pr n f g) lm rs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   959
    rec_ci f = (ab, ac, bc)\<rbrakk>  \<Longrightarrow> ac = rs_pos - Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   960
apply(subgoal_tac "\<exists>rs. rec_calc_rel f (butlast lm) rs", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   961
      erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   962
apply(drule_tac f = f and lm = "butlast lm" in para_pattern, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   963
      simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   964
apply(drule_tac para_pattern, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   965
apply(subgoal_tac "lm \<noteq> []", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   966
apply(erule_tac calc_pr_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   967
apply(erule calc_pr_zero_ex)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   968
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   969
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   970
lemma ci_pr_md_ge_f:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   971
        rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> Suc bc \<le> a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   972
apply(case_tac "rec_ci g")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   973
apply(simp add: rec_ci.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   974
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   975
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   976
lemma ci_pr_md_ge_g:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   977
        rec_ci g = (ab, ac, bc)\<rbrakk> \<Longrightarrow> bc < a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   978
apply(case_tac "rec_ci f")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   979
apply(simp add: rec_ci.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   980
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   981
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   982
lemma rec_calc_rel_def0: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   983
  "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; rec_calc_rel f (butlast lm) rsa\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   984
  \<Longrightarrow> rec_calc_rel (Pr n f g) (butlast lm @ [0]) rsa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   985
  apply(rule_tac calc_pr_zero, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   986
apply(erule_tac calc_pr_reverse, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   987
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   988
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   989
lemma [simp]:  "length (empty m n) = 3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   990
by (auto simp: empty.simps)
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
lemma
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   993
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   994
  rec_calc_rel (Pr n f g) lm rs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   995
  rec_ci g = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   996
  rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   997
\<Longrightarrow> \<exists>ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 + length ab \<and> bp = recursive.empty (n - Suc 0) n 3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   998
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   999
apply(rule_tac x = "recursive.empty (n - Suc 0) (max (Suc (Suc n)) (max bc ba)) 3 [+] ab" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1000
apply(rule_tac x = "([Dec (max (Suc (Suc n)) (max bc ba)) (length a + 7)] [+] a [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1001
  [Inc (n - Suc 0), Dec n 3, Goto (Suc 0)]) @ [Dec (Suc n) 0, Inc n, Goto (length a + 4)]" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1002
apply(auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1003
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1004
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1005
lemma  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1006
        rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1007
    \<Longrightarrow> \<exists>ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> bp = ab"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1008
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1009
apply(rule_tac x = "recursive.empty (n - Suc 0) (max (Suc (Suc n)) (max bc ba)) 3" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1010
apply(rule_tac x = "recursive.empty (n - Suc 0) n 3 [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1011
     ([Dec (max (Suc (Suc n)) (max bc ba)) (length a + 7)] [+] a 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1012
  [+] [Inc (n - Suc 0), Dec n 3, Goto (Suc 0)]) @ [Dec (Suc n) 0, Inc n, Goto (length a + 4)]" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1013
apply(simp add: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1014
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1015
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1016
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1017
lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1018
    \<Longrightarrow> rs_pos = Suc n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1019
apply(simp add: ci_pr_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1020
done
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1023
lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1024
    \<Longrightarrow> length lm = Suc n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1025
apply(subgoal_tac "rs_pos = Suc n", rule_tac para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1026
apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1027
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1028
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1029
lemma [simp]: "rec_ci (Pr n f g) = (a, rs_pos, a_md) \<Longrightarrow> Suc (Suc n) < a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1030
apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1031
apply arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1032
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1033
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1034
lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) \<Longrightarrow> 0 < rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1035
apply(case_tac "rec_ci f", case_tac "rec_ci g")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1036
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1037
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1038
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1039
lemma [simp]: "Suc (Suc rs_pos) < a_md \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1040
       butlast lm @ (last lm - xa) # (rsa::nat) # 0 # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1041
       butlast lm @ (last lm - xa) # rsa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1042
apply(simp add: exp_ind_def[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1043
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1044
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1045
lemma pr_cycle_part_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1046
  assumes g_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1047
  "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1048
  \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1049
                    (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1050
  and ap_def: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1051
  "ap = ([Dec (a_md - Suc 0) (length a + 7)] [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1052
        (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1053
         [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1054
  and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1055
         "rec_calc_rel (Pr n f g) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1056
                   (butlast lm @ [last lm - Suc xa]) rsxa" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1057
         "Suc xa \<le> last lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1058
         "rec_ci g = (a, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1059
         "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1060
         "lm \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1061
  shows 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1062
  "\<exists>stp. abc_steps_l 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1063
     (0, butlast lm @ (last lm - Suc xa) # rsxa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1064
               0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1065
     (0, butlast lm @ (last lm - xa) # rsa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1066
                 # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1067
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1068
  have k1: "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1069
    rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1070
         (length a + 4, butlast lm @ (last lm - xa) # 0 # rsa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1071
                           0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1072
    apply(simp add: ap_def, rule_tac abc_add_exc1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1073
    apply(rule_tac as = "Suc 0" and 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1074
      bm = "butlast lm @ (last lm - Suc xa) # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1075
      rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm" in abc_append_exc2,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1076
      auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1077
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1078
    show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1079
      "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1080
                   # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1081
              [Dec (a_md - Suc 0)(length a + 7)] astp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1082
      (Suc 0, butlast lm @ (last lm - Suc xa) # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1083
             rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1084
      apply(rule_tac x = "Suc 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1085
          simp add: abc_steps_l.simps abc_step_l.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1086
                     abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1087
      apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1088
                              a_md > Suc (Suc rs_pos)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1089
      apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1090
      apply(insert nth_append[of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1091
                 "(last lm - Suc xa) # rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup>" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1092
                 "Suc xa # suf_lm" "(a_md - rs_pos)"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1093
      apply(simp add: list_update_append del: list_update.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1094
      apply(insert list_update_append[of "(last lm - Suc xa) # rsxa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1095
                                           0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup>" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1096
                    "Suc xa # suf_lm" "a_md - rs_pos" "xa"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1097
      apply(case_tac a_md, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1098
      apply(insert h, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1099
      apply(insert para_pattern[of "Pr n f g" aprog rs_pos a_md 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1100
                    "(butlast lm @ [last lm - Suc xa])" rsxa], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1101
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1102
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1103
    show "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1104
           rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm) (a [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1105
            [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]) bstp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1106
         (3 + length a, butlast lm @ (last lm - xa) # 0 # rsa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1107
                          0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1108
      apply(rule_tac as = "length a" and
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1109
               bm = "butlast lm @ (last lm - Suc xa) # rsxa # rsa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1110
                     0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1111
        in abc_append_exc2, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1112
    proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1113
      from h have j1: "aa = Suc rs_pos \<and> a_md > ba \<and> ba > Suc rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1114
	apply(insert h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1115
	apply(insert ci_pr_g_paras[of n f g aprog rs_pos
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1116
                 a_md a aa ba "butlast lm" "last lm - xa" rsa], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1117
	apply(drule_tac ci_pr_md_ge_g, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1118
	apply(erule_tac ci_ad_ge_paras)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1119
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1120
      from h have j2: "rec_calc_rel g (butlast lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1121
                                  [last lm - Suc xa, rsxa]) rsa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1122
	apply(rule_tac  calc_pr_g_def, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1123
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1124
      from j1 and j2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1125
        "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1126
                rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm) a astp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1127
        (length a, butlast lm @ (last lm - Suc xa) # rsxa # rsa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1128
                         # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1129
	apply(insert g_ind[of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1130
          "butlast lm @ (last lm - Suc xa) # [rsxa]" rsa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1131
          "0\<^bsup>a_md - ba - Suc 0 \<^esup> @ xa # suf_lm"], simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1132
	apply(simp add: exponent_add_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1133
	apply(rule_tac x = stp in exI, simp add: numeral_3_eq_3)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1134
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1135
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1136
      from h have j3: "length lm = rs_pos \<and> rs_pos > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1137
	apply(rule_tac conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1138
	apply(drule_tac lm = "(butlast lm @ [last lm - Suc xa])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1139
                          and xs = rsxa in para_pattern, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1140
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1141
      from h have j4: "Suc (last lm - Suc xa) = last lm - xa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1142
	apply(case_tac "last lm", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1143
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1144
      from j3 and j4 show
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1145
      "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1146
                     rsa # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1147
            [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)] bstp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1148
        (3, butlast lm @ (last lm - xa) # 0 # rsa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1149
                       0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1150
	apply(insert pr_cycle_part_middle_inv[of "butlast lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1151
          "rs_pos - Suc 0" "(last lm - Suc xa)" rsxa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1152
          "rsa # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1153
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1154
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1155
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1156
  from h have k2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1157
    "\<exists>stp. abc_steps_l (length a + 4, butlast lm @ (last lm - xa) # 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1158
           # rsa # 0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1159
    (0, butlast lm @ (last lm - xa) # rsa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1160
    apply(insert switch_para_inv[of ap 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1161
      "([Dec (a_md - Suc 0) (length a + 7)] [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1162
      (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1163
      n "length a + 4" f g aprog rs_pos a_md 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1164
      "butlast lm @ [last lm - xa]" 0 rsa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1165
      "0\<^bsup>a_md - Suc (Suc (Suc rs_pos))\<^esup> @ xa # suf_lm"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1166
    apply(simp add: h ap_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1167
    apply(subgoal_tac "length lm = Suc n \<and> Suc (Suc rs_pos) < a_md", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1168
          simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1169
    apply(insert h, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1170
    apply(frule_tac lm = "(butlast lm @ [last lm - Suc xa])" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1171
      and xs = rsxa in para_pattern, simp, simp)
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
  from k1 and k2 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1174
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1175
    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
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
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1178
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1179
lemma ci_pr_ex1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1180
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1181
    rec_ci g = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1182
    rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1183
\<Longrightarrow> \<exists>ap bp. length ap = 6 + length ab \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1184
    aprog = ap [+] bp \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1185
    bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] (a [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1186
         [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1187
         [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1188
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1189
apply(rule_tac x = "recursive.empty n (max (Suc (Suc (Suc n)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1190
    (max bc ba)) [+] ab [+] recursive.empty n (Suc n)" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1191
     simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1192
apply(auto simp add: abc_append_commute add3_Suc)
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 pr_cycle_part:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1196
  "\<lbrakk>\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1197
     \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1198
                        (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1199
  rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1200
  rec_calc_rel (Pr n f g) lm rs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1201
  rec_ci g = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1202
  rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1203
  rec_ci f = (ab, ac, bc);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1204
  lm \<noteq> [];
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1205
  x \<le> last lm\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1206
  \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1207
              rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ x # suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1208
  (6 + length ab, butlast lm @ last lm # rs #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1209
                                0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1210
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1211
  assume g_ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1212
    "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1213
    \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1214
                      (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1215
    and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1216
           "rec_calc_rel (Pr n f g) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1217
           "rec_ci g = (a, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1218
           "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1219
           "lm \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1220
           "x \<le> last lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1221
           "rec_ci f = (ab, ac, bc)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1222
  from h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1223
    "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1224
            rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ x # suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1225
    (6 + length ab, butlast lm @ last lm # rs #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1226
                               0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1227
  proof(induct x arbitrary: rsx, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1228
    fix rsxa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1229
    assume "rec_calc_rel (Pr n f g) lm rsxa" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1230
           "rec_calc_rel (Pr n f g) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1231
    from h and this have "rs = rsxa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1232
      apply(subgoal_tac "lm \<noteq> [] \<and> rs_pos = Suc n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1233
      apply(rule_tac rec_calc_inj, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1234
      apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1235
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1236
    thus "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @  last lm # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1237
             rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1238
      (6 + length ab, butlast lm @ last lm # rs #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1239
                               0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1240
      by(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1241
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1242
    fix xa rsxa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1243
    assume ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1244
   "\<And>rsx. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsx 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1245
  \<Longrightarrow> \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - xa) #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1246
             rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1247
      (6 + length ab, butlast lm @ last lm # rs # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1248
                               0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1249
      and g: "rec_calc_rel (Pr n f g) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1250
                      (butlast lm @ [last lm - Suc xa]) rsxa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1251
      "Suc xa \<le> last lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1252
      "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1253
      "rec_calc_rel (Pr n f g) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1254
      "rec_ci g = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1255
      "rec_ci f = (ab, ac, bc)" "lm \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1256
    from g have k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1257
      "\<exists> rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1258
      apply(rule_tac rs = rs in  calc_pr_less_ex, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1259
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1260
    from g and this show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1261
      "\<exists>stp. abc_steps_l (6 + length ab, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1262
           butlast lm @ (last lm - Suc xa) # rsxa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1263
              0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1264
              (6 + length ab, butlast lm @ last lm # rs # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1265
                                0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1266
    proof(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1267
      fix rsa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1268
      assume k2: "rec_calc_rel (Pr n f g) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1269
                           (butlast lm @ [last lm - xa]) rsa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1270
      from g and k2 have
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1271
      "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1272
       (last lm - Suc xa) # rsxa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1273
               0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1274
        = (6 + length ab, butlast lm @ (last lm - xa) # rsa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1275
                               0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1276
	proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1277
	  from g have k2_1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1278
            "\<exists> ap bp. length ap = 6 + length ab \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1279
                   aprog = ap [+] bp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1280
                   bp = ([Dec (a_md - Suc 0) (length a + 7)] [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1281
                  (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1282
                  Goto (Suc 0)])) @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1283
                  [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1284
            apply(rule_tac ci_pr_ex1, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1285
	    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1286
	  from k2_1 and k2 and g show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1287
	    proof(erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1288
	      fix ap bp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1289
	      assume 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1290
                "length ap = 6 + length ab \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1291
                 aprog = ap [+] bp \<and> bp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1292
                ([Dec (a_md - Suc 0) (length a + 7)] [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1293
                (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1294
                Goto (Suc 0)])) @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1295
                [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1296
	      from g and this and k2 and g_ind show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1297
		apply(insert abc_append_exc3[of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1298
                  "butlast lm @ (last lm - Suc xa) # rsxa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1299
                  0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # suf_lm" bp 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1300
                  "butlast lm @ (last lm - xa) # rsa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1301
                0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm" "length ap" ap],
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1302
                 simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1303
		apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1304
                "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1305
                           # rsxa # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ Suc xa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1306
                              suf_lm) bp stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1307
	          (0, butlast lm @ (last lm - xa) # rsa #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1308
                           0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ xa # suf_lm)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1309
                      simp, erule_tac conjE, erule conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1310
		apply(erule pr_cycle_part_ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1311
		done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1312
	    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1313
	  qed  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1314
      from g and k2 and this show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1315
	apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1316
	apply(insert ind[of rsa], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1317
	apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1318
	apply(rule_tac x = "stp + stpa" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1319
              simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1320
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1321
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1322
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1323
qed
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 ci_pr_length: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1326
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1327
    rec_ci g = (a, aa, ba);  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1328
    rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1329
    \<Longrightarrow>  length aprog = 13 + length ab + length a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1330
apply(auto simp: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1331
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1332
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1333
thm empty.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1334
term max
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1335
fun empty_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1336
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1337
  "empty_inv (as, lm) m n initlm = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1338
         (let plus = initlm ! m + initlm ! n in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1339
           length initlm > max m n \<and> m \<noteq> n \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1340
              (if as = 0 then \<exists> k l. lm = initlm[m := k, n := l] \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1341
                    k + l = plus \<and> k \<le> initlm ! m 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1342
              else if as = 1 then \<exists> k l. lm = initlm[m := k, n := l]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1343
                             \<and> k + l + 1 = plus \<and> k < initlm ! m 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1344
              else if as = 2 then \<exists> k l. lm = initlm[m := k, n := l] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1345
                              \<and> k + l = plus \<and> k \<le> initlm ! m
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1346
              else if as = 3 then lm = initlm[m := 0, n := plus]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1347
              else False))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1348
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1349
fun empty_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1350
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1351
  "empty_stage1 (as, lm) m  = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1352
            (if as = 3 then 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1353
             else 1)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1354
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1355
fun empty_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1356
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1357
  "empty_stage2 (as, lm) m = (lm ! m)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1358
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1359
fun empty_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1360
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1361
  "empty_stage3 (as, lm) m = (if as = 1 then 3 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1362
                                else if as = 2 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1363
                                else if as = 0 then 1 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1364
                                else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1365
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1366
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1367
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1368
fun empty_measure :: "((nat \<times> nat list) \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1369
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1370
  "empty_measure ((as, lm), m) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1371
     (empty_stage1 (as, lm) m, empty_stage2 (as, lm) m,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1372
      empty_stage3 (as, lm) m)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1373
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1374
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
  1375
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1376
  "lex_pair = less_than <*lex*> less_than"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1377
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1378
definition lex_triple :: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1379
 "((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
  1380
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1381
  "lex_triple \<equiv> less_than <*lex*> lex_pair"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1382
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1383
definition empty_LE :: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1384
 "(((nat \<times> nat list) \<times> nat) \<times> ((nat \<times> nat list) \<times> nat)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1385
  where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1386
  "empty_LE \<equiv> (inv_image lex_triple empty_measure)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1387
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1388
lemma wf_lex_triple: "wf lex_triple"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1389
  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
  1390
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1391
lemma wf_empty_le[intro]: "wf empty_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1392
by(auto intro:wf_inv_image wf_lex_triple simp: empty_LE_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1393
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1394
declare empty_inv.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1395
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1396
lemma empty_inv_init:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1397
"\<lbrakk>m < length initlm; n < length initlm; m \<noteq> n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1398
  empty_inv (0, initlm) m n initlm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1399
apply(simp add: abc_steps_l.simps empty_inv.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1400
apply(rule_tac x = "initlm ! m" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1401
      rule_tac x = "initlm ! n" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1402
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1403
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1404
lemma [simp]: "abc_fetch 0 (recursive.empty m n) = Some (Dec m 3)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1405
apply(simp add: empty.simps abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1406
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1407
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1408
lemma [simp]: "abc_fetch (Suc 0) (recursive.empty m n) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1409
               Some (Inc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1410
apply(simp add: empty.simps abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1411
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1412
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1413
lemma [simp]: "abc_fetch 2 (recursive.empty m n) = Some (Goto 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1414
apply(simp add: empty.simps abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1415
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1416
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1417
lemma [simp]: "abc_fetch 3 (recursive.empty m n) = None"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1418
apply(simp add: empty.simps abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1419
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1420
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1421
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1422
  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1423
    k + l = initlm ! m + initlm ! n; k \<le> initlm ! m; 0 < k\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1424
 \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, m := k - Suc 0] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1425
     initlm[m := ka, n := la] \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1426
     Suc (ka + la) = initlm ! m + initlm ! n \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1427
     ka < initlm ! m"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1428
apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1429
      simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1430
apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1431
      "initlm[m := k, n := l, m := k - Suc 0] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1432
       initlm[n := l, m := k, m := k - Suc 0]")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1433
apply(simp add: list_update_overwrite )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1434
apply(simp add: list_update_swap)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1435
apply(simp add: list_update_swap)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1436
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1437
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1438
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1439
  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1440
    Suc (k + l) = initlm ! m + initlm ! n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1441
    k < initlm ! m\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1442
    \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, n := Suc l] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1443
                initlm[m := ka, n := la] \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1444
                ka + la = initlm ! m + initlm ! n \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1445
                ka \<le> initlm ! m"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1446
apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1447
done
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]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1450
  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1451
   \<forall>na. \<not> (\<lambda>(as, lm) m. as = 3) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1452
    (abc_steps_l (0, initlm) (recursive.empty m n) na) m \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1453
  empty_inv (abc_steps_l (0, initlm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1454
           (recursive.empty m n) na) m n initlm \<longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1455
  empty_inv (abc_steps_l (0, initlm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1456
           (recursive.empty m n) (Suc na)) m n initlm \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1457
  ((abc_steps_l (0, initlm) (recursive.empty m n) (Suc na), m),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1458
   abc_steps_l (0, initlm) (recursive.empty m n) na, m) \<in> empty_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1459
apply(rule allI, rule impI, simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1460
apply(case_tac "(abc_steps_l (0, initlm) (recursive.empty m n) na)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1461
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1462
apply(auto split:if_splits simp add:abc_steps_l.simps empty_inv.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1463
apply(auto simp add: empty_LE_def lex_triple_def lex_pair_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1464
                     abc_step_l.simps abc_steps_l.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1465
                     empty_inv.simps abc_lm_v.simps abc_lm_s.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1466
                split: if_splits )
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1467
apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1468
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1469
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1470
lemma empty_inv_halt: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1471
  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1472
  \<exists> stp. (\<lambda> (as, lm). as = 3 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1473
  empty_inv (as, lm) m n initlm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1474
             (abc_steps_l (0::nat, initlm) (empty m n) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1475
apply(insert halt_lemma2[of empty_LE
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1476
  "\<lambda> ((as, lm), m). as = (3::nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1477
  "\<lambda> stp. (abc_steps_l (0, initlm) (recursive.empty m n) stp, m)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1478
  "\<lambda> ((as, lm), m). empty_inv (as, lm) m n initlm"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1479
apply(insert wf_empty_le, simp add: empty_inv_init abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1480
apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1481
apply(rule_tac x = na in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1482
apply(case_tac "(abc_steps_l (0, initlm) (recursive.empty m n) na)",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1483
      simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1484
done
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 empty_halt_cond:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1487
  "\<lbrakk>m \<noteq> n; empty_inv (a, b) m n lm; a = 3\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1488
  b = lm[n := lm ! m + lm ! n, m := 0]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1489
apply(simp add: empty_inv.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1490
apply(simp add: list_update_swap)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1491
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1492
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1493
lemma empty_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1494
  "\<lbrakk>length lm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1495
  \<exists> stp. abc_steps_l (0::nat, lm) (empty m n) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1496
  = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1497
apply(drule empty_inv_halt, simp, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1498
apply(rule_tac x = stp in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1499
apply(case_tac "abc_steps_l (0, lm) (recursive.empty m n) stp",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1500
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1501
apply(erule_tac empty_halt_cond, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1502
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1503
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1504
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1505
  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1506
   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1507
  \<Longrightarrow> n - Suc 0 < length lm + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1508
  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1509
   Suc (Suc n) < length lm + (Suc (max (Suc (Suc n)) (max bc ba)) -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1510
  rs_pos + length suf_lm) \<and> bc < length lm + (Suc (max (Suc (Suc n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1511
 (max bc ba)) - rs_pos + length suf_lm) \<and> ba < length lm + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1512
  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1513
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1514
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1515
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1516
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1517
  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1518
   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1519
 \<Longrightarrow> n - Suc 0 < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1520
     Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1521
     bc < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1522
     ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1523
apply(arith)
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1526
lemma [simp]: "n - Suc 0 \<noteq> max (Suc (Suc n)) (max bc ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1527
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1528
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1529
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1530
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1531
  "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1532
 bc - (rs_pos - Suc 0) + a_md - Suc bc = Suc (a_md - rs_pos - Suc 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1533
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1534
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1535
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1536
lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1537
                                                  Suc rs_pos < a_md 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1538
       \<Longrightarrow> n - Suc 0 < Suc (Suc (a_md + length suf_lm - Suc (Suc 0))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1539
        \<and> n < Suc (Suc (a_md + length suf_lm - Suc (Suc 0)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1540
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1541
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1542
     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1543
lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1544
               Suc rs_pos < a_md \<Longrightarrow> n - Suc 0 \<noteq> n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1545
by arith
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 ci_pr_ex2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1548
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1549
    rec_calc_rel (Pr n f g) lm rs; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1550
    rec_ci g = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1551
    rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1552
  \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1553
         ap = empty n (max (Suc (Suc (Suc n))) (max bc ba))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1554
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1555
apply(rule_tac x = "(ab [+] (recursive.empty n (Suc n) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1556
              ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1557
      [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1558
      [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1559
apply(simp add: abc_append_commute add3_Suc)
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
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1563
  "max (Suc (Suc (Suc n))) (max bc ba) - n < 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1564
     Suc (max (Suc (Suc (Suc n))) (max bc ba)) - n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1565
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1566
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1567
lemma exp_nth[simp]: "n < m \<Longrightarrow> a\<^bsup>m\<^esup> ! n = a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1568
apply(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1569
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1570
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1571
lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < n \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1572
                      lm[n - Suc 0 := 0::nat] = butlast lm @ [0]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1573
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1574
apply(insert list_update_append[of "butlast lm" "[last lm]" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1575
                                   "length lm - Suc 0" "0"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1576
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1577
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1578
lemma [simp]: "\<lbrakk>length lm = n; 0 < n\<rbrakk>  \<Longrightarrow> lm ! (n - Suc 0) = last lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1579
apply(insert nth_append[of "butlast lm" "[last lm]" "n - Suc 0"],
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1580
      simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1581
apply(insert butlast_append_last[of lm], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1582
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1583
lemma exp_suc_iff: "a\<^bsup>b\<^esup> @ [a] = a\<^bsup>b + Suc 0\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1584
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
  1585
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1586
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1587
lemma less_not_less[simp]: "n > 0 \<Longrightarrow> \<not> n < n - Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1588
by auto
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
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1591
  "Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1592
  bc < Suc (length suf_lm + max (Suc (Suc n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1593
  (max bc ba)) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1594
  ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1595
  by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1596
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1597
lemma [simp]: "length lm = n \<and> rs_pos = n \<and> n > 0 \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1598
(lm @ 0\<^bsup>Suc (max (Suc (Suc n)) (max bc ba)) - n\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1599
  [max (Suc (Suc n)) (max bc ba) :=
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1600
   (lm @ 0\<^bsup>Suc (max (Suc (Suc n)) (max bc ba)) - n\<^esup> @ suf_lm) ! (n - Suc 0) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1601
       (lm @ 0\<^bsup>Suc (max (Suc (Suc n)) (max bc ba)) - n\<^esup> @ suf_lm) ! 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1602
                   max (Suc (Suc n)) (max bc ba), n - Suc 0 := 0::nat]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1603
 = butlast lm @ 0 # 0\<^bsup>max (Suc (Suc n)) (max bc ba) - n\<^esup> @ last lm # suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1604
apply(simp add: nth_append exp_nth list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1605
apply(insert list_update_append[of "0\<^bsup>(max (Suc (Suc n)) (max bc ba)) - n\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1606
         "[0]" "max (Suc (Suc n)) (max bc ba) - n" "last lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1607
apply(simp add: exp_suc_iff Suc_diff_le del: list_update.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1608
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1609
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1610
lemma exp_eq: "(a = b) = (c\<^bsup>a\<^esup> = c\<^bsup>b\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1611
apply(auto simp: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1612
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1613
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1614
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1615
  "\<lbrakk>length lm = n; 0 < n;  Suc n < a_md\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1616
   (butlast lm @ rsa # 0\<^bsup>a_md - Suc n\<^esup> @ last lm # suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1617
    [n := (butlast lm @ rsa # 0\<^bsup>a_md - Suc n\<^esup> @ last lm # suf_lm) ! 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1618
        (n - Suc 0) + (butlast lm @ rsa # (0::nat)\<^bsup>a_md - Suc n\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1619
                                last lm # suf_lm) ! n, n - Suc 0 := 0]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1620
 = butlast lm @ 0 # rsa # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1621
apply(simp add: nth_append exp_nth list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1622
apply(case_tac "a_md - Suc n", simp, simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1623
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1624
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1625
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1626
  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1627
  \<Longrightarrow> a_md - Suc 0 < 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1628
          Suc (Suc (Suc (a_md + length suf_lm - Suc (Suc (Suc 0)))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1629
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1630
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1631
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1632
  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1633
                                   \<not> a_md - Suc 0 < rs_pos - Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1634
by arith
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 [simp]: "Suc (Suc rs_pos) \<le> a_md \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1637
                                \<not> a_md - Suc 0 < rs_pos - Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1638
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1639
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1640
lemma [simp]: "\<lbrakk>Suc (Suc rs_pos) \<le> a_md\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1641
               \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1642
by arith 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1643
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1644
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1645
  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1646
 \<Longrightarrow> (abc_lm_v (butlast lm @ last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1647
        0 # suf_lm) (a_md - Suc 0) = 0 \<longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1648
      abc_lm_s (butlast lm @ last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1649
        0 # suf_lm) (a_md - Suc 0) 0 = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1650
         lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1651
     abc_lm_v (butlast lm @ last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1652
               0 # suf_lm) (a_md - Suc 0) = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1653
apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1654
apply(insert nth_append[of "last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup>" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1655
               "0 # suf_lm" "(a_md - rs_pos)"], auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1656
apply(simp only: exp_suc_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1657
apply(subgoal_tac "a_md - Suc 0 < a_md + length suf_lm", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1658
apply(case_tac "lm = []", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1659
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1660
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1661
lemma pr_prog_ex[simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1662
      rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1663
    \<Longrightarrow> \<exists>cp. aprog = recursive.empty n (max (n + 3) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1664
                    (max bc ba)) [+] cp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1665
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1666
apply(rule_tac x = "(ab [+] (recursive.empty n (Suc n) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1667
              ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1668
             [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1669
             @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1670
apply(auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1671
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1672
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1673
lemma [simp]: "empty m n \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1674
by (simp add: empty.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1675
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1676
lemma [simp]: "\<lbrakk>rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1677
                        n - Suc 0 < a_md + length suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1678
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1679
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1680
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1681
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1682
    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1683
   \<exists>ap. (\<exists>cp. aprog = ap [+] ab [+] cp) \<and> length ap = 3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1684
apply(case_tac "rec_ci g", simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1685
apply(rule_tac x = "empty n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1686
              (max (n + 3) (max bc c))" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1687
apply(rule_tac x = "recursive.empty n (Suc n) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1688
                 ([Dec (max (n + 3) (max bc c)) (length a + 7)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1689
                 [+] a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1690
               @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1691
      auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1692
apply(simp add: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1693
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1694
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1695
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1696
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1697
    rec_ci g = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1698
    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1699
    \<exists>ap. (\<exists>cp. aprog = ap [+] recursive.empty n (Suc n) [+] cp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1700
      \<and> length ap = 3 + length ab"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1701
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1702
apply(rule_tac x = "recursive.empty n (max (n + 3)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1703
                                (max bc ba)) [+] ab" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1704
apply(rule_tac x = "([Dec (max (n + 3) (max bc ba))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1705
  (length a + 7)] [+] a [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1706
  [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1707
  [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1708
apply(auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1709
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1710
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 [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1713
  "n - Suc 0 < Suc (max (Suc (Suc n)) (max bc ba) + length suf_lm) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1714
  Suc n < max (Suc (Suc n)) (max bc ba) + length suf_lm \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1715
  bc < Suc (max (Suc (Suc n)) (max bc ba) + length suf_lm) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1716
  ba < Suc (max (Suc (Suc n)) (max bc ba) + length suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1717
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1718
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1719
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1720
lemma [intro]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1721
  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1722
    rec_ci g = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1723
    rec_ci f = (ab, ac, bc)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1724
    \<Longrightarrow> \<exists>ap. (\<exists>cp. aprog = ap [+] ([Dec (a_md - Suc 0) (length a + 7)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1725
             [+] (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1726
             Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1727
             Goto (length a + 4)] [+] cp) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1728
             length ap = 6 + length ab"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1729
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1730
apply(rule_tac x = "recursive.empty n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1731
    (max (n + 3) (max bc ba)) [+] ab [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1732
     recursive.empty n (Suc n)" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1733
apply(rule_tac x = "[]" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1734
apply(simp add: abc_append_commute)
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
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1738
lemma [simp]: "\<lbrakk>rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1739
     n - Suc 0 < Suc (Suc (a_md + length suf_lm - 2)) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1740
     n < Suc (Suc (a_md + length suf_lm - 2))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1741
by arith
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1744
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1745
  "n < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1746
   Suc (Suc n) < max (n + 3) (max bc ba) + length suf_lm \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1747
   bc < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1748
   ba < Suc (max (n + 3) (max bc ba) + length suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1749
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1750
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1751
lemma [simp]: "n \<noteq> max (n + (3::nat)) (max bc ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1752
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1753
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1754
lemma [simp]:"length lm = Suc n \<Longrightarrow> lm[n := (0::nat)] = butlast lm @ [0]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1755
apply(subgoal_tac "\<exists> xs x. lm = xs @ [x]", auto simp: list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1756
apply(rule_tac x = "butlast lm" in exI, rule_tac x = "last lm" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1757
apply(case_tac lm, auto)
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]:  "length lm = Suc n \<Longrightarrow> lm ! n =last lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1761
apply(subgoal_tac "lm \<noteq> []")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1762
apply(simp add: last_conv_nth, case_tac lm, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1763
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1764
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1765
lemma [simp]: "length lm = Suc n \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1766
      (lm @ (0::nat)\<^bsup>max (n + 3) (max bc ba) - n\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1767
           [max (n + 3) (max bc ba) := (lm @ 0\<^bsup>max (n + 3) (max bc ba) - n\<^esup> @ suf_lm) ! n + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1768
                  (lm @ 0\<^bsup>max (n + 3) (max bc ba) - n\<^esup> @ suf_lm) ! max (n + 3) (max bc ba), n := 0]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1769
       = butlast lm @ 0 # 0\<^bsup>max (n + 3) (max bc ba) - Suc n\<^esup> @ last lm # suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1770
apply(auto simp: list_update_append nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1771
apply(subgoal_tac "(0\<^bsup>max (n + 3) (max bc ba) - n\<^esup>) = 0\<^bsup>max (n + 3) (max bc ba) - Suc n\<^esup> @ [0::nat]")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1772
apply(simp add: list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1773
apply(simp add: exp_suc_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1774
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1775
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1776
lemma [simp]: "Suc (Suc n) < a_md \<Longrightarrow>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1777
      n < Suc (Suc (a_md + length suf_lm - 2)) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1778
        n < Suc (a_md + length suf_lm - 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1779
by(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1780
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1781
lemma [simp]: "\<lbrakk>length lm = Suc n; Suc (Suc n) < a_md\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1782
        \<Longrightarrow>(butlast lm @ (rsa::nat) # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1783
          [Suc n := (butlast lm @ rsa # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm) ! n +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1784
                  (butlast lm @ rsa # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ last lm # suf_lm) ! Suc n, n := 0]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1785
    = butlast lm @ 0 # rsa # 0\<^bsup>a_md - Suc (Suc (Suc n))\<^esup> @ last lm # suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1786
apply(auto simp: list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1787
apply(subgoal_tac "(0\<^bsup>a_md - Suc (Suc n)\<^esup>) = (0::nat) # (0\<^bsup>a_md - Suc (Suc (Suc n))\<^esup>)", simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1788
apply(simp add: exp_ind_def[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1789
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1790
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1791
lemma pr_case:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1792
  assumes nf_ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1793
  "\<And> lm rs suf_lm. rec_calc_rel f lm rs \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1794
  \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>bc - ac\<^esup> @ suf_lm) ab stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1795
                (length ab, lm @ rs # 0\<^bsup>bc - Suc ac\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1796
  and ng_ind: "\<And> lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1797
        \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>ba - aa\<^esup> @ suf_lm) a stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1798
                       (length a, lm @ rs # 0\<^bsup>ba - Suc aa\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1799
    and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)"  "rec_calc_rel (Pr n f g) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1800
           "rec_ci g = (a, aa, ba)" "rec_ci f = (ab, ac, bc)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1801
  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1802
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1803
  from h have k1: "\<exists> stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1804
    = (3, butlast lm @ 0 # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ last lm # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1805
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1806
    have "\<exists>bp cp. aprog = bp [+] cp \<and> bp = empty n 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1807
                 (max (n + 3) (max bc ba))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1808
      apply(insert h, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1809
      apply(erule pr_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1810
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1811
    thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1812
      apply(erule_tac exE, erule_tac exE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1813
      apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1814
           "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1815
              ([] [+] recursive.empty n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1816
                  (max (n + 3) (max bc ba)) [+] cp) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1817
             (0 + 3, butlast lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1818
                                        last lm # suf_lm)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1819
      apply(rule_tac abc_append_exc1, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1820
      apply(insert empty_ex[of "n" "(max (n + 3) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1821
                 (max bc ba))" "lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1822
      apply(subgoal_tac "a_md = Suc (max (n + 3) (max bc ba))",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1823
            simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1824
      apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1825
      apply(insert h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1826
      apply(simp add: para_pattern ci_pr_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1827
      apply(rule ci_pr_md_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1828
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1829
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1830
  from h have k2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1831
  "\<exists> stp. abc_steps_l (3,  butlast lm @ 0 # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1832
             last lm # suf_lm) aprog stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1833
    = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1834
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1835
    from h have k2_1: "\<exists> rs. rec_calc_rel f (butlast lm) rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1836
      apply(erule_tac calc_pr_zero_ex)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1837
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1838
    thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1839
    proof(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1840
      fix rsa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1841
      assume k2_2: "rec_calc_rel f (butlast lm) rsa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1842
      from h and k2_2 have k2_2_1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1843
       "\<exists> stp. abc_steps_l (3, butlast lm @ 0 # 0\<^bsup>a_md - rs_pos - 1\<^esup> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1844
                 @ last lm # suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1845
        = (3 + length ab, butlast lm @ rsa # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1846
                                             last lm # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1847
      proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1848
	from h have j1: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1849
          \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1850
              bp = ab"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1851
	  apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1852
	  done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1853
	from h have j2: "ac = rs_pos - 1"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1854
	  apply(drule_tac ci_pr_f_paras, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1855
	  done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1856
	from h and j2 have j3: "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1857
	  apply(rule_tac conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1858
	  apply(erule_tac ab = ab and ac = ac in ci_pr_md_ge_f, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1859
	  apply(rule_tac context_conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1860
          apply(simp_all add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1861
	  apply(drule_tac ci_ad_ge_paras, drule_tac ci_ad_ge_paras)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1862
	  apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1863
	  done	  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1864
	from j1 and j2 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1865
	  apply(auto simp del: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1866
	  apply(rule_tac abc_append_exc1, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1867
	  apply(insert nf_ind[of "butlast lm" "rsa" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1868
                "0\<^bsup>a_md - bc - Suc 0\<^esup> @ last lm # suf_lm"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1869
               simp add: k2_2 j2, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1870
	  apply(simp add: exponent_add_iff j3)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1871
	  apply(rule_tac x = "stp" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1872
	  done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1873
      qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1874
      from h have k2_2_2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1875
      "\<exists> stp. abc_steps_l (3 + length ab, butlast lm @ rsa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1876
                  0\<^bsup>a_md - rs_pos - 1\<^esup> @ last lm # suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1877
        = (6 + length ab, butlast lm @ 0 # rsa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1878
                       0\<^bsup>a_md - rs_pos - 2\<^esup> @ last lm # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1879
      proof -	     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1880
	from h have "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1881
          length ap = 3 + length ab \<and> bp = recursive.empty n (Suc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1882
	  by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1883
	thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1884
	proof(erule_tac exE, erule_tac exE, erule_tac exE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1885
              erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1886
	  fix ap cp bp apa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1887
	  assume "aprog = ap [+] bp [+] cp \<and> length ap = 3 + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1888
                    length ab \<and> bp = recursive.empty n (Suc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1889
	  thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1890
	    apply(simp del: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1891
	    apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1892
              "\<exists>stp. abc_steps_l (3 + length ab, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1893
               butlast lm @ rsa # 0\<^bsup>a_md - Suc rs_pos\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1894
                 last lm # suf_lm) (ap [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1895
                   recursive.empty n (Suc n) [+] cp) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1896
              ((3 + length ab) + 3, butlast lm @ 0 # rsa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1897
                  0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ last lm # suf_lm)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1898
	    apply(rule_tac abc_append_exc1, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1899
	    apply(insert empty_ex[of n "Suc n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1900
                    "butlast lm @ rsa # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1901
                          last lm # suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1902
	    apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and> a_md > Suc (Suc n)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1903
	    apply(insert h, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1904
            done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1905
	qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1906
      qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1907
      from h have k2_3: "lm \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1908
	apply(rule_tac calc_pr_para_not_null, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1909
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1910
      from h and k2_2 and k2_3 have k2_2_3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1911
      "\<exists> stp. abc_steps_l (6 + length ab, butlast lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1912
          (last lm - last lm) # rsa # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1913
            0\<^bsup>a_md - (Suc (Suc rs_pos))\<^esup> @ last lm # suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1914
        = (6 + length ab, butlast lm @ last lm # rs # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1915
                        0\<^bsup>a_md - Suc (Suc (rs_pos))\<^esup> @ 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1916
	apply(rule_tac x = "last lm" and g = g in pr_cycle_part, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1917
	apply(rule_tac ng_ind, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1918
	apply(rule_tac rec_calc_rel_def0, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1919
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1920
      from h  have k2_2_4: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1921
       "\<exists> stp. abc_steps_l (6 + length ab,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1922
             butlast lm @ last lm # rs # 0\<^bsup>a_md - rs_pos - 2\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1923
                  0 # suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1924
        = (13 + length ab + length a,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1925
                   lm @ rs # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1926
      proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1927
	from h have 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1928
        "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1929
                     length ap = 6 + length ab \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1930
                    bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1931
                         (a [+] [Inc (rs_pos - Suc 0), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1932
                         Dec rs_pos 3, Goto (Suc 0)])) @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1933
                        [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1934
	  by auto
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1935
	thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1936
	  apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1937
	  apply(subgoal_tac  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1938
            "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1939
                last lm # rs # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ 0 # suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1940
                (ap [+] ([Dec (a_md - Suc 0) (length a + 7)] [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1941
                (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1942
                Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1943
                Goto (length a + 4)] [+] cp) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1944
            (6 + length ab + (length a + 7) , 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1945
                 lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1946
	  apply(subgoal_tac "13 + (length ab + length a) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1947
                              13 + length ab + length a", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1948
	  apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1949
	  apply(rule abc_append_exc1, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1950
	  apply(rule_tac x = "Suc 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1951
                simp add: abc_steps_l.simps abc_fetch.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1952
                         nth_append abc_append_nth abc_step_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1953
	  apply(subgoal_tac "a_md > Suc (Suc rs_pos) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1954
                            length lm = rs_pos \<and> rs_pos > 0", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1955
	  apply(insert h, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1956
	  apply(subgoal_tac "rs_pos = Suc n", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1957
          done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1958
      qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1959
      from h have k2_2_5: "length aprog = 13 + length ab + length a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1960
	apply(rule_tac ci_pr_length, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1961
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1962
      from k2_2_1 and k2_2_2 and k2_2_3 and k2_2_4 and k2_2_5 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1963
      show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1964
	apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1965
	apply(rule_tac x = "stp + stpa + stpb + stpc" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1966
              simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1967
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1968
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1969
  qed	
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1970
  from k1 and k2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1971
    "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1972
               = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1973
    apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1974
    apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1975
    apply(rule_tac x = "stp + stpa" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1976
    apply(simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1977
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1978
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1979
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1980
thm rec_calc_rel.induct
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1981
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1982
lemma eq_switch: "x = y \<Longrightarrow> y = x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1983
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1984
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1985
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1986
  "\<lbrakk>rec_ci f = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1987
    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk> \<Longrightarrow> \<exists>bp. aprog = a @ bp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1988
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1989
apply(rule_tac x = "[Dec (Suc n) (length a + 5), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1990
      Dec (Suc n) (length a + 3), Goto (Suc (length a)), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1991
      Inc n, Goto 0]" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1992
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1993
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1994
lemma ci_mn_para_eq[simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1995
  "rec_ci (Mn n f) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1996
apply(case_tac "rec_ci f", simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1997
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1998
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  1999
lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> aa = Suc rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2000
apply(rule_tac calc_mn_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2001
apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2002
apply(subgoal_tac "rs_pos = length lm", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2003
apply(drule_tac ci_mn_para_eq, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2004
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2005
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2006
lemma [simp]: "rec_ci f = (a, aa, ba) \<Longrightarrow> aa < ba"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2007
apply(simp add: ci_ad_ge_paras)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2008
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2009
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2010
lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2011
                rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2012
    \<Longrightarrow> ba \<le> a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2013
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2014
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2015
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2016
lemma mn_calc_f: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2017
  assumes ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2018
  "\<And>aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2019
  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk>  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2020
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2021
           = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2022
  and h: "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2023
         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2024
         "rec_calc_rel f (lm @ [x]) rsx" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2025
         "aa = Suc n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2026
  shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2027
                  aprog stp = (length a, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2028
                   lm @ x # rsx # 0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2029
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2030
  from h have k1: "\<exists> ap bp. aprog = ap @ bp \<and> ap = a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2031
    by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2032
  from h have k2: "rs_pos = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2033
    apply(erule_tac ci_mn_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2034
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2035
  from h and k1 and k2 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2036
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2037
  proof(erule_tac exE, erule_tac exE, simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2038
        rule_tac abc_add_exc1, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2039
    fix bp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2040
    show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2041
      "\<exists>astp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - Suc n\<^esup> @ suf_lm) a astp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2042
      = (length a, lm @ x # rsx # 0\<^bsup>a_md - Suc (Suc n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2043
      apply(insert ind[of a "Suc n" ba  "lm @ [x]" rsx 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2044
             "0\<^bsup>a_md - ba\<^esup> @ suf_lm"], simp add: exponent_add_iff h k2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2045
      apply(subgoal_tac "ba > aa \<and> a_md \<ge> ba \<and> aa = Suc n", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2046
            insert h, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2047
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2048
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2049
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2050
thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2051
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2052
fun mn_ind_inv ::
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2053
  "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2054
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2055
  "mn_ind_inv (as, lm') ss x rsx suf_lm lm = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2056
           (if as = ss then lm' = lm @ x # rsx # suf_lm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2057
            else if as = ss + 1 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2058
                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2059
            else if as = ss + 2 then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2060
                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2061
            else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2062
            else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2063
            else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2064
            else False
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2065
)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2066
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2067
fun mn_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2068
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2069
  "mn_stage1 (as, lm) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2070
            (if as = 0 then 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2071
             else if as = ss + 4 then 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2072
             else if as = ss + 3 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2073
             else if as = ss + 2 \<or> as = ss + 1 then 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2074
             else if as = ss then 4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2075
             else 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2076
)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2077
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2078
fun mn_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2079
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2080
  "mn_stage2 (as, lm) ss n = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2081
            (if as = ss + 1 \<or> as = ss + 2 then (lm ! (Suc n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2082
             else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2083
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2084
fun mn_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2085
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2086
  "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2087
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2088
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2089
fun mn_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2090
                                                (nat \<times> nat \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2091
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2092
  "mn_measure ((as, lm), ss, n) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2093
     (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2094
                                       mn_stage3 (as, lm) ss n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2095
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2096
definition mn_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2097
                     ((nat \<times> nat list) \<times> nat \<times> nat)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2098
  where "mn_LE \<equiv> (inv_image lex_triple mn_measure)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2099
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2100
thm halt_lemma2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2101
lemma wf_mn_le[intro]: "wf mn_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2102
by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2103
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2104
declare mn_ind_inv.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2105
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2106
lemma mn_inv_init: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2107
  "mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2108
                                         (length a) x rsx suf_lm lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2109
apply(simp add: mn_ind_inv.simps abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2110
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2111
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2112
lemma mn_halt_init: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2113
  "rec_ci f = (a, aa, ba) \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2114
  \<not> (\<lambda>(as, lm') (ss, n). as = 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2115
    (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2116
                                                       (length a, n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2117
apply(simp add: abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2118
apply(erule_tac rec_ci_not_null)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2119
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2120
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2121
thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2122
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2123
  "\<lbrakk>rec_ci f = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2124
    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2125
    \<Longrightarrow> abc_fetch (length a) aprog =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2126
                      Some (Dec (Suc n) (length a + 5))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2127
apply(simp add: rec_ci.simps abc_fetch.simps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2128
                erule_tac conjE, erule_tac conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2129
apply(drule_tac eq_switch, drule_tac eq_switch, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2130
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2131
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2132
lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2133
    \<Longrightarrow> abc_fetch (Suc (length a)) aprog = Some (Dec (Suc n) (length a + 3))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2134
apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2135
apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2136
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2137
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2138
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2139
  "\<lbrakk>rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2140
    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2141
    \<Longrightarrow> abc_fetch (Suc (Suc (length a))) aprog = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2142
                                     Some (Goto (length a + 1))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2143
apply(simp add: rec_ci.simps abc_fetch.simps,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2144
      erule_tac conjE, erule_tac conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2145
apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2146
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2147
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2148
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2149
  "\<lbrakk>rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2150
    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2151
    \<Longrightarrow> abc_fetch (length a + 3) aprog = Some (Inc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2152
apply(simp add: rec_ci.simps abc_fetch.simps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2153
      erule_tac conjE, erule_tac conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2154
apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2155
done
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
lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2158
    \<Longrightarrow> abc_fetch (length a + 4) aprog = Some (Goto 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2159
apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2160
apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
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]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2164
  "0 < rsx
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2165
   \<Longrightarrow> \<exists>y. (lm @ x # rsx # suf_lm)[Suc (length lm) := rsx - Suc 0]   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2166
    = lm @ x # y # suf_lm \<and> y \<le> rsx"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2167
apply(case_tac rsx, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2168
apply(rule_tac x = nat in exI, simp add: list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2169
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2170
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2171
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2172
  "\<lbrakk>y \<le> rsx; 0 < y\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2173
   \<Longrightarrow> \<exists>ya. (lm @ x # y # suf_lm)[Suc (length lm) := y - Suc 0] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2174
          = lm @ x # ya # suf_lm \<and> ya \<le> rsx"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2175
apply(case_tac y, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2176
apply(rule_tac x = nat in exI, simp add: list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2177
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2178
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2179
lemma mn_halt_lemma: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2180
  "\<lbrakk>rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2181
    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2182
     0 < rsx; length lm = n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2183
    \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2184
  \<forall>na. \<not> (\<lambda>(as, lm') (ss, n). as = 0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2185
  (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2186
                                                       (length a, n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2187
 \<and> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2188
                       aprog na) (length a) x rsx suf_lm lm 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2189
\<longrightarrow> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2190
                         (Suc na)) (length a) x rsx suf_lm lm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2191
 \<and> ((abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog (Suc na), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2192
                                                    length a, n), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2193
    abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2194
                              length a, n) \<in> mn_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2195
apply(rule allI, rule impI, simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2196
apply(case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2197
                                                   aprog na)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2198
apply(auto split:if_splits simp add:abc_steps_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2199
                           mn_ind_inv.simps abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2200
apply(auto simp add: mn_LE_def lex_triple_def lex_pair_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2201
            abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2202
            abc_lm_v.simps abc_lm_s.simps nth_append
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2203
           split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2204
apply(drule_tac  rec_ci_not_null, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2205
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2206
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2207
lemma mn_halt:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2208
  "\<lbrakk>rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2209
    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2210
    0 < rsx; length lm = n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2211
    \<Longrightarrow> \<exists> stp. (\<lambda> (as, lm'). (as = 0 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2212
           mn_ind_inv (as, lm')  (length a) x rsx suf_lm lm))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2213
            (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2214
apply(insert wf_mn_le)	  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2215
apply(insert halt_lemma2[of mn_LE
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2216
  "\<lambda> ((as, lm'), ss, n). as = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2217
  "\<lambda> stp. (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2218
   length a, n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2219
   "\<lambda> ((as, lm'), ss, n). mn_ind_inv (as, lm') ss x rsx suf_lm lm"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2220
   simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2221
apply(simp add: mn_halt_init mn_inv_init)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2222
apply(drule_tac x = x and suf_lm = suf_lm in mn_halt_lemma, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2223
apply(rule_tac x = n in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2224
      case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2225
                              aprog n)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2226
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2227
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2228
lemma [simp]: "Suc rs_pos < a_md \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2229
                Suc (a_md - Suc (Suc rs_pos)) = a_md - Suc rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2230
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2231
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2232
term rec_ci
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2233
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2234
lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2235
apply(case_tac "rec_ci f")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2236
apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2237
apply(arith, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2238
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2239
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2240
lemma mn_ind_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2241
  assumes ind:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2242
  "\<And>aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2243
  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2244
   rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2245
  \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2246
            = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2247
  and h: "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2248
         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2249
         "rec_calc_rel f (lm @ [x]) rsx" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2250
         "rsx > 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2251
         "Suc rs_pos < a_md" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2252
         "aa = Suc rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2253
  shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2254
             aprog stp = (0, lm @ Suc x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2255
thm abc_add_exc1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2256
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2257
  have k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2258
    "\<exists> stp. abc_steps_l (0, lm @ x #  0\<^bsup>a_md - Suc (rs_pos)\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2259
         aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2260
       (length a, lm @ x # rsx # 0\<^bsup>a_md  - Suc (Suc rs_pos) \<^esup>@ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2261
    apply(insert h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2262
    apply(auto intro: mn_calc_f ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2263
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2264
  from h have k2: "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2265
    apply(subgoal_tac "rs_pos = n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2266
    apply(drule_tac  para_pattern, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2267
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2268
  from h have k3: "a_md > (Suc rs_pos)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2269
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2270
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2271
  from k2 and h and k3 have k4: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2272
    "\<exists> stp. abc_steps_l (length a,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2273
       lm @ x # rsx # 0\<^bsup>a_md  - Suc (Suc rs_pos)  \<^esup>@ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2274
        (0, lm @ Suc x # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2275
    apply(frule_tac x = x and 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2276
       suf_lm = "0\<^bsup>a_md - Suc (Suc rs_pos)\<^esup> @ suf_lm" in mn_halt, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2277
    apply(rule_tac x = "stp" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2278
          simp add: mn_ind_inv.simps rec_ci_not_null exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2279
    apply(simp only: replicate.simps[THEN sym], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2280
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2281
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2282
  from k1 and k4 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2283
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2284
    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2285
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2286
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2287
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2288
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2289
  "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2290
    rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> aa = Suc rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2291
apply(rule_tac calc_mn_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2292
apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2293
apply(subgoal_tac "rs_pos = length lm", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2294
apply(drule_tac ci_mn_para_eq, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2295
done
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
lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);      
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2298
                rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2299
apply(case_tac "rec_ci f")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2300
apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2301
apply(arith, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2302
done
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
lemma mn_ind_steps:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2305
  assumes ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2306
  "\<And>aprog a_md rs_pos rs suf_lm lm. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2307
  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2308
  \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2309
              (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2310
  and h: "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2311
  "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2312
  "rec_calc_rel (Mn n f) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2313
  "rec_calc_rel f (lm @ [rs]) 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2314
  "\<forall>x<rs. (\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2315
  "n = length lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2316
  "x \<le> rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2317
  shows "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2318
                 aprog stp = (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2319
apply(insert h, induct x, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2320
      rule_tac x = 0 in exI, simp add: abc_steps_zero, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2321
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2322
  fix x
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2323
  assume k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2324
    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2325
                aprog stp = (0, lm @ x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2326
  and k2: "rec_ci (Mn (length lm) f) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2327
          "rec_calc_rel (Mn (length lm) f) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2328
          "rec_calc_rel f (lm @ [rs]) 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2329
          "\<forall>x<rs.(\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> v > 0)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2330
          "n = length lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2331
          "Suc x \<le> rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2332
          "rec_ci f = (a, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2333
  hence k2:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2334
    "\<exists>stp. abc_steps_l (0, lm @ x # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm) aprog
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2335
               stp = (0, lm @ Suc x # 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2336
    apply(erule_tac x = x in allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2337
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2338
    apply(rule_tac  x = x in mn_ind_step)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2339
    apply(rule_tac ind, auto)      
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2340
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2341
  from k1 and k2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2342
    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2343
          aprog stp = (0, lm @ Suc x # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2344
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2345
    apply(rule_tac x = "stp + stpa" in exI, simp only: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2346
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2347
qed
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 [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2350
"\<lbrakk>rec_ci f = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2351
  rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2352
  rec_calc_rel (Mn n f) lm rs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2353
  length lm = n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2354
 \<Longrightarrow> abc_lm_v (lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) (Suc n) = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2355
apply(auto simp: abc_lm_v.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2356
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2357
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2358
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2359
  "\<lbrakk>rec_ci f = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2360
    rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2361
    rec_calc_rel (Mn n f) lm rs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2362
     length lm = n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2363
    \<Longrightarrow> abc_lm_s (lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) (Suc n) 0 =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2364
                           lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2365
apply(auto simp: abc_lm_s.simps list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2366
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2367
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2368
lemma mn_length: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2369
  "\<lbrakk>rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2370
    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2371
  \<Longrightarrow> length aprog = length a + 5"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2372
apply(simp add: rec_ci.simps, erule_tac conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2373
apply(drule_tac eq_switch, drule_tac eq_switch, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2374
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2375
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2376
lemma mn_final_step:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2377
  assumes ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2378
  "\<And>aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2379
  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2380
  rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2381
  \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2382
              (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2383
  and h: "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2384
         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2385
         "rec_calc_rel (Mn n f) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2386
         "rec_calc_rel f (lm @ [rs]) 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2387
  shows "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2388
     aprog stp = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2389
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2390
  from h and ind have k1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2391
    "\<exists>stp.  abc_steps_l (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2392
        aprog stp = (length a,  lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2393
    thm mn_calc_f
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2394
    apply(insert mn_calc_f[of f a aa ba n aprog 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2395
                               rs_pos a_md lm rs 0 suf_lm], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2396
    apply(subgoal_tac "aa = Suc n", simp add: exponent_cons_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2397
    apply(subgoal_tac "rs_pos = n", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2398
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2399
  from h have k2: "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2400
    apply(subgoal_tac "rs_pos = n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2401
    apply(drule_tac f = "Mn n f" in para_pattern, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2402
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2403
  from h and k2 have k3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2404
  "\<exists>stp. abc_steps_l (length a, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2405
    aprog stp = (length a + 5, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2406
    apply(rule_tac x = "Suc 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2407
          simp add: abc_step_l.simps abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2408
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2409
  from h have k4: "length aprog = length a + 5"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2410
    apply(simp add: mn_length)
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 k1 and k3 and k4 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2413
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2414
    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2415
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2416
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2417
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2418
lemma mn_case: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2419
  assumes ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2420
  "\<And>aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2421
  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2422
  \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2423
               (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2424
  and h: "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2425
         "rec_calc_rel (Mn n f) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2426
  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2427
  = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2428
apply(case_tac "rec_ci f", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2429
apply(insert h, rule_tac calc_mn_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2430
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2431
  fix a b c v
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2432
  assume h: "rec_ci f = (a, b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2433
            "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2434
            "rec_calc_rel (Mn n f) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2435
            "rec_calc_rel f (lm @ [rs]) 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2436
            "\<forall>x<rs. \<exists>v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2437
            "n = length lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2438
  hence k1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2439
    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) aprog
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2440
                  stp = (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2441
    thm mn_ind_steps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2442
    apply(auto intro: mn_ind_steps ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2443
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2444
  from h have k2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2445
    "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) aprog
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2446
         stp = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2447
    apply(auto intro: mn_final_step ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2448
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2449
  from k1 and k2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2450
    "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2451
  (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2452
    apply(auto, insert h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2453
    apply(subgoal_tac "Suc rs_pos < a_md")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2454
    apply(rule_tac x = "stp + stpa" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2455
      simp only: abc_steps_add exponent_cons_iff, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2456
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2457
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2458
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2459
lemma z_rs: "rec_calc_rel z lm rs \<Longrightarrow> rs = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2460
apply(rule_tac calc_z_reverse, auto)
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 z_case:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2464
  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2465
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2466
           (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2467
apply(simp add: rec_ci.simps rec_ci_z_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2468
apply(rule_tac x = "Suc 0" in exI, simp add: abc_steps_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2469
                               abc_fetch.simps abc_step_l.simps z_rs)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2470
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2471
thm addition.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2472
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2473
thm addition.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2474
thm rec_ci_s_def
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2475
fun addition_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2476
                     nat list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2477
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2478
  "addition_inv (as, lm') m n p lm = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2479
        (let sn = lm ! n in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2480
         let sm = lm ! m in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2481
         lm ! p = 0 \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2482
             (if as = 0 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2483
                                    n := (sn + sm - x), p := (sm - x)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2484
             else if as = 1 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2485
                            n := (sn + sm - x - 1), p := (sm - x - 1)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2486
             else if as = 2 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2487
                               n := (sn + sm - x), p := (sm - x - 1)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2488
             else if as = 3 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2489
                                   n := (sn + sm - x), p := (sm - x)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2490
             else if as = 4 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2491
                                       n := (sn + sm), p := (sm - x)] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2492
             else if as = 5 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2493
                                  n := (sn + sm), p := (sm - x - 1)] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2494
             else if as = 6 then \<exists> x. x < lm ! m \<and> lm' =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2495
                     lm[m := Suc x, n := (sn + sm), p := (sm - x - 1)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2496
             else if as = 7 then lm' = lm[m := sm, n := (sn + sm)]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2497
             else False))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2498
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2499
fun addition_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2500
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2501
  "addition_stage1 (as, lm) m p = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2502
          (if as = 0 \<or> as = 1 \<or> as = 2 \<or> as = 3 then 2 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2503
           else if as = 4 \<or> as = 5 \<or> as = 6 then 1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2504
           else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2505
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2506
fun addition_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow>  nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2507
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2508
  "addition_stage2 (as, lm) m p = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2509
              (if 0 \<le> as \<and> as \<le> 3 then lm ! m
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2510
               else if 4 \<le> as \<and> as \<le> 6 then lm ! p
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2511
               else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2512
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2513
fun addition_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2514
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2515
  "addition_stage3 (as, lm) m p = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2516
             (if as = 1 then 4  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2517
              else if as = 2 then 3 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2518
              else if as = 3 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2519
              else if as = 0 then 1 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2520
              else if as = 5 then 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2521
              else if as = 6 then 1 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2522
              else if as = 4 then 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2523
              else 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2524
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2525
fun addition_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2526
                                                 (nat \<times> nat \<times> nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2527
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2528
  "addition_measure ((as, lm), m, p) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2529
     (addition_stage1 (as, lm) m p, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2530
      addition_stage2 (as, lm) m p,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2531
      addition_stage3 (as, lm) m p)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2532
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2533
definition addition_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2534
                          ((nat \<times> nat list) \<times> nat \<times> nat)) set"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2535
  where "addition_LE \<equiv> (inv_image lex_triple addition_measure)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2536
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2537
lemma [simp]: "wf addition_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2538
by(simp add: wf_inv_image wf_lex_triple addition_LE_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2539
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2540
declare addition_inv.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2541
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2542
lemma addition_inv_init: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2543
  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2544
                                   addition_inv (0, lm) m n p lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2545
apply(simp add: addition_inv.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2546
apply(rule_tac x = "lm ! m" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2547
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2548
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2549
thm addition.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2550
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2551
lemma [simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2552
by(simp add: abc_fetch.simps addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2553
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2554
lemma [simp]: "abc_fetch (Suc 0) (addition m n p) = Some (Inc n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2555
by(simp add: abc_fetch.simps addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2556
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2557
lemma [simp]: "abc_fetch 2 (addition m n p) = Some (Inc p)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2558
by(simp add: abc_fetch.simps addition.simps)
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
lemma [simp]: "abc_fetch 3 (addition m n p) = Some (Goto 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2561
by(simp add: abc_fetch.simps addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2562
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2563
lemma [simp]: "abc_fetch 4 (addition m n p) = Some (Dec p 7)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2564
by(simp add: abc_fetch.simps addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2565
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2566
lemma [simp]: "abc_fetch 5 (addition m n p) = Some (Inc m)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2567
by(simp add: abc_fetch.simps addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2568
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2569
lemma [simp]: "abc_fetch 6 (addition m n p) = Some (Goto 4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2570
by(simp add: abc_fetch.simps addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2571
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2572
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2573
  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x \<le> lm ! m; 0 < x\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2574
 \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2575
                    p := lm ! m - x, m := x - Suc 0] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2576
                 lm[m := xa, n := lm ! n + lm ! m - Suc xa,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2577
                    p := lm ! m - Suc xa]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2578
apply(case_tac x, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2579
apply(rule_tac x = nat in exI, simp add: list_update_swap 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2580
                                         list_update_overwrite)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2581
done
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
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2584
  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2585
   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - Suc x,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2586
                      p := lm ! m - Suc x, n := lm ! n + lm ! m - x]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2587
                 = lm[m := xa, n := lm ! n + lm ! m - xa, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2588
                      p := lm ! m - Suc xa]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2589
apply(rule_tac x = x in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2590
      simp add: list_update_swap list_update_overwrite)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2591
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2592
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2593
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2594
  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2595
   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2596
                          p := lm ! m - Suc x, p := lm ! m - x]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2597
                 = lm[m := xa, n := lm ! n + lm ! m - xa, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2598
                          p := lm ! m - xa]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2599
apply(rule_tac x = x in exI, simp add: list_update_overwrite)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2600
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2601
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2602
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2603
  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2604
  \<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := x, n := lm ! n + lm ! m - x,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2605
                                   p := lm ! m - x] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2606
                  lm[m := xa, n := lm ! n + lm ! m - xa, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2607
                                   p := lm ! m - xa]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2608
apply(rule_tac x = x in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2609
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2610
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2611
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2612
  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2613
    x \<le> lm ! m; lm ! m \<noteq> x\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2614
  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2615
                       p := lm ! m - x, p := lm ! m - Suc x] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2616
               = lm[m := xa, n := lm ! n + lm ! m, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2617
                       p := lm ! m - Suc xa]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2618
apply(rule_tac x = x in exI, simp add: list_update_overwrite)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2619
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2620
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2621
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2622
  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2623
  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2624
                             p := lm ! m - Suc x, m := Suc x]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2625
                = lm[m := Suc xa, n := lm ! n + lm ! m, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2626
                             p := lm ! m - Suc xa]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2627
apply(rule_tac x = x in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2628
     simp add: list_update_swap list_update_overwrite)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2629
done
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
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2632
  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2633
  \<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := Suc x, n := lm ! n + lm ! m, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2634
                             p := lm ! m - Suc x] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2635
               = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2636
apply(rule_tac x = "Suc x" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2637
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2638
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2639
lemma addition_halt_lemma: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2640
  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2641
  \<forall>na. \<not> (\<lambda>(as, lm') (m, p). as = 7) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2642
        (abc_steps_l (0, lm) (addition m n p) na) (m, p) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2643
  addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2644
\<longrightarrow> addition_inv (abc_steps_l (0, lm) (addition m n p) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2645
                                 (Suc na)) m n p lm 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2646
  \<and> ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2647
     abc_steps_l (0, lm) (addition m n p) na, m, p) \<in> addition_LE"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2648
apply(rule allI, rule impI, simp add: abc_steps_ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2649
apply(case_tac "(abc_steps_l (0, lm) (addition m n p) na)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2650
apply(auto split:if_splits simp add: addition_inv.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2651
                                 abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2652
apply(simp_all add: abc_steps_l.simps abc_steps_zero)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2653
apply(auto simp add: addition_LE_def lex_triple_def lex_pair_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2654
                     abc_step_l.simps addition_inv.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2655
                     abc_lm_v.simps abc_lm_s.simps nth_append
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2656
                split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2657
apply(rule_tac x = x in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2658
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2659
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2660
lemma  addition_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2661
  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2662
  \<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2663
                        (abc_steps_l (0, lm) (addition m n p) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2664
apply(insert halt_lemma2[of addition_LE 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2665
  "\<lambda> ((as, lm'), m, p). as = 7"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2666
  "\<lambda> stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2667
  "\<lambda> ((as, lm'), m, p). addition_inv (as, lm') m n p lm"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2668
  simp add: abc_steps_zero addition_inv_init)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2669
apply(drule_tac addition_halt_lemma, simp, simp, simp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2670
      simp, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2671
apply(rule_tac x = na in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2672
      case_tac "(abc_steps_l (0, lm) (addition m n p) na)", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2673
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2674
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2675
lemma [simp]: "length (addition m n p) = 7"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2676
by (simp add: addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2677
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2678
lemma [elim]: "addition 0 (Suc 0) 2 = [] \<Longrightarrow> RR"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2679
by(simp add: addition.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2680
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2681
lemma [simp]: "(0\<^bsup>2\<^esup>)[0 := n] = [n, 0::nat]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2682
apply(subgoal_tac "2 = Suc 1", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2683
      simp only: replicate.simps exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2684
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2685
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2686
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2687
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2688
  "\<exists>stp. abc_steps_l (0, n # 0\<^bsup>2\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2689
     (addition 0 (Suc 0) 2 [+] [Inc (Suc 0)]) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2690
                                      (8, n # Suc n # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2691
apply(rule_tac bm = "n # n # 0 # suf_lm" in abc_append_exc2, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2692
apply(insert addition_ex[of 0 "Suc 0" 2 "n # 0\<^bsup>2\<^esup> @ suf_lm"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2693
      simp add: nth_append numeral_2_eq_2, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2694
apply(rule_tac x = stp in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2695
      case_tac "(abc_steps_l (0, n # 0\<^bsup>2\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2696
                      (addition 0 (Suc 0) 2) stp)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2697
      simp add: addition_inv.simps nth_append list_update_append numeral_2_eq_2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2698
apply(simp add: nth_append numeral_2_eq_2, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2699
apply(rule_tac x = "Suc 0" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2700
      simp add: abc_steps_l.simps abc_fetch.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2701
      abc_steps_zero abc_step_l.simps abc_lm_s.simps abc_lm_v.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2702
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2703
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2704
lemma s_case:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2705
  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2706
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2707
               (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2708
apply(simp add: rec_ci.simps rec_ci_s_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2709
apply(rule_tac calc_s_reverse, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2710
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2711
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2712
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2713
  "\<lbrakk>n < length lm; lm ! n = rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2714
    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0 # 0 #suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2715
                     (addition n (length lm) (Suc (length lm))) stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2716
             = (7, lm @ rs # 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2717
apply(insert addition_ex[of n "length lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2718
                           "Suc (length lm)" "lm @ 0 # 0 # suf_lm"])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2719
apply(simp add: nth_append, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2720
apply(rule_tac x = stp in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2721
apply(case_tac "abc_steps_l (0, lm @ 0 # 0 # suf_lm) (addition n (length lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2722
                 (Suc (length lm))) stp", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2723
apply(simp add: addition_inv.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2724
apply(insert nth_append[of lm "0 # 0 # suf_lm" "n"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2725
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2726
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2727
lemma [simp]: "0\<^bsup>2\<^esup> = [0, 0::nat]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2728
apply(auto simp: exponent_def numeral_2_eq_2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2729
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2730
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2731
lemma id_case: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2732
  "\<lbrakk>rec_ci (id m n) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2733
    rec_calc_rel (id m n) lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2734
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2735
               (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2736
apply(simp add: rec_ci.simps rec_ci_id.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2737
apply(rule_tac calc_id_reverse, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2738
done   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2739
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2740
lemma list_tl_induct:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2741
  "\<lbrakk>P []; \<And>a list. P list \<Longrightarrow> P (list @ [a::'a])\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2742
                                            P ((list::'a list))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2743
apply(case_tac "length list", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2744
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2745
  fix nat
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2746
  assume ind: "\<And>a list. P list \<Longrightarrow> P (list @ [a])"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2747
  and h: "length list = Suc nat" "P []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2748
  from h show "P list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2749
  proof(induct nat arbitrary: list, case_tac lista, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2750
    fix lista a listaa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2751
    from h show "P [a]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2752
      by(insert ind[of "[]"], simp add: h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2753
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2754
    fix nat list
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2755
    assume nind: "\<And>list. \<lbrakk>length list = Suc nat; P []\<rbrakk> \<Longrightarrow> P list" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2756
    and g: "length (list:: 'a list) = Suc (Suc nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2757
    from g show "P (list::'a list)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2758
      apply(insert nind[of "butlast list"], simp add: h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2759
      apply(insert ind[of "butlast list" "last list"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2760
      apply(subgoal_tac "butlast list @ [last list] = list", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2761
      apply(case_tac "list::'a list", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2762
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2763
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2764
qed      
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2765
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2766
thm list.induct
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2767
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2768
lemma nth_eq_butlast_nth: "\<lbrakk>length ys > Suc k\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2769
                                        ys ! k = butlast ys ! k"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2770
apply(subgoal_tac "\<exists> xs y. ys = xs @ [y]", auto simp: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2771
apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2772
apply(case_tac "ys = []", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2773
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2774
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2775
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2776
"\<lbrakk>\<forall>k<Suc (length list). rec_calc_rel ((list @ [a]) ! k) lm (ys ! k);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2777
  length ys = Suc (length list)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2778
   \<Longrightarrow> \<forall>k<length list. rec_calc_rel (list ! k) lm (butlast ys ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2779
apply(rule allI, rule impI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2780
apply(erule_tac  x = k in allE, simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2781
apply(subgoal_tac "ys ! k = butlast ys ! k", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2782
apply(rule_tac nth_eq_butlast_nth, arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2783
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2784
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2785
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2786
thm cn_merge_gs.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2787
lemma cn_merge_gs_tl_app: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2788
  "cn_merge_gs (gs @ [g]) pstr = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2789
        cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2790
apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2791
apply(case_tac a, simp add: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2792
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2793
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2794
lemma cn_merge_gs_length: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2795
  "length (cn_merge_gs (map rec_ci list) pstr) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2796
      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2797
apply(induct list arbitrary: pstr, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2798
apply(case_tac "rec_ci a", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2799
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2800
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2801
lemma [simp]: "Suc n \<le> pstr \<Longrightarrow> pstr + x - n > 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2802
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2803
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2804
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2805
  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2806
    length ys = Suc (length list);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2807
    length lm = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2808
     Suc n \<le> pstr\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2809
   \<Longrightarrow>  (ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2810
             0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) ! 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2811
                      (pstr + length list - n) = (0 :: nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2812
apply(insert nth_append[of "ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2813
     butlast ys" "0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2814
      "(pstr + length list - n)"], simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2815
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2816
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2817
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2818
  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2819
    length ys = Suc (length list);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2820
    length lm = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2821
     Suc n \<le> pstr\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2822
    \<Longrightarrow> (lm @ last ys # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2823
         0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm)[pstr + length list := 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2824
                                        last ys, n := 0] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2825
        lm @ 0::nat\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2826
apply(insert list_update_length[of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2827
   "lm @ last ys # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys" 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2828
   "0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm" "last ys"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2829
apply(simp add: exponent_cons_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2830
apply(insert list_update_length[of "lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2831
        "last ys" "0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2832
      last ys # 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm" 0], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2833
apply(simp add: exponent_cons_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2834
apply(case_tac "ys = []", simp_all add: append_butlast_last_id)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2835
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2836
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2837
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2838
lemma cn_merge_gs_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2839
  "\<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2840
    \<lbrakk>x \<in> set gs; rec_ci x = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2841
     rec_calc_rel x lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2842
     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2843
           = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2844
   pstr + length gs\<le> a_md;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2845
   \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2846
   length ys = length gs; length lm = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2847
   pstr \<ge> Max (set (Suc n # map (\<lambda>(aprog, p, n). n) (map rec_ci gs)))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2848
  \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2849
                   (cn_merge_gs (map rec_ci gs) pstr) stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2850
   = (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2851
  3 * length gs, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - (pstr + length gs)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2852
apply(induct gs arbitrary: ys rule: list_tl_induct)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2853
apply(simp add: exponent_add_iff, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2854
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2855
  fix a list ys
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2856
  assume ind: "\<And>x aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2857
    \<lbrakk>x = a \<or> x \<in> set list; rec_ci x = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2858
     rec_calc_rel x lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2859
    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2860
                (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2861
  and ind2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2862
    "\<And>ys. \<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2863
    \<lbrakk>x \<in> set list; rec_ci x = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2864
     rec_calc_rel x lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2865
    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2866
        = (length aprog, lm @ rs # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2867
    \<forall>k<length list. rec_calc_rel (list ! k) lm (ys ! k); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2868
    length ys = length list\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2869
    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2870
                   (cn_merge_gs (map rec_ci list) pstr) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2871
    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2872
     3 * length list,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2873
                lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2874
    and h: "Suc (pstr + length list) \<le> a_md" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2875
            "\<forall>k<Suc (length list). 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2876
                   rec_calc_rel ((list @ [a]) ! k) lm (ys ! k)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2877
            "length ys = Suc (length list)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2878
            "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2879
            "Suc n \<le> pstr \<and> (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2880
            (\<forall>a\<in>set list. (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2881
  from h have k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2882
    "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2883
                     (cn_merge_gs (map rec_ci list) pstr) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2884
    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2885
     3 * length list, lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2886
                               0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2887
    apply(rule_tac ind2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2888
    apply(rule_tac ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2889
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2890
  from k1 and h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2891
    "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2892
          (cn_merge_gs (map rec_ci list @ [rec_ci a]) pstr) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2893
        (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2894
        (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2895
             lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2896
    apply(simp add: cn_merge_gs_tl_app)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2897
    thm abc_append_exc2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2898
    apply(rule_tac as = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2899
  "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list"    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2900
      and bm = "lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2901
                              0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2902
      and bs = "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2903
      and bm' = "lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2904
                                  suf_lm" in abc_append_exc2, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2905
    apply(simp add: cn_merge_gs_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2906
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2907
    from h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2908
      "\<exists>bstp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2909
                                  0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2910
              ((\<lambda>(gprog, gpara, gn). gprog [+] recursive.empty gpara 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2911
              (pstr + length list)) (rec_ci a)) bstp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2912
              ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2913
             lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2914
      apply(case_tac "rec_ci a", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2915
      apply(rule_tac as = "length aa" and 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2916
                     bm = "lm @ (ys ! (length list)) # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2917
          0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2918
        and bs = "3" and bm' = "lm @ 0\<^bsup>pstr - n\<^esup> @ ys @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2919
             0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm" in abc_append_exc2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2920
    proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2921
      fix aa b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2922
      assume g: "rec_ci a = (aa, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2923
      from h and g have k2: "b = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2924
	apply(erule_tac x = "length list" in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2925
	apply(subgoal_tac "length lm = b", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2926
	apply(rule para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2927
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2928
      from h and g and this show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2929
        "\<exists>astp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ butlast ys @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2930
                         0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm) aa astp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2931
        (length aa, lm @ ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2932
                       butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2933
	apply(subgoal_tac "c \<ge> Suc n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2934
	apply(insert ind[of a aa b c lm "ys ! length list" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2935
     "0\<^bsup>pstr - c\<^esup> @ butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2936
	apply(erule_tac x = "length list" in allE, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2937
              simp add: exponent_add_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2938
	apply(rule_tac Suc_leI, rule_tac ci_ad_ge_paras, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2939
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2940
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2941
      fix aa b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2942
      show "length aa = length aa" by simp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2943
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2944
      fix aa b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2945
      assume "rec_ci a = (aa, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2946
      from h and this show     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2947
      "\<exists>bstp. abc_steps_l (0, lm @ ys ! length list #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2948
          0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ 0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2949
                 (recursive.empty b (pstr + length list)) bstp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2950
       (3, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - Suc (pstr + length list)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2951
	apply(insert empty_ex [of b "pstr + length list" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2952
         "lm @ ys ! length list # 0\<^bsup>pstr - Suc n\<^esup> @ butlast ys @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2953
         0\<^bsup>a_md - (pstr + length list)\<^esup> @ suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2954
        apply(subgoal_tac "b = n")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2955
	apply(simp add: nth_append split: if_splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2956
	apply(erule_tac x = "length list" in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2957
        apply(drule para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2958
	done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2959
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2960
      fix aa b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2961
      show "3 = length (recursive.empty b (pstr + length list))" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2962
        by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2963
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2964
      fix aa b aaa ba
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2965
      show "length aa + 3 = length aa + 3" by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2966
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2967
      fix aa b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2968
      show "empty b (pstr + length list) \<noteq> []" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2969
        by(simp add: empty.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2970
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2971
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2972
    show "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3 = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2973
        length ((\<lambda>(gprog, gpara, gn). gprog [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2974
           recursive.empty gpara (pstr + length list)) (rec_ci a))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2975
      by(case_tac "rec_ci a", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2976
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2977
    show "listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2978
      (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list)=
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2979
      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2980
                ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3)" by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2981
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2982
    show "(\<lambda>(gprog, gpara, gn). gprog [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2983
      recursive.empty gpara (pstr + length list)) (rec_ci a) \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2984
      by(case_tac "rec_ci a", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2985
         simp add: abc_append.simps abc_shift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2986
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2987
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2988
   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2989
declare drop_abc_lm_v_simp[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2990
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2991
lemma [simp]: "length (mv_boxes aa ba n) = 3*n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2992
by(induct n, auto simp: mv_boxes.simps)
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 exp_suc: "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
  2995
by(simp add: exponent_def rep_ind del: replicate.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2996
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2997
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2998
  "\<lbrakk>Suc n \<le> ba - aa;  length lm2 = Suc n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  2999
    length lm3 = ba - Suc (aa + n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3000
  \<Longrightarrow> (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3001
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3002
  assume h: "Suc n \<le> ba - aa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3003
  and g: "length lm2 = Suc n" "length lm3 = ba - Suc (aa + n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3004
  from h and g have k: "ba - aa = Suc (length lm3 + n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3005
    by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3006
  from  k show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3007
    "(last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3008
    apply(simp, insert g)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3009
    apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3010
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3011
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3012
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3013
lemma [simp]: "length lm1 = aa \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3014
      (lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3015
apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3016
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3017
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3018
lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3019
                    (ba < Suc (aa + (ba - Suc (aa + n) + n))) = False"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3020
apply arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3021
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3022
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3023
lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3024
       length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3025
     \<Longrightarrow> (lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3026
using nth_append[of "lm1 @ 0\<Colon>'a\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3027
                     "(0\<Colon>'a) # lm4" "ba + n"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3028
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3029
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3030
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3031
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3032
 "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3033
                 length lm3 = ba - Suc (aa + n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3034
  \<Longrightarrow> (lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3035
  [ba + n := last lm2, aa + n := 0] = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3036
  lm1 @ 0 # 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3037
using list_update_append[of "lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2" "0 # lm4" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3038
                            "ba + n" "last lm2"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3039
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3040
apply(simp add: list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3041
apply(simp only: exponent_cons_iff exp_suc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3042
apply(case_tac lm2, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3043
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3044
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3045
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3046
lemma mv_boxes_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3047
  "\<lbrakk>n \<le> ba - aa; ba > aa; length lm1 = aa; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3048
    length (lm2::nat list) = n; length lm3 = ba - aa - n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3049
     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ lm4)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3050
       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3051
apply(induct n arbitrary: lm2 lm3 lm4, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3052
apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3053
              simp add: mv_boxes.simps del: exp_suc_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3054
apply(rule_tac as = "3 *n" and bm = "lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3055
               butlast lm2 @ 0 # lm4" in abc_append_exc2, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3056
apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3057
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3058
  fix n lm2 lm3 lm4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3059
  assume ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3060
    "\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = n; length lm3 = ba - (aa + n)\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3061
    \<exists>stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ lm4) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3062
       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3063
  and h: "Suc n \<le> ba - aa" "aa < ba" "length (lm1::nat list) = aa" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3064
         "length (lm2::nat list) = Suc n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3065
         "length (lm3::nat list) = ba - Suc (aa + n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3066
  from h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3067
    "\<exists>astp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ 0 # lm4) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3068
                       (mv_boxes aa ba n) astp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3069
        (3 * n, lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3070
    apply(insert ind[of "butlast lm2" "last lm2 # lm3" "0 # lm4"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3071
          simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3072
    apply(subgoal_tac "lm1 @ butlast lm2 @ last lm2 # lm3 @ 0\<^bsup>n\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3073
              0 # lm4 = lm1 @ lm2 @ lm3 @ 0\<^bsup>n\<^esup> @ 0 # lm4", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3074
    apply(case_tac "lm2 = []", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3075
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3076
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3077
  fix n lm2 lm3 lm4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3078
  assume h: "Suc n \<le> ba - aa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3079
            "aa < ba" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3080
            "length (lm1::nat list) = aa" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3081
            "length (lm2::nat list) = Suc n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3082
            "length (lm3::nat list) = ba - Suc (aa + n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3083
  thus " \<exists>bstp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3084
                       butlast lm2 @ 0 # lm4) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3085
                         (recursive.empty (aa + n) (ba + n)) bstp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3086
               = (3, lm1 @ 0 # 0\<^bsup>n\<^esup> @ lm3 @ lm2 @ lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3087
    apply(insert empty_ex[of "aa + n" "ba + n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3088
       "lm1 @ 0\<^bsup>n\<^esup> @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3089
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3090
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3091
(*    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3092
lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3093
                ba < aa; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3094
               length lm2 = aa - Suc (ba + n)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3095
      \<Longrightarrow> ((0::nat) # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (aa - ba)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3096
         = last lm3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3097
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3098
  assume h: "Suc n \<le> aa - ba"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3099
    and g: " ba < aa" "length lm2 = aa - Suc (ba + n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3100
  from h and g have k: "aa - ba = Suc (length lm2 + n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3101
    by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3102
  thus "((0::nat) # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (aa - ba) = last lm3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3103
    apply(simp,  simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3104
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3105
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3106
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3107
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3108
lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3109
        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3110
   \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (aa + n) = last lm3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3111
using nth_append[of "lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup>" "last lm3 # lm4" "aa + n"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3112
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3113
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3114
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3115
lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3116
        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3117
     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4) ! (ba + n) = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3118
apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3119
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3120
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3121
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3122
lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3123
        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3124
     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4)[ba + n := last lm3, aa + n := 0]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3125
      = lm1 @ lm3 @ lm2 @ 0 # 0\<^bsup>n\<^esup> @ lm4"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3126
using list_update_append[of "lm1 @ butlast lm3" "(0\<Colon>'a) # lm2 @ 0\<Colon>'a\<^bsup>n\<^esup> @ last lm3 # lm4"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3127
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3128
using list_update_append[of "lm1 @ butlast lm3 @ last lm3 # lm2 @ 0\<Colon>'a\<^bsup>n\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3129
                            "last lm3 # lm4" "aa + n" "0"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3130
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3131
apply(simp only: exp_ind_def[THEN sym] exp_suc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3132
apply(case_tac lm3, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3133
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3134
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3135
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3136
lemma mv_boxes_ex2:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3137
  "\<lbrakk>n \<le> aa - ba; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3138
    ba < aa; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3139
    length (lm1::nat list) = ba;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3140
    length (lm2::nat list) = aa - ba - n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3141
    length (lm3::nat list) = n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3142
     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ lm2 @ lm3 @ lm4) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3143
                (mv_boxes aa ba n) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3144
                    (3 * n, lm1 @ lm3 @ lm2 @ 0\<^bsup>n\<^esup> @ lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3145
apply(induct n arbitrary: lm2 lm3 lm4, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3146
apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3147
                   simp add: mv_boxes.simps del: exp_suc_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3148
apply(rule_tac as = "3 *n" and bm = "lm1 @ butlast lm3 @ 0 # lm2 @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3149
                  0\<^bsup>n\<^esup> @ last lm3 # lm4" in abc_append_exc2, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3150
apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3151
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3152
  fix n lm2 lm3 lm4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3153
  assume ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3154
"\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = aa - (ba + n); length lm3 = n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3155
  \<exists>stp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ lm2 @ lm3 @ lm4) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3156
                 (mv_boxes aa ba n) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3157
                            (3 * n, lm1 @ lm3 @ lm2 @ 0\<^bsup>n\<^esup> @ lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3158
  and h: "Suc n \<le> aa - ba" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3159
         "ba < aa"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3160
         "length (lm1::nat list) = ba" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3161
         "length (lm2::nat list) = aa - Suc (ba + n)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3162
         "length (lm3::nat list) = Suc n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3163
  from h show
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3164
    "\<exists>astp. abc_steps_l (0, lm1 @ 0\<^bsup>n\<^esup> @ 0 # lm2 @ lm3 @ lm4)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3165
        (mv_boxes aa ba n) astp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3166
          (3 * n, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3167
    apply(insert ind[of "0 # lm2" "butlast lm3" "last lm3 # lm4"],
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3168
          simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3169
    apply(subgoal_tac
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3170
      "lm1 @ 0\<^bsup>n\<^esup> @ 0 # lm2 @ butlast lm3 @ last lm3 # lm4 =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3171
           lm1 @ 0\<^bsup>n\<^esup> @ 0 # lm2 @ lm3 @ lm4", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3172
    apply(case_tac "lm3 = []", simp, simp)
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
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3175
  fix n lm2 lm3 lm4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3176
  assume h:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3177
    "Suc n \<le> aa - ba" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3178
    "ba < aa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3179
    "length lm1 = ba"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3180
    "length (lm2::nat list) = aa - Suc (ba + n)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3181
    "length (lm3::nat list) = Suc n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3182
  thus
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3183
    "\<exists>bstp. abc_steps_l (0, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<^bsup>n\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3184
                               last lm3 # lm4) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3185
           (recursive.empty (aa + n) (ba + n)) bstp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3186
                 (3, lm1 @ lm3 @ lm2 @ 0 # 0\<^bsup>n\<^esup> @ lm4)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3187
    apply(insert empty_ex[of "aa + n" "ba + n" "lm1 @ butlast lm3 @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3188
                          0 # lm2 @ 0\<^bsup>n\<^esup> @ last lm3 # lm4"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3189
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3190
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3191
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3192
lemma cn_merge_gs_len: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3193
  "length (cn_merge_gs (map rec_ci gs) pstr) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3194
      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3195
apply(induct gs arbitrary: pstr, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3196
apply(case_tac "rec_ci a", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3197
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3198
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3199
lemma [simp]: "n < pstr \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3200
     Suc (pstr + length ys - n) = Suc (pstr + length ys) - n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3201
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3202
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3203
lemma save_paras':  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3204
  "\<lbrakk>length lm = n; pstr > n; a_md > pstr + length ys + n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3205
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3206
               0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3207
                 (mv_boxes 0 (pstr + Suc (length ys)) n) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3208
        = (3 * n, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3209
thm mv_boxes_ex
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3210
apply(insert mv_boxes_ex[of n "pstr + Suc (length ys)" 0 "[]" "lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3211
         "0\<^bsup>pstr - n\<^esup> @ ys @ [0]" "0\<^bsup>a_md - pstr - length ys - n - Suc 0\<^esup> @ suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3212
apply(erule_tac exE, rule_tac x = stp in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3213
                            simp add: exponent_add_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3214
apply(simp only: exponent_cons_iff, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3215
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3216
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3217
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3218
 "(max ba (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3219
 = (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3220
apply(rule min_max.sup_absorb2, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3221
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3222
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3223
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3224
  "((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3225
                  (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3226
apply(induct gs)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3227
apply(simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3228
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3229
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3230
lemma ci_cn_md_def:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3231
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3232
  rec_ci f = (a, aa, ba)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3233
    \<Longrightarrow> a_md = max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3234
  rec_ci) ` set gs))) + Suc (length gs) + n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3235
apply(simp add: rec_ci.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3236
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3237
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3238
lemma save_paras_prog_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3239
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3240
    rec_ci f = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3241
    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3242
                                    (map rec_ci (f # gs))))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3243
    \<Longrightarrow> \<exists>ap bp cp. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3244
      aprog = ap [+] bp [+] cp \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3245
      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3246
              3 * length gs \<and> bp = mv_boxes 0 (pstr + Suc (length gs)) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3247
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3248
apply(rule_tac x = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3249
  "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3250
      (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3251
      simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3252
apply(rule_tac x = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3253
  "mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3254
   0 (length gs) [+] a [+]recursive.empty aa (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3255
   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3256
   empty_boxes (length gs) [+] recursive.empty (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3257
  (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3258
   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3259
   ` set gs))) + length gs)) 0 n" in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3260
apply(simp add: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3261
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3262
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3263
lemma save_paras: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3264
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3265
    rs_pos = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3266
    \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3267
    length ys = length gs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3268
    length lm = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3269
    rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3270
    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3271
                                          (map rec_ci (f # gs))))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3272
  \<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3273
          3 * length gs, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3274
                 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3275
           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3276
                      3 * length gs + 3 * n, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3277
             0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup>  @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3278
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3279
  assume h:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3280
    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3281
    "rs_pos = n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3282
    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3283
    "length ys = length gs"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3284
    "length lm = n"    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3285
    "rec_ci f = (a, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3286
    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3287
                                        (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3288
  from h and g have k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3289
    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3290
    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3291
                3 *length gs \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3292
    apply(drule_tac save_paras_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3293
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3294
  from h have k2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3295
    "\<exists> stp. abc_steps_l (0, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3296
                         0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3297
         (mv_boxes 0 (pstr + Suc (length ys)) n) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3298
        (3 * n, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup>  @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3299
    apply(rule_tac save_paras', simp, simp_all add: g)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3300
    apply(drule_tac a = a and aa = aa and ba = ba in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3301
                                        ci_cn_md_def, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3302
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3303
  from k1 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3304
    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3305
         3 * length gs, lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3306
                 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3307
             ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3308
               3 * length gs + 3 * n, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3309
                0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3310
  proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3311
    fix ap bp apa cp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3312
    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3313
            (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3314
            \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3315
    from this and k2 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3316
      apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3317
      apply(rule_tac abc_append_exc1, simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3318
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3319
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3320
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3321
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3322
lemma ci_cn_para_eq:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3323
  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3324
apply(simp add: rec_ci.simps, case_tac "rec_ci f", simp)
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 calc_gs_prog_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3328
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3329
    rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3330
    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3331
                         (map rec_ci (f # gs)))) = pstr\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3332
   \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3333
                 ap = cn_merge_gs (map rec_ci gs) pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3334
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3335
apply(rule_tac x = "mv_boxes 0 (Suc (max (Suc n)  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3336
   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3337
   mv_boxes (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3338
  (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3339
   a [+] recursive.empty aa (max (Suc n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3340
    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3341
   empty_boxes (length gs) [+] recursive.empty (max (Suc n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3342
    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3343
    mv_boxes (Suc (max (Suc n) (Max 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3344
    (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3345
   in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3346
apply(auto simp: abc_append_commute)
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 cn_calc_gs: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3350
  assumes ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3351
  "\<And>x aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3352
  \<lbrakk>x \<in> set gs; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3353
   rec_ci x = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3354
   rec_calc_rel x lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3355
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3356
     (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3357
  and h:  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3358
          "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3359
          "length ys = length gs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3360
          "length lm = n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3361
          "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3362
          "Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3363
                               (map rec_ci (f # gs)))) = pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3364
  shows  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3365
  "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3366
  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3367
   lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -pstr - length ys\<^esup> @ suf_lm) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3368
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3369
  from h have k1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3370
    "\<exists> ap bp. aprog = ap [+] bp \<and> ap = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3371
                        cn_merge_gs (map rec_ci gs) pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3372
    by(erule_tac calc_gs_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3373
  from h have j1: "rs_pos = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3374
    by(simp add: ci_cn_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3375
  from h have j2: "a_md \<ge> pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3376
    by(drule_tac a = a and aa = aa and ba = ba in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3377
                                ci_cn_md_def, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3378
  from h have j3: "pstr > n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3379
    by(auto)    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3380
  from j1 and j2 and j3 and h have k2:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3381
    "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3382
                         (cn_merge_gs (map rec_ci gs) pstr) stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3383
    = ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3384
                  lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3385
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3386
    apply(rule_tac cn_merge_gs_ex, rule_tac ind, simp, simp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3387
    apply(drule_tac a = a and aa = aa and ba = ba in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3388
                                 ci_cn_md_def, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3389
    apply(rule min_max.le_supI2, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3390
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3391
  from k1 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3392
  proof(erule_tac exE, erule_tac exE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3393
    fix ap bp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3394
    from k2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3395
      "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3396
           (cn_merge_gs (map rec_ci gs) pstr [+] bp) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3397
      (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3398
         3 * length gs, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3399
         lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - (pstr + length ys)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3400
      apply(insert abc_append_exc1[of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3401
        "lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3402
        "(cn_merge_gs (map rec_ci gs) pstr)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3403
        "length (cn_merge_gs (map rec_ci gs) pstr)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3404
        "lm @ 0\<^bsup>pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - pstr - length ys\<^esup> @ suf_lm" 0 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3405
        "[]" bp], simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3406
      done      
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3407
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3408
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3409
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3410
lemma reset_new_paras': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3411
  "\<lbrakk>length lm = n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3412
    pstr > 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3413
    a_md \<ge> pstr + length ys + n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3414
     pstr > length ys\<rbrakk> \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3415
   \<exists>stp. abc_steps_l (0, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @  0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3416
          suf_lm) (mv_boxes pstr 0 (length ys)) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3417
  (3 * length ys, ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3418
thm mv_boxes_ex2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3419
apply(insert mv_boxes_ex2[of "length ys" "pstr" 0 "[]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3420
     "0\<^bsup>pstr - length ys\<^esup>" "ys" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3421
     "0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3422
     simp add: exponent_add_iff)
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]:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3426
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3427
  rec_calc_rel f ys rs; rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3428
  pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3429
               (map rec_ci (f # gs))))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3430
  \<Longrightarrow> length ys < pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3431
apply(subgoal_tac "length ys = aa", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3432
apply(subgoal_tac "aa < ba \<and> ba \<le> pstr", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3433
      rule basic_trans_rules(22), auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3434
apply(rule min_max.le_supI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3435
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3436
apply(erule_tac para_pattern, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3437
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3438
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3439
lemma reset_new_paras_prog_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3440
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3441
   rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3442
   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3443
  (map rec_ci (f # gs)))) = pstr\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3444
  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3445
  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3446
           3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length gs)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3447
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3448
apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3449
          (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3450
          mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3451
           (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3452
       simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3453
apply(rule_tac x = "a [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3454
     recursive.empty aa (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3455
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3456
     empty_boxes (length gs) [+] recursive.empty 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3457
     (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3458
      [+] mv_boxes (Suc (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3459
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3460
       auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3461
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3462
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3463
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3464
lemma reset_new_paras:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3465
       "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3466
        rs_pos = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3467
        \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3468
        length ys = length gs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3469
        length lm = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3470
        length ys = aa;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3471
        rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3472
        pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3473
                                    (map rec_ci (f # gs))))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3474
\<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3475
                                               3 * length gs + 3 * n,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3476
        0\<^bsup>pstr \<^esup>@ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3477
  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3478
           ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3479
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3480
  assume h:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3481
    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3482
    "rs_pos = n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3483
    "length ys = aa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3484
    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3485
    "length ys = length gs"  "length lm = n"    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3486
    "rec_ci f = (a, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3487
    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3488
                                         (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3489
  thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3490
  from h and g have k1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3491
    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3492
    (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3493
          3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3494
    by(drule_tac reset_new_paras_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3495
  from h have k2:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3496
    "\<exists> stp. abc_steps_l (0, 0\<^bsup>pstr \<^esup>@ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3497
              suf_lm) (mv_boxes pstr 0 (length ys)) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3498
    (3 * (length ys), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3499
     ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3500
    apply(rule_tac reset_new_paras', simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3501
    apply(simp add: g)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3502
    apply(drule_tac a = a and aa = aa and ba = ba in ci_cn_md_def,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3503
      simp, simp add: g, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3504
    apply(subgoal_tac "length gs = aa \<and> aa < ba \<and> ba \<le> pstr", arith,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3505
          simp add: para_pattern)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3506
    apply(insert g, auto intro: min_max.le_supI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3507
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3508
  from k1 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3509
    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3510
    * length gs + 3 * n, 0\<^bsup>pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3511
     suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3512
    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3513
      3 * n, ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3514
  proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3515
    fix ap bp apa cp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3516
    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3517
      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3518
                  3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3519
    from this and k2 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3520
      apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3521
      apply(drule_tac as = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3522
        "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3523
        3 * n" and ap = ap and cp = cp in abc_append_exc1, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3524
      apply(rule_tac x = stp in exI, simp add: h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3525
      using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3526
      apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3527
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3528
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3529
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3530
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3531
thm rec_ci.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3532
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3533
lemma calc_f_prog_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3534
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3535
    rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3536
    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3537
                   (map rec_ci (f # gs)))) = pstr\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3538
   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3539
  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3540
                                6 *length gs + 3 * n \<and> bp = a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3541
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3542
apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3543
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3544
     mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3545
            (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3546
     mv_boxes (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3547
      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs)" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3548
     simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3549
apply(rule_tac x = "recursive.empty aa (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3550
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3551
     empty_boxes (length gs) [+] recursive.empty (max (Suc n) (
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3552
     Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3553
     mv_boxes (Suc (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3554
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3555
  auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3556
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3557
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3558
lemma calc_cn_f:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3559
  assumes ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3560
  "\<And>x aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3561
  \<lbrakk>x \<in> set (f # gs);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3562
  rec_ci x = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3563
  rec_calc_rel x lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3564
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3565
  (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3566
  and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3567
  "rec_calc_rel (Cn n f gs) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3568
  "length ys = length gs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3569
  "rec_calc_rel f ys rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3570
  "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3571
  "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3572
  and p: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3573
                                (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3574
  shows "\<exists>stp. abc_steps_l   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3575
  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3576
  ys @ 0\<^bsup>pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3577
  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3578
                3 * n + length a,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3579
  ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3580
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3581
  from h have k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3582
    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3583
    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3584
    6 *length gs + 3 * n \<and> bp = a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3585
    by(drule_tac calc_f_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3586
  from h and k1 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3587
  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3588
    fix ap bp apa cp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3589
    assume
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3590
      "aprog = ap [+] bp [+] cp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3591
      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3592
      6 * length gs + 3 * n \<and> bp = a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3593
    from h and this show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3594
      apply(simp, rule_tac abc_append_exc1, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3595
      apply(insert ind[of f "a" aa ba ys rs 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3596
        "0\<^bsup>pstr - ba + length gs \<^esup> @ 0 # lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3597
        0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3598
      apply(subgoal_tac "ba > aa \<and> aa = length gs\<and> pstr \<ge> ba", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3599
      apply(simp add: exponent_add_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3600
      apply(case_tac pstr, simp add: p)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3601
      apply(simp only: exp_suc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3602
      apply(rule conjI, rule ci_ad_ge_paras, simp, rule conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3603
      apply(subgoal_tac "length ys = aa", simp,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3604
        rule para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3605
      apply(insert p, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3606
      apply(auto intro: min_max.le_supI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3607
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3608
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3609
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3610
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3611
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3612
  "\<lbrakk>pstr + length ys + n \<le> a_md; ys \<noteq> []\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3613
                          pstr < a_md + length suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3614
apply(case_tac "length ys", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3615
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3616
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3617
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3618
lemma [simp]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3619
  "pstr > length ys 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3620
  \<Longrightarrow> (ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3621
  0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) ! pstr = (0::nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3622
apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3623
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3624
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3625
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3626
lemma [simp]: "\<lbrakk>length ys < pstr; pstr - length ys = Suc x\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3627
  \<Longrightarrow> pstr - Suc (length ys) = x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3628
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3629
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3630
lemma [simp]: "pstr > length ys \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3631
      (ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3632
                                         [pstr := rs, length ys := 0] =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3633
       ys @ 0\<^bsup>pstr - length ys\<^esup> @ (rs::nat) # 0\<^bsup>length ys\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3634
apply(auto simp: list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3635
apply(case_tac "pstr - length ys",simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3636
using list_update_length[of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3637
  "0\<^bsup>pstr - Suc (length ys)\<^esup>" "0" "0\<^bsup>length ys\<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3638
  0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm" rs]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3639
apply(simp only: exponent_cons_iff exponent_add_iff, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3640
apply(subgoal_tac "pstr - Suc (length ys) = nat", simp, 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 save_rs': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3644
  "\<lbrakk>pstr > length ys\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3645
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3646
  0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3647
  (recursive.empty (length ys) pstr) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3648
  (3, ys @ 0\<^bsup>pstr - (length ys)\<^esup> @ rs # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3649
  0\<^bsup>length ys \<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3650
apply(insert empty_ex[of "length ys" pstr 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3651
  "ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @ 0\<^bsup>a_md - Suc(pstr + length ys + n)\<^esup> @ suf_lm"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3652
  simp)
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3656
lemma save_rs_prog_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3657
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3658
  rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3659
  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3660
                        (map rec_ci (f # gs)))) = pstr\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3661
  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3662
  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3663
              6 *length gs + 3 * n + length a
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3664
  \<and> bp = empty aa pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3665
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3666
apply(rule_tac x =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3667
  "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3668
   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3669
   [+] mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3670
   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3671
   mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3672
    0 (length gs) [+] a" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3673
  in exI, simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3674
apply(rule_tac x = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3675
  "empty_boxes (length gs) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3676
   recursive.empty (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3677
    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3678
   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3679
    + length gs)) 0 n" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3680
  auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3681
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3682
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3683
lemma save_rs:  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3684
  assumes h: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3685
  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3686
  "rec_calc_rel (Cn n f gs) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3687
  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3688
  "length ys = length gs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3689
  "rec_calc_rel f ys rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3690
  "rec_ci f = (a, aa, ba)"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3691
  "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3692
  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3693
                                            (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3694
  shows "\<exists>stp. abc_steps_l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3695
           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3696
          + 3 * n + length a, ys @ rs # 0\<^bsup>pstr\<^esup> @ lm @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3697
             0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3698
  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3699
  + 3 * n + length a + 3,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3700
  ys @ 0\<^bsup>pstr - length ys \<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3701
                               0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3702
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3703
  thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3704
  from h and pdef have k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3705
    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3706
    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3707
    6 *length gs + 3 * n + length a \<and> bp = empty (length ys) pstr "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3708
    apply(subgoal_tac "length ys = aa")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3709
    apply(drule_tac a = a and aa = aa and ba = ba in save_rs_prog_ex, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3710
      simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3711
    by(rule_tac para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3712
  from k1 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3713
    "\<exists>stp. abc_steps_l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3714
    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3715
    + length a, ys @ rs # 0\<^bsup>pstr \<^esup>@ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3716
    @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3717
    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3718
    + length a + 3, ys @ 0\<^bsup>pstr - length ys\<^esup> @ rs # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3719
    0\<^bsup>length ys\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3720
  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3721
    fix ap bp apa cp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3722
    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3723
      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3724
      3 * n + length a \<and> bp = recursive.empty (length ys) pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3725
    thus"?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3726
      apply(simp, rule_tac abc_append_exc1, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3727
      apply(rule_tac save_rs', insert h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3728
      apply(subgoal_tac "length gs = aa \<and> pstr \<ge> ba \<and> ba > aa",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3729
            arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3730
      apply(simp add: para_pattern, insert pdef, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3731
      apply(rule_tac min_max.le_supI2, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3732
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3733
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3734
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3735
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3736
lemma [simp]: "length (empty_boxes n) = 2*n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3737
apply(induct n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3738
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3739
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3740
lemma empty_step_ex: "length lm = n \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3741
      \<exists>stp. abc_steps_l (0, lm @ Suc x # suf_lm) [Dec n 2, Goto 0] stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3742
  = (0, lm @ x # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3743
apply(rule_tac x = "Suc (Suc 0)" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3744
  simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3745
         abc_lm_v.simps abc_lm_s.simps nth_append list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3746
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3747
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3748
lemma empty_box_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3749
  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3750
  \<exists> stp. abc_steps_l (0, lm @ x # suf_lm) [Dec n 2, Goto 0] stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3751
  (Suc (Suc 0), lm @ 0 # suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3752
apply(induct x)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3753
apply(rule_tac x = "Suc 0" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3754
  simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3755
            abc_lm_v.simps nth_append abc_lm_s.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3756
apply(drule_tac x = x and suf_lm = suf_lm in empty_step_ex, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3757
      erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3758
apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3759
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3760
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3761
lemma [simp]: "drop n lm = a # list \<Longrightarrow> list = drop (Suc n) lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3762
apply(induct n arbitrary: lm a list, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3763
apply(case_tac "lm", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3764
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3765
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3766
lemma empty_boxes_ex: "\<lbrakk>length lm \<ge> n\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3767
     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm) (empty_boxes n) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3768
                                          (2*n, 0\<^bsup>n\<^esup> @ drop n lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3769
apply(induct n, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3770
apply(rule_tac abc_append_exc2, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3771
apply(case_tac "drop n lm", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3772
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3773
  fix n stp a list
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3774
  assume h: "Suc n \<le> length lm"  "drop n lm = a # list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3775
  thus "\<exists>bstp. abc_steps_l (0, 0\<^bsup>n\<^esup> @ a # list) [Dec n 2, Goto 0] bstp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3776
                       (Suc (Suc 0), 0 # 0\<^bsup>n\<^esup> @ drop (Suc n) lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3777
    apply(insert empty_box_ex[of "0\<^bsup>n\<^esup>" n a list], simp, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3778
    apply(rule_tac x = stp in exI, simp, simp only: exponent_cons_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3779
    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
  3780
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3781
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3782
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3783
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3784
lemma empty_paras_prog_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3785
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3786
  rec_ci f = (a, aa, ba); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3787
  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3788
                    (map rec_ci (f # gs)))) = pstr\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3789
  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3790
  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3791
  6 *length gs + 3 * n + length a + 3 \<and> bp = empty_boxes (length gs)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3792
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3793
apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3794
    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3795
    mv_boxes 0 (Suc (max (Suc n) (Max 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3796
     (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3797
    [+] mv_boxes (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3798
    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3799
     a [+] recursive.empty aa (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3800
   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3801
    in exI, simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3802
apply(rule_tac x = " recursive.empty (max (Suc n) (Max (insert ba
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3803
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3804
     mv_boxes (Suc (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3805
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3806
  auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3807
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3808
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3809
lemma empty_paras: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3810
 assumes h: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3811
  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3812
  "rec_calc_rel (Cn n f gs) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3813
  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3814
  "length ys = length gs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3815
  "rec_calc_rel f ys rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3816
  "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3817
  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3818
                                             (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3819
  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3820
                              6 * length gs + 3 * n + length a + 3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3821
  shows "\<exists>stp. abc_steps_l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3822
           (ss, ys @ 0\<^bsup>pstr - length ys\<^esup> @ rs # 0\<^bsup>length ys\<^esup> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3823
               @ lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3824
   (ss + 2 * length gs, 0\<^bsup>pstr\<^esup> @ rs # 0\<^bsup>length ys \<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3825
                                0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3826
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3827
  from h and pdef and starts have k1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3828
    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3829
    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3830
                               6 *length gs + 3 * n + length a + 3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3831
    \<and> bp = empty_boxes (length ys)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3832
    by(drule_tac empty_paras_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3833
  from h have j1: "aa < ba"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3834
    by(simp add: ci_ad_ge_paras)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3835
  from h have j2: "length gs = aa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3836
    by(drule_tac f = f in para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3837
  from h and pdef have j3: "ba \<le> pstr"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3838
    apply simp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3839
    apply(rule_tac min_max.le_supI2, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3840
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3841
  from k1 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3842
  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3843
    fix ap bp apa cp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3844
    assume "aprog = ap [+] bp [+] cp \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3845
      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3846
      6 * length gs + 3 * n + length a + 3 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3847
      bp = empty_boxes (length ys)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3848
    thus"?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3849
      apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3850
      apply(insert empty_boxes_ex[of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3851
        "length gs" "ys @ 0\<^bsup>pstr - (length gs)\<^esup> @ rs #
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3852
        0\<^bsup>length gs\<^esup> @ lm @ 0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3853
        simp add: h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3854
      apply(erule_tac exE, rule_tac x = stp in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3855
        simp add: exponent_def replicate.simps[THEN sym]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3856
        replicate_add[THEN sym] del: replicate.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3857
      apply(subgoal_tac "pstr >(length gs)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3858
      apply(subgoal_tac "ba > aa \<and> length gs = aa \<and> pstr \<ge> ba", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3859
      apply(simp add: j1 j2 j3)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3860
      done     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3861
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3862
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3863
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3864
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3865
lemma [simp]: " n < pstr \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3866
  (0\<^bsup>pstr\<^esup>)[n := rs] @ [0::nat] = 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3867
apply(insert list_update_length[of "0\<^bsup>n\<^esup>" 0 "0\<^bsup>pstr - Suc n\<^esup>" rs])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3868
apply(insert exponent_cons_iff[of "0::nat" "pstr - Suc n" "[]"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3869
apply(insert exponent_add_iff[of "0::nat" n "pstr - n" "[]"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3870
apply(case_tac "pstr - n", simp, simp only: exp_suc, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3871
apply(subgoal_tac "pstr - Suc n = nat", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3872
by arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3873
*)
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 restore_rs_prog_ex:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3876
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3877
  rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3878
  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3879
  (map rec_ci (f # gs)))) = pstr;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3880
  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3881
  8 * length gs + 3 * n + length a + 3\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3882
  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3883
                                           bp = empty pstr n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3884
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3885
apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3886
      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3887
      mv_boxes 0 (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3888
        \<circ> rec_ci) ` set gs))) + length gs)) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3889
     mv_boxes (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3890
      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3891
     a [+] recursive.empty aa (max (Suc n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3892
       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3893
     empty_boxes (length gs)" in exI, simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3894
apply(rule_tac x = "mv_boxes (Suc (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3895
       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3896
        + length gs)) 0 n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3897
  in exI, auto simp: abc_append_commute)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3898
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3899
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3900
lemma exp_add: "a\<^bsup>b+c\<^esup> = a\<^bsup>b\<^esup> @ a\<^bsup>c\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3901
apply(simp add: exponent_def replicate_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3902
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3903
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3904
lemma [simp]: "n < pstr \<Longrightarrow> (0\<^bsup>pstr\<^esup>)[n := rs] @ [0::nat] = 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3905
using list_update_length[of "0\<^bsup>n\<^esup>" "0::nat" "0\<^bsup>pstr - Suc n\<^esup>" rs]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3906
apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym] exp_suc[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3907
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3908
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3909
lemma restore_rs:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3910
  assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3911
  "rec_calc_rel (Cn n f gs) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3912
  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3913
  "length ys = length gs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3914
  "rec_calc_rel f ys rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3915
  "rec_ci f = (a, aa, ba)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3916
  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3917
                                        (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3918
  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3919
                              8 * length gs + 3 * n + length a + 3" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3920
  shows "\<exists>stp. abc_steps_l
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3921
           (ss, 0\<^bsup>pstr\<^esup> @ rs # 0\<^bsup>length ys \<^esup> @ lm @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3922
                    0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3923
  (ss + 3, 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n\<^esup> @ 0\<^bsup>length ys \<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3924
                                   0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3925
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3926
 from h and pdef and starts have k1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3927
   "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3928
                                            bp = empty pstr n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3929
   by(drule_tac restore_rs_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3930
 from k1 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3931
 proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3932
   fix ap bp apa cp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3933
   assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3934
                                 bp = recursive.empty pstr n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3935
   thus"?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3936
     apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3937
     apply(insert empty_ex[of pstr n "0\<^bsup>pstr\<^esup> @ rs # 0\<^bsup>length gs\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3938
                     lm @ 0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3939
     apply(subgoal_tac "pstr > n", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3940
     apply(erule_tac exE, rule_tac x = stp in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3941
                         simp add: nth_append list_update_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3942
     apply(simp add: pdef)
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
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3945
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3946
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3947
lemma [simp]:"xs \<noteq> [] \<Longrightarrow> length xs \<ge> Suc 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3948
by(case_tac xs, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3949
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3950
lemma  [simp]: "n < max (Suc n) (max ba (Max (((\<lambda>(aprog, p, n). n) o 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3951
                                                  rec_ci) ` set gs)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3952
by(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3953
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3954
lemma restore_paras_prog_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3955
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3956
  rec_ci f = (a, aa, ba);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3957
  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3958
                          (map rec_ci (f # gs)))) = pstr;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3959
  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3960
                         8 * length gs + 3 * n + length a + 6\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3961
  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3962
                      bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3963
apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3964
apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3965
      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3966
      [+] mv_boxes 0 (Suc (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3967
       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3968
     + length gs)) n [+] mv_boxes (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3969
    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3970
     a [+] recursive.empty aa (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3971
      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3972
     empty_boxes (length gs) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3973
     recursive.empty (max (Suc n) (Max (insert ba 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3974
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n" in exI, simp add: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3975
apply(rule_tac x = "[]" in exI, auto simp: abc_append_commute)
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 restore_paras: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3979
  assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3980
  "rec_calc_rel (Cn n f gs) lm rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3981
  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3982
  "length ys = length gs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3983
  "rec_calc_rel f ys rs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3984
  "rec_ci f = (a, aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3985
  and pdef: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3986
  "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3987
                         (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3988
  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3989
                              8 * length gs + 3 * n + length a + 6" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3990
  shows "\<exists>stp. abc_steps_l (ss, 0\<^bsup>n\<^esup> @ rs # 0\<^bsup>pstr - n+ length ys\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3991
                         lm @ 0\<^bsup>a_md - Suc (pstr + length ys + n)\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3992
  aprog stp = (ss + 3 * n, lm @ rs # 0\<^bsup>a_md - Suc n\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3993
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3994
  thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3995
  from h and pdef and starts have k1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3996
    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3997
                     bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3998
    by(drule_tac restore_paras_prog_ex, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  3999
  from k1 show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4000
  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4001
    fix ap bp apa cp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4002
    assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4003
                              bp = mv_boxes (pstr + Suc (length gs)) 0 n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4004
    thus"?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4005
      apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4006
      apply(insert mv_boxes_ex2[of n "pstr + Suc (length gs)" 0 "[]" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4007
        "rs # 0\<^bsup>pstr - n + length gs\<^esup>" "lm" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4008
        "0\<^bsup>a_md - Suc (pstr + length gs + n)\<^esup> @ suf_lm"], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4009
      apply(subgoal_tac "pstr > n \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4010
        a_md > pstr + length gs + n \<and> length lm = n" , simp add: exponent_add_iff h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4011
      using h pdef
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4012
      apply(simp)     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4013
      apply(frule_tac a = a and 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4014
        aa = aa and ba = ba in ci_cn_md_def, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4015
      apply(subgoal_tac "length lm = rs_pos",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4016
        simp add: ci_cn_para_eq, erule_tac para_pattern, 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
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4019
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4020
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4021
lemma ci_cn_length:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4022
  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4023
  rec_calc_rel (Cn n f gs) lm rs;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4024
  rec_ci f = (a, aa, ba)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4025
  \<Longrightarrow> length aprog = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4026
                             8 * length gs + 6 * n + length a + 6"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4027
apply(simp add: rec_ci.simps, auto simp: cn_merge_gs_len)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4028
done
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4031
lemma  cn_case: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4032
  assumes ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4033
  "\<And>x aprog a_md rs_pos rs suf_lm lm.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4034
  \<lbrakk>x \<in> set (f # gs);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4035
  rec_ci x = (aprog, rs_pos, a_md);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4036
  rec_calc_rel x lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4037
  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4038
               (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4039
  and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4040
         "rec_calc_rel (Cn n f gs) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4041
         
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4042
  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4043
  = (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4044
apply(insert h, case_tac "rec_ci f",  rule_tac calc_cn_reverse, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4045
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4046
  fix a b c ys
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4047
  let ?pstr = "Max (set (Suc n # c # (map (\<lambda>(aprog, p, n). n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4048
                                         (map rec_ci (f # gs)))))"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4049
  let ?gs_len = "listsum (map (\<lambda> (ap, pos, n). length ap) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4050
                                                (map rec_ci (gs)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4051
  assume g: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4052
    "rec_calc_rel (Cn n f gs) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4053
    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4054
    "length ys = length gs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4055
    "rec_calc_rel f ys rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4056
    "n = length lm"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4057
    "rec_ci f = (a, b, c)"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4058
  hence k1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4059
    "\<exists> stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4060
    (?gs_len + 3 * length gs, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4061
                               0\<^bsup>a_md - ?pstr - length ys\<^esup> @ suf_lm)"	
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4062
    apply(rule_tac a = a and aa = b and ba = c in cn_calc_gs)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4063
    apply(rule_tac ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4064
    done  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4065
  thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4066
  from g have k2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4067
    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs,  lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4068
        0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md - ?pstr - length ys\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4069
    (?gs_len + 3 * length gs + 3 * n, 0\<^bsup>?pstr\<^esup> @ ys @ 0 # lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4070
                              0\<^bsup>a_md - Suc (?pstr + length ys + n )\<^esup>  @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4071
    thm save_paras
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4072
    apply(erule_tac ba = c in save_paras, auto intro: ci_cn_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4073
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4074
  from g have k3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4075
    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs + 3 * n,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4076
    0\<^bsup>?pstr\<^esup> @ ys @ 0 # lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup>  @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4077
    (?gs_len + 6 * length gs + 3 * n,  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4078
           ys @ 0\<^bsup>?pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup>  @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4079
    apply(erule_tac ba = c in reset_new_paras, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4080
          auto intro: ci_cn_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4081
    using para_pattern[of f a b c ys rs]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4082
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4083
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4084
  from g have k4: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4085
    "\<exists>stp. abc_steps_l  (?gs_len + 6 * length gs + 3 * n,  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4086
    ys @ 0\<^bsup>?pstr\<^esup> @ 0 # lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup>  @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4087
    (?gs_len + 6 * length gs + 3 * n + length a, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4088
   ys @ rs # 0\<^bsup>?pstr \<^esup> @ lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup>  @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4089
    apply(rule_tac ba = c in calc_cn_f, rule_tac ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4090
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4091
thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4092
  from g h have k5:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4093
    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + length a,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4094
    ys @ rs # 0\<^bsup>?pstr \<^esup>@ lm @ 0\<^bsup>a_md - Suc (?pstr + length ys + n)\<^esup>  @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4095
    aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4096
    (?gs_len + 6 * length gs + 3 * n + length a + 3,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4097
    ys @ 0\<^bsup>?pstr - length ys\<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4098
    0\<^bsup>a_md  - Suc (?pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4099
    apply(rule_tac save_rs, auto simp: h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4100
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4101
  thm rec_ci.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4102
  thm empty_boxes.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4103
  from g have k6: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4104
    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4105
    length a + 3, ys @ 0\<^bsup>?pstr - length ys\<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4106
    0\<^bsup>a_md  - Suc (?pstr + length ys + n)\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4107
    aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4108
    (?gs_len + 8 * length gs + 3 *n + length a + 3,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4109
    0\<^bsup>?pstr \<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4110
                        0\<^bsup>a_md -Suc (?pstr + length ys + n)\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4111
    apply(drule_tac suf_lm = suf_lm in empty_paras, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4112
    apply(rule_tac x = stp in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4113
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4114
  from g have k7: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4115
    "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 3 *n + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4116
    length a + 3, 0\<^bsup>?pstr \<^esup> @ rs # 0\<^bsup>length ys\<^esup> @ lm @ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4117
    0\<^bsup>a_md -Suc (?pstr + length ys + n)\<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4118
    (?gs_len + 8 * length gs + 3 * n + length a + 6, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4119
    0\<^bsup>n\<^esup> @ rs # 0\<^bsup>?pstr  - n\<^esup> @ 0\<^bsup>length ys\<^esup> @ lm @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4120
                        0\<^bsup>a_md -Suc (?pstr + length ys + n) \<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4121
    apply(drule_tac suf_lm = suf_lm in restore_rs, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4122
    apply(rule_tac x = stp in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4123
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4124
  from g have k8: "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4125
    3 * n + length a + 6,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4126
    0\<^bsup>n\<^esup> @ rs # 0\<^bsup>?pstr  - n\<^esup> @ 0\<^bsup>length ys\<^esup> @ lm @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4127
                      0\<^bsup>a_md -Suc (?pstr + length ys + n) \<^esup> @ suf_lm) aprog stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4128
    (?gs_len + 8 * length gs + 6 * n + length a + 6,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4129
                           lm @ rs # 0\<^bsup>a_md - Suc n \<^esup>@ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4130
    apply(drule_tac suf_lm = suf_lm in restore_paras, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4131
    apply(simp add: exponent_add_iff)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4132
    apply(rule_tac x = stp in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4133
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4134
  from g have j1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4135
    "length aprog = ?gs_len + 8 * length gs + 6 * n + length a + 6"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4136
    by(drule_tac a = a and aa = b and ba = c in ci_cn_length,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4137
      simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4138
  from g have j2: "rs_pos = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4139
    by(simp add: ci_cn_para_eq)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4140
  from k1 and k2 and k3 and k4 and k5 and k6 and k7 and k8
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4141
    and j1 and j2 show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4142
    "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4143
    (length aprog, lm @ [rs] @ 0\<^bsup>a_md - rs_pos - 1\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4144
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4145
    apply(rule_tac x = "stp + stpa + stpb + stpc +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4146
      stpd + stpe + stpf + stpg" in exI, simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4147
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4148
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4149
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4150
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4151
  Correctness of the complier (terminate case), which says if the execution of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4152
  a recursive function @{text "recf"} terminates and gives result, then 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4153
  the Abacus program compiled from @{text "recf"} termintes and gives the same result.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4154
  Additionally, to facilitate induction proof, we append @{text "anything"} to the
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4155
  end of Abacus memory.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4156
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4157
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4158
lemma aba_rec_equality:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4159
  "\<lbrakk>rec_ci recf = (ap, arity, fp);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4160
    rec_calc_rel recf args r\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4161
  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4162
              (length ap, args@[r]@0\<^bsup>fp - arity - 1\<^esup> @ anything))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4163
apply(induct arbitrary: ap fp arity r anything args
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4164
  rule: rec_ci.induct)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4165
prefer 5
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4166
proof(case_tac "rec_ci g", case_tac "rec_ci f", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4167
  fix n f g ap fp arity r anything args  a b c aa ba ca
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4168
  assume f_ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4169
    "\<And>ap fp arity r anything args.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4170
    \<lbrakk>aa = ap \<and> ba = arity \<and> ca = fp; rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4171
    \<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4172
    (length ap, args @ r # 0\<^bsup>fp - Suc arity\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4173
    and g_ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4174
    "\<And>x xa y xb ya ap fp arity r anything args.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4175
    \<lbrakk>x = (aa, ba, ca); xa = aa \<and> y = (ba, ca); xb = ba \<and> ya = ca; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4176
    a = ap \<and> b = arity \<and> c = fp; rec_calc_rel g args r\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4177
    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4178
    (length ap, args @ r # 0\<^bsup>fp - Suc arity\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4179
    and h: "rec_ci (Pr n f g) = (ap, arity, fp)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4180
    "rec_calc_rel (Pr n f g) args r" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4181
    "rec_ci g = (a, b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4182
    "rec_ci f = (aa, ba, ca)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4183
  from h have nf_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4184
    "\<And> args r anything. rec_calc_rel f args r \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4185
    \<exists>stp. abc_steps_l (0, args @ 0\<^bsup>ca - ba\<^esup> @ anything) aa stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4186
    (length aa, args @ r # 0\<^bsup>ca - Suc ba\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4187
    and ng_ind: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4188
    "\<And> args r anything. rec_calc_rel g args r \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4189
    \<exists>stp. abc_steps_l (0, args @ 0\<^bsup>c - b\<^esup> @ anything) a stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4190
         (length a, args @ r # 0\<^bsup>c - Suc b \<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4191
    apply(insert f_ind[of aa ba ca], simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4192
    apply(insert g_ind[of "(aa, ba, ca)" aa "(ba, ca)" ba ca a b c],
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
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4195
  from nf_ind and ng_ind and h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4196
    "\<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4197
    (length ap, args @ r # 0\<^bsup>fp - Suc arity\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4198
    apply(auto intro: nf_ind ng_ind pr_case)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4199
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4200
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4201
  fix ap fp arity r anything args
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4202
  assume h:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4203
    "rec_ci z = (ap, arity, fp)" "rec_calc_rel z args r"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4204
  thus "\<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4205
    (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4206
    by (rule_tac z_case)    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4207
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4208
  fix ap fp arity r anything args
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4209
  assume h: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4210
    "rec_ci s = (ap, arity, fp)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4211
    "rec_calc_rel s args r"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4212
  thus 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4213
    "\<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4214
    (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4215
    by(erule_tac s_case, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4216
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4217
  fix m n ap fp arity r anything args
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4218
  assume h: "rec_ci (id m n) = (ap, arity, fp)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4219
    "rec_calc_rel (id m n) args r"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4220
  thus "\<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4221
    = (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4222
    by(erule_tac id_case)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4223
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4224
  fix n f gs ap fp arity r anything args
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4225
  assume ind: "\<And>x ap fp arity r anything args.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4226
    \<lbrakk>x \<in> set (f # gs); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4227
    rec_ci x = (ap, arity, fp); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4228
    rec_calc_rel x args r\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4229
    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4230
    (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4231
  and h: "rec_ci (Cn n f gs) = (ap, arity, fp)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4232
    "rec_calc_rel (Cn n f gs) args r"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4233
  from h show
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4234
    "\<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4235
       ap stp = (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4236
    apply(rule_tac cn_case, rule_tac ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4237
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4238
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4239
  fix n f ap fp arity r anything args
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4240
  assume ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4241
    "\<And>ap fp arity r anything args.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4242
    \<lbrakk>rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4243
    \<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4244
    (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4245
  and h: "rec_ci (Mn n f) = (ap, arity, fp)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4246
    "rec_calc_rel (Mn n f) args r"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4247
  from h show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4248
    "\<exists>stp. abc_steps_l (0, args @ 0\<^bsup>fp - arity\<^esup> @ anything) ap stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4249
              (length ap, args @ [r] @ 0\<^bsup>fp - arity - 1\<^esup> @ anything)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4250
    apply(rule_tac mn_case, rule_tac ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4251
    done    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4252
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4253
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4254
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4255
thm abc_append_state_in_exc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4256
lemma abc_append_uhalt1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4257
  "\<lbrakk>\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4258
    p = ap [+] bp [+] cp\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4259
  \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4260
                     (abc_steps_l (length ap, lm) p stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4261
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4262
apply(erule_tac x = stp in allE, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4263
apply(frule_tac ap = ap and cp = cp in abc_append_state_in_exc, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4264
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4265
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4266
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4267
lemma abc_append_unhalt2:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4268
  "\<lbrakk>abc_steps_l (0, am) ap stp = (length ap, lm); bp \<noteq> [];
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4269
  \<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4270
  p = ap [+] bp [+] cp\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4271
  \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) (abc_steps_l (0, am) p stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4272
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4273
  assume h: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4274
    "abc_steps_l (0, am) ap stp = (length ap, lm)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4275
    "bp \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4276
    "\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4277
    "p = ap [+] bp [+] cp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4278
  have "\<exists> stp. (abc_steps_l (0, am) p stp) = (length ap, lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4279
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4280
    thm abc_add_exc1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4281
    apply(simp add: abc_append.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4282
    apply(rule_tac abc_add_exc1, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4283
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4284
  from this obtain stpa where g1: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4285
    "(abc_steps_l (0, am) p stpa) = (length ap, lm)" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4286
  moreover have g2: "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4287
                          (abc_steps_l (length ap, lm) p stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4288
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4289
    apply(erule_tac abc_append_uhalt1, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4290
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4291
  moreover from g1 and g2 have
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4292
    "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4293
                    (abc_steps_l (0, am) p (stpa + stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4294
    apply(simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4295
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4296
  thus "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4297
                           (abc_steps_l (0, am) p stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4298
    apply(rule_tac allI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4299
    apply(case_tac "stp \<ge>  stpa")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4300
    apply(erule_tac x = "stp - stpa" in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4301
  proof - 	
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4302
    fix stp a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4303
    assume g3:  "abc_steps_l (0, am) p stp = (a, b)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4304
                "\<not> stpa \<le> stp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4305
    thus "a < length p"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4306
      using g1 h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4307
      apply(case_tac "a < length p", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4308
      apply(subgoal_tac "\<exists> d. stpa = stp + d")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4309
      using  abc_state_keep[of p a b "stpa - stp"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4310
      apply(erule_tac exE, simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4311
      apply(rule_tac x = "stpa - stp" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4312
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4313
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4314
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4315
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4316
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4317
  Correctness of the complier (non-terminating case for Mn). There are many cases when a 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4318
  recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4319
  need to prove the case for @{text "Mn"} and @{text "Cn"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4320
  This lemma is for @{text "Mn"}. For @{text "Mn n f"}, this lemma describes what 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4321
  happens when @{text "f"} always terminates but always does not return zero, so that
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4322
  @{text "Mn"} has to loop forever.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4323
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4324
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4325
lemma Mn_unhalt:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4326
  assumes mn_rf: "rf = Mn n f"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4327
  and compiled_mnrf: "rec_ci rf = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4328
  and compiled_f: "rec_ci f = (aprog', rs_pos', a_md')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4329
  and args: "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4330
  and unhalt_condition: "\<forall> y. (\<exists> rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4331
  shows "\<forall> stp. case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4332
               aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4333
  using mn_rf compiled_mnrf compiled_f args unhalt_condition
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4334
proof(rule_tac allI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4335
  fix stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4336
  assume h: "rf = Mn n f" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4337
            "rec_ci rf = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4338
            "rec_ci f = (aprog', rs_pos', a_md')" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4339
            "\<forall>y. \<exists>rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0" "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4340
  thm mn_ind_step
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4341
  have "\<exists>stpa \<ge> stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) aprog stpa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4342
         = (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4343
  proof(induct stp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4344
    show "\<exists>stpa. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4345
          aprog stpa = (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4346
      apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4347
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4348
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4349
    fix stp stpa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4350
    assume g1: "stp \<le> stpa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4351
      and g2: "abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4352
                            aprog stpa
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4353
               = (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4354
    have "\<exists>rs. rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4355
      using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4356
      apply(erule_tac x = stp in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4357
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4358
    from this obtain rs where g3:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4359
      "rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4360
    hence "\<exists> stpb. abc_steps_l (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4361
                     suf_lm) aprog stpb 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4362
      = (0, lm @ Suc stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4363
      using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4364
      apply(rule_tac mn_ind_step)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4365
      apply(rule_tac aba_rec_equality, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4366
    proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4367
      show "rec_ci f = ((aprog', rs_pos', a_md'))" using h by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4368
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4369
      show "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" using h by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4370
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4371
      show "rec_calc_rel f (lm @ [stp]) rs" using g3 by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4372
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4373
      show "0 < rs" using g3 by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4374
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4375
      show "Suc rs_pos < a_md"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4376
        using g3 h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4377
        apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4378
        apply(frule_tac f = f in para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4379
        apply(simp add: rec_ci.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4380
        apply(subgoal_tac "Suc (length lm) < a_md'")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4381
        apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4382
        apply(simp add: ci_ad_ge_paras)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4383
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4384
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4385
      show "rs_pos' = Suc rs_pos"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4386
        using g3 h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4387
        apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4388
        apply(frule_tac f = f in para_pattern, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4389
        apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4390
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4391
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4392
    thus "\<exists>stpa\<ge>Suc stp. abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4393
                 suf_lm) aprog stpa 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4394
      = (0, lm @ Suc stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4395
      using g2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4396
      apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4397
      apply(case_tac "stpb = 0", simp add: abc_steps_l.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4398
      apply(rule_tac x = "stpa + stpb" in exI, simp add:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4399
        abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4400
      using g1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4401
      apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4402
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4403
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4404
  from this obtain stpa where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4405
    "stp \<le> stpa \<and> abc_steps_l (0, lm @ 0 # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4406
         aprog stpa = (0, lm @ stp # 0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm)" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4407
  thus "case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4408
    of (ss, e) \<Rightarrow> ss < length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4409
    apply(case_tac "abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm) aprog
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4410
      stp", simp, case_tac "a \<ge> length aprog", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4411
        simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4412
    apply(subgoal_tac "\<exists> d. stpa = stp + d", erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4413
    apply(subgoal_tac "lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suf_lm = lm @ 0 # 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4414
             0\<^bsup>a_md - Suc rs_pos\<^esup> @ suf_lm", simp add: abc_steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4415
    apply(frule_tac as = a and lm = b and stp = d in abc_state_keep, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4416
          simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4417
    using h  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4418
    apply(simp add: rec_ci.simps, simp, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4419
              simp only: exp_ind_def[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4420
    apply(case_tac rs_pos, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4421
    apply(rule_tac x = "stpa - stp" in exI, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4422
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4423
qed   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4424
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4425
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4426
lemma abc_append_cons_eq[intro!]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4427
  "\<lbrakk>ap = bp; cp = dp\<rbrakk> \<Longrightarrow> ap [+] cp = bp [+] dp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4428
by simp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4429
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4430
lemma cn_merge_gs_split: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4431
  "\<lbrakk>i < length gs; rec_ci (gs!i) = (ga, gb, gc)\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4432
     cn_merge_gs (map rec_ci gs) p = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4433
        cn_merge_gs (map rec_ci (take i gs)) p [+] ga [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4434
       empty gb (p + i) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4435
      cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4436
apply(induct i arbitrary: gs p, case_tac gs, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4437
apply(case_tac gs, simp, case_tac "rec_ci a", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4438
       simp add: abc_append_commute[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4439
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4440
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4441
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4442
  Correctness of the complier (non-terminating case for Mn). There are many cases when a 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4443
  recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4444
  need to prove the case for @{text "Mn"} and @{text "Cn"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4445
  This lemma is for @{text "Cn"}. For @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}, this lemma describes what 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4446
  happens when every one of @{text "g1, g2, \<dots> gi"} terminates, but 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4447
  @{text "gi+1"} does not terminate, so that whole function @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4448
  does not terminate.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4449
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4450
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4451
lemma cn_gi_uhalt: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4452
  assumes cn_recf: "rf = Cn n f gs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4453
  and compiled_cn_recf: "rec_ci rf = (aprog, rs_pos, a_md)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4454
  and args_length: "length lm = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4455
  and exist_unhalt_recf: "i < length gs" "gi = gs ! i"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4456
  and complied_unhalt_recf: "rec_ci gi = (ga, gb, gc)"  "gb = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4457
  and all_halt_before_gi: "\<forall> j < i. (\<exists> rs. rec_calc_rel (gs!j) lm rs)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4458
  and unhalt_condition: "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<^bsup>gc - gb\<^esup> @ slm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4459
     ga stp of (se, e) \<Rightarrow> se < length ga"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4460
  shows " \<forall> stp. case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suflm) aprog
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4461
  stp of (ss, e) \<Rightarrow> ss < length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4462
  using cn_recf compiled_cn_recf args_length exist_unhalt_recf complied_unhalt_recf
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4463
        all_halt_before_gi unhalt_condition
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4464
proof(case_tac "rec_ci f", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4465
  fix a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4466
  assume h1: "rf = Cn n f gs" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4467
    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4468
    "length lm = n" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4469
    "gi = gs ! i" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4470
    "rec_ci (gs!i) = (ga, n, gc)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4471
    "gb = n" "rec_ci f = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4472
    and h2: "\<forall>j<i. \<exists>rs. rec_calc_rel (gs ! j) lm rs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4473
    "i < length gs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4474
  and ind:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4475
    "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<^bsup>gc - n\<^esup> @ slm) ga stp of (se, e) \<Rightarrow> se < length ga"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4476
  have h3: "rs_pos = n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4477
    using h1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4478
    by(rule_tac ci_cn_para_eq, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4479
  let ?ggs = "take i gs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4480
  have "\<exists> ys. (length ys = i \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4481
    (\<forall> k < i. rec_calc_rel (?ggs ! k) lm (ys ! k)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4482
    using h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4483
    apply(induct i, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4484
    apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4485
    apply(erule_tac x = ia in allE, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4486
    apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4487
    apply(rule_tac x = "ys @ [x]" in exI, simp add: nth_append, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4488
    apply(subgoal_tac "k = length ys", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4489
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4490
  from this obtain ys where g1:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4491
    "(length ys = i \<and> (\<forall> k < i. rec_calc_rel (?ggs ! k)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4492
                        lm (ys ! k)))" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4493
  let ?pstr = "Max (set (Suc n # c # map (\<lambda>(aprog, p, n). n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4494
    (map rec_ci (f # gs))))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4495
  have "\<exists>stp. abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suflm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4496
    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stp =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4497
    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4498
    3 * length ?ggs, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -(?pstr + length ?ggs)\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4499
    suflm) "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4500
    apply(rule_tac  cn_merge_gs_ex)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4501
    apply(rule_tac  aba_rec_equality, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4502
    using h1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4503
    apply(simp add: rec_ci.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4504
    using g1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4505
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4506
    using h2 g1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4507
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4508
    apply(rule_tac min_max.le_supI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4509
    apply(rule_tac Max_ge, simp, simp, rule_tac disjI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4510
    apply(subgoal_tac "aa \<in> set gs", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4511
    using h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4512
    apply(rule_tac A = "set (take i gs)" in subsetD, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4513
      simp add: set_take_subset, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4514
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4515
  thm cn_merge_gs.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4516
  from this obtain stpa where g2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4517
    "abc_steps_l (0, lm @ 0\<^bsup>a_md - n\<^esup> @ suflm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4518
    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4519
    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4520
    3 * length ?ggs, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -(?pstr + length ?ggs)\<^esup> @
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4521
    suflm)" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4522
  moreover have 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4523
    "\<exists> cp. aprog = (cn_merge_gs
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4524
    (map rec_ci ?ggs) ?pstr) [+] ga [+] cp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4525
    using h1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4526
    apply(simp add: rec_ci.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4527
    apply(rule_tac x = "empty n (?pstr + i) [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4528
      (cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?pstr + Suc i))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4529
      [+]mv_boxes 0 (Suc (max (Suc n) (Max (insert c 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4530
     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) +
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4531
      length gs)) n [+] mv_boxes (max (Suc n) (Max (insert c 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4532
      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4533
      a [+] recursive.empty b (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4534
      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4535
     empty_boxes (length gs) [+] recursive.empty (max (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4536
      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4537
      mv_boxes (Suc (max (Suc n) (Max (insert c 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4538
    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4539
    apply(simp add: abc_append_commute [THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4540
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4541
    using cn_merge_gs_split[of i gs ga "length lm" gc 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4542
      "(max (Suc (length lm))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4543
       (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))"] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4544
      h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4545
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4546
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4547
  from this obtain cp where g3: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4548
    "aprog = (cn_merge_gs (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4549
  show "\<forall> stp. case abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suflm) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4550
    aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4551
  proof(rule_tac abc_append_unhalt2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4552
    show "abc_steps_l (0, lm @ 0\<^bsup>a_md - rs_pos\<^esup> @ suflm) (
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4553
      cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4554
         (length ((cn_merge_gs (map rec_ci ?ggs) ?pstr)),  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4555
          lm @ 0\<^bsup>?pstr - n\<^esup> @ ys @ 0\<^bsup>a_md -(?pstr + length ?ggs)\<^esup> @ suflm)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4556
      using h3 g2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4557
      apply(simp add: cn_merge_gs_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4558
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4559
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4560
    show "ga \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4561
      using h1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4562
      apply(simp add: rec_ci_not_null)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4563
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4564
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4565
    show "\<forall>stp. case abc_steps_l (0, lm @ 0\<^bsup>?pstr - n\<^esup> @ ys
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4566
      @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup> @ suflm) ga  stp of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4567
          (ss, e) \<Rightarrow> ss < length ga"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4568
      using ind[of "0\<^bsup>?pstr -gc\<^esup> @ ys @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4569
        @ suflm"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4570
      apply(subgoal_tac "lm @ 0\<^bsup>?pstr - n\<^esup> @ ys
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4571
        @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup> @ suflm
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4572
                       = lm @ 0\<^bsup>gc - n \<^esup>@ 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4573
        0\<^bsup>?pstr -gc\<^esup> @ ys @ 0\<^bsup>a_md - (?pstr + length (take i gs))\<^esup> @ suflm", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4574
      apply(simp add: exponent_def replicate_add[THEN sym])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4575
      apply(subgoal_tac "gc > n \<and> ?pstr \<ge> gc")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4576
      apply(erule_tac conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4577
      apply(simp add: h1)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4578
      using h1
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4579
      apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4580
      apply(rule_tac min_max.le_supI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4581
      apply(rule_tac Max_ge, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4582
      apply(rule_tac disjI2)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4583
      using h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4584
      thm rev_image_eqI
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4585
      apply(rule_tac x = "gs!i" in rev_image_eqI, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4586
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4587
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4588
    show "aprog = cn_merge_gs (map rec_ci (take i gs)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4589
              ?pstr [+] ga [+] cp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4590
      using g3 by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4591
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4592
qed
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4595
lemma abc_rec_halt_eq': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4596
  "\<lbrakk>rec_ci re = (ap, ary, fp); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4597
    rec_calc_rel re args r\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4598
  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<^bsup>fp - ary\<^esup>) ap stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4599
                     (length ap, args@[r]@0\<^bsup>fp - ary - 1\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4600
using aba_rec_equality[of re ap ary fp args r "[]"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4601
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4602
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4603
thm abc_step_l.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4604
definition dummy_abc :: "nat \<Rightarrow> abc_inst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4605
where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4606
"dummy_abc k = [Inc k, Dec k 0, Goto 3]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4607
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4608
lemma abc_rec_halt_eq'': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4609
  "\<lbrakk>rec_ci re = (aprog, rs_pos, a_md);  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4610
  rec_calc_rel re lm rs\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4611
  \<Longrightarrow> (\<exists> stp lm' m. (abc_steps_l (0, lm) aprog stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4612
  (length aprog, lm') \<and> abc_list_crsp lm' (lm @ rs # 0\<^bsup>m\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4613
apply(frule_tac abc_rec_halt_eq', auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4614
apply(drule_tac abc_list_crsp_steps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4615
apply(rule_tac rec_ci_not_null, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4616
apply(erule_tac exE, rule_tac x = stp in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4617
  auto simp: abc_list_crsp_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4618
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4619
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4620
lemma [simp]: "length (dummy_abc (length lm)) = 3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4621
apply(simp add: dummy_abc_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4622
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4623
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4624
lemma [simp]: "dummy_abc (length lm) \<noteq> []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4625
apply(simp add: dummy_abc_def)
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 dummy_abc_steps_ex: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4629
  "\<exists>bstp. abc_steps_l (0, lm') (dummy_abc (length lm)) bstp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4630
  ((Suc (Suc (Suc 0))), abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4631
apply(rule_tac x = "Suc (Suc (Suc 0))" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4632
apply(auto simp: abc_steps_l.simps abc_step_l.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4633
  dummy_abc_def abc_fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4634
apply(auto simp: abc_lm_s.simps abc_lm_v.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4635
apply(simp add: butlast_append)
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 [elim]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4639
  "lm @ rs # 0\<^bsup>m\<^esup> = lm' @ 0\<^bsup>n\<^esup> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4640
  \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4641
                            lm @ rs # 0\<^bsup>m\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4642
proof(cases "length lm' > length lm")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4643
  case True 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4644
  assume h: "lm @ rs # 0\<^bsup>m\<^esup> = lm' @ 0\<^bsup>n\<^esup>" "length lm < length lm'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4645
  hence "m \<ge> n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4646
    apply(drule_tac list_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4647
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4648
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4649
  hence "\<exists> d. m = d + n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4650
    apply(rule_tac x = "m - n" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4651
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4652
  from this obtain d where "m = d + n" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4653
  from h and this show "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4654
    apply(auto simp: abc_lm_s.simps abc_lm_v.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4655
                     exponent_def replicate_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4656
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4657
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4658
  case False
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4659
  assume h:"lm @ rs # 0\<^bsup>m\<^esup> = lm' @ 0\<^bsup>n\<^esup>" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4660
    and    g: "\<not> length lm < length lm'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4661
  have "take (Suc (length lm)) (lm @ rs # 0\<^bsup>m\<^esup>) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4662
                        take (Suc (length lm)) (lm' @ 0\<^bsup>n\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4663
    using h by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4664
  moreover have "n \<ge> (Suc (length lm) - length lm')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4665
    using h g
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4666
    apply(drule_tac list_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4667
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4668
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4669
  ultimately show 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4670
    "\<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4671
                                                       lm @ rs # 0\<^bsup>m\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4672
    using g h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4673
    apply(simp add: abc_lm_s.simps abc_lm_v.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4674
                                        exponent_def min_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4675
    apply(rule_tac x = 0 in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4676
      simp add:replicate_append_same replicate_Suc[THEN sym]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4677
                                      del:replicate_Suc)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4678
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4679
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4680
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4681
lemma [elim]: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4682
  "abc_list_crsp lm' (lm @ rs # 0\<^bsup>m\<^esup>)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4683
  \<Longrightarrow> \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4684
             = lm @ rs # 0\<^bsup>m\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4685
apply(auto simp: abc_list_crsp_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4686
apply(simp add: abc_lm_v.simps abc_lm_s.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4687
apply(rule_tac x =  "m + n" in exI, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4688
      simp add: exponent_def replicate_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4689
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4690
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4691
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4692
lemma abc_append_dummy_complie:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4693
  "\<lbrakk>rec_ci recf = (ap, ary, fp);  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4694
    rec_calc_rel recf args r; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4695
    length args = k\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4696
  \<Longrightarrow> (\<exists> stp m. (abc_steps_l (0, args) (ap [+] dummy_abc k) stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4697
                  (length ap + 3, args @ r # 0\<^bsup>m\<^esup>))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4698
apply(drule_tac abc_rec_halt_eq'', auto simp: numeral_3_eq_3)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4699
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4700
  fix stp lm' m
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4701
  assume h: "rec_calc_rel recf args r"  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4702
    "abc_steps_l (0, args) ap stp = (length ap, lm')" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4703
    "abc_list_crsp lm' (args @ r # 0\<^bsup>m\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4704
  thm abc_append_exc2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4705
  thm abc_lm_s.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4706
  have "\<exists>stp. abc_steps_l (0, args) (ap [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4707
    (dummy_abc (length args))) stp = (length ap + 3, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4708
    abc_lm_s lm' (length args) (abc_lm_v lm' (length args)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4709
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4710
    apply(rule_tac bm = lm' in abc_append_exc2,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4711
          auto intro: dummy_abc_steps_ex simp: numeral_3_eq_3)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4712
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4713
  thus "\<exists>stp m. abc_steps_l (0, args) (ap [+] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4714
    dummy_abc (length args)) stp = (Suc (Suc (Suc (length ap))), args @ r # 0\<^bsup>m\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4715
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4716
    apply(erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4717
    apply(rule_tac x = stpa in exI, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4718
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4719
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4720
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4721
lemma [simp]: "length (dummy_abc k) = 3"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4722
apply(simp add: dummy_abc_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4723
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4724
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4725
lemma [simp]: "length args = k \<Longrightarrow> abc_lm_v (args @ r # 0\<^bsup>m\<^esup>) k = r "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4726
apply(simp add: abc_lm_v.simps nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4727
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4728
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4729
lemma t_compiled_by_rec: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4730
  "\<lbrakk>rec_ci recf = (ap, ary, fp); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4731
    rec_calc_rel recf args r;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4732
    length args = k;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4733
    ly = layout_of (ap [+] dummy_abc k);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4734
    mop_ss = start_of ly (length (ap [+] dummy_abc k));
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4735
    tp = tm_of (ap [+] dummy_abc k)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4736
  \<Longrightarrow> \<exists> stp m l. steps (Suc 0, Bk # Bk # ires, <args> @ Bk\<^bsup>rn\<^esup>) (tp @ (tMp k (mop_ss - 1))) stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4737
                      = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc r\<^esup> @ Bk\<^bsup>l\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4738
  using abc_append_dummy_complie[of recf ap ary fp args r k]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4739
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4740
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4741
apply(frule_tac tprog = tp and as = "length ap + 3" and n = k 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4742
               and ires = ires and rn = rn in abacus_turing_eq_halt, simp_all, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4743
apply(erule_tac exE)+
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4744
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4745
apply(rule_tac x = stpa in exI, rule_tac x = ma in exI, rule_tac x = l in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4746
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4747
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4748
thm tms_of.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4749
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4750
lemma [simp]:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4751
  "list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4752
  list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4753
apply(induct xs, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4754
apply(case_tac a, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4755
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4756
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4757
(*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4758
lemma [simp]: "t_correct (tMp n 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4759
apply(simp add: t_correct.simps tMp.simps shift_length mp_up_def iseven_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4760
apply(rule_tac x = "2*n + 6" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4761
apply(induct n, auto simp: mop_bef.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4762
apply(simp add: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4763
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4764
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4765
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4766
lemma tshift_append: "tshift (xs @ ys) n = tshift xs n @ tshift ys n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4767
apply(simp add: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4768
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4769
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4770
lemma [simp]: "length (tMp n ss) = 4 * n + 12"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4771
apply(auto simp: tMp.simps tshift_append shift_length mp_up_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4772
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4773
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4774
lemma length_tm_even[intro]: "\<exists>x. length (tm_of ap) = 2*x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4775
apply(subgoal_tac "t_ncorrect (tm_of ap)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4776
apply(simp add: t_ncorrect.simps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4777
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4778
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4779
lemma [simp]: "k < length ap \<Longrightarrow> tms_of ap ! k  = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4780
 ci (layout_of ap) (start_of (layout_of ap) k) (ap ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4781
apply(simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4782
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4783
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4784
lemma [elim]: "\<lbrakk>k < length ap; ap ! k = Inc n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4785
       (a, b) \<in> set (abacus.tshift (abacus.tshift tinc_b (2 * n)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4786
                            (start_of (layout_of ap) k - Suc 0))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4787
       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4788
apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4789
apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4790
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4791
apply(case_tac "Suc k = length ap", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4792
apply(rule_tac start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4793
apply(auto simp: tinc_b_def tshift.simps start_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4794
  layout_of.simps length_of.simps startof_not0)
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4797
lemma findnth_le[elim]: "(a, b) \<in> set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4798
        \<Longrightarrow> b \<le> Suc (start_of (layout_of ap) k + 2 * n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4799
apply(induct n, simp add: findnth.simps tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4800
apply(simp add: findnth.simps tshift_append, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4801
apply(auto simp: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4802
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4803
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4804
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4805
lemma  [elim]: "\<lbrakk>k < length ap; ap ! k = Inc n; (a, b) \<in> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4806
  set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4807
  \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4808
apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4809
apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4810
apply(arith)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4811
apply(case_tac "Suc k = length ap", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4812
apply(rule_tac start_of_le, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4813
apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4814
     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k)", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4815
apply(auto simp: tinc_b_def tshift.simps start_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4816
  layout_of.simps length_of.simps startof_not0)
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 start_of_eq: "length ap < as \<Longrightarrow> start_of (layout_of ap) as = start_of (layout_of ap) (length ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4820
apply(induct as, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4821
apply(case_tac "length ap < as", simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4822
apply(subgoal_tac "as = length ap")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4823
apply(simp add: start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4824
apply arith
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4825
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4826
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4827
lemma start_of_all_le: "start_of (layout_of ap) as \<le> start_of (layout_of ap) (length ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4828
apply(subgoal_tac "as > length ap \<or> as = length ap \<or> as < length ap", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4829
      auto simp: start_of_eq start_of_le)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4830
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4831
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4832
lemma [elim]: "\<lbrakk>k < length ap; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4833
        ap ! k = Dec n e;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4834
         (a, b) \<in> set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4835
       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4836
apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4837
     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k) \<and>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4838
      start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4839
apply(simp add:  tshift.simps start_of.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4840
  layout_of.simps length_of.simps startof_not0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4841
apply(rule_tac start_of_all_le)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4842
done
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
thm length_of.simps
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4845
lemma [elim]: "\<lbrakk>k < length ap; ap ! k = Dec n e; (a, b) \<in> set (abacus.tshift (abacus.tshift tdec_b (2 * n))
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4846
                  (start_of (layout_of ap) k - Suc 0))\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4847
       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4848
apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \<le> start_of (layout_of ap) (length ap) \<and> start_of (layout_of ap) k > 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4849
prefer 2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4850
apply(subgoal_tac "2 * n + start_of (layout_of ap) k + 16 = start_of (layout_of ap) (Suc k)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4851
                 \<and> start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4852
apply(simp add: startof_not0, rule_tac conjI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4853
apply(simp add: start_of.simps layout_of.simps length_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4854
apply(rule_tac start_of_all_le)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4855
apply(auto simp: tdec_b_def tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4856
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4857
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4858
lemma tms_any_less: "\<lbrakk>k < length ap; (a, b) \<in> set (tms_of ap ! k)\<rbrakk> \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4859
apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4860
apply(case_tac "ap!k", simp_all add: ci.simps tshift_append, auto intro: start_of_all_le)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4861
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4862
lemma concat_in: "i < length (concat xs) \<Longrightarrow> \<exists>k < length xs. concat xs ! i \<in> set (xs ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4863
apply(induct xs rule: list_tl_induct, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4864
apply(case_tac "i < length (concat list)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4865
apply(erule_tac exE, rule_tac x = k in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4866
apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4867
apply(rule_tac x = "length list" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4868
apply(simp add: nth_append)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4869
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4870
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4871
lemma [simp]: "length (tms_of ap) = length ap"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4872
apply(simp add: tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4873
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4874
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4875
lemma in_tms: "i < length (tm_of ap) \<Longrightarrow> \<exists> k < length ap. (tm_of ap ! i) \<in> set (tms_of ap ! k)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4876
apply(simp add: tm_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4877
using concat_in[of i "tms_of ap"]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4878
by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4879
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4880
lemma all_le_start_of: "list_all (\<lambda>(acn, s). s \<le> start_of (layout_of ap) (length ap)) (tm_of ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4881
apply(simp add: list_all_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4882
apply(rule_tac allI, rule_tac impI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4883
apply(drule_tac in_tms, auto elim: tms_any_less)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4884
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4885
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4886
lemma length_ci: "\<lbrakk>k < length ap; length (ci ly y (ap ! k)) = 2 * qa\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4887
      \<Longrightarrow> layout_of ap ! k = qa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4888
apply(case_tac "ap ! k")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4889
apply(auto simp: layout_of.simps ci.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4890
  length_of.simps shift_length tinc_b_def tdec_b_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4891
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4892
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4893
lemma [intro]: "length (ci ly y i) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4894
apply(auto simp: ci.simps shift_length tinc_b_def tdec_b_def
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4895
      split: abc_inst.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4896
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4897
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4898
lemma [intro]: "listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) zs) mod 2 = 0"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4899
apply(induct zs rule: list_tl_induct, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4900
apply(case_tac a, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4901
apply(subgoal_tac "length (ci ly aa b) mod 2 = 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4902
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4903
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4904
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4905
lemma zip_pre:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4906
  "(length ys) \<le> length ap \<Longrightarrow>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4907
  zip ys ap = zip ys (take (length ys) (ap::'a list))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4908
proof(induct ys arbitrary: ap, simp, case_tac ap, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4909
  fix a ys ap aa list
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4910
  assume ind: "\<And>(ap::'a list). length ys \<le> length ap \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4911
    zip ys ap = zip ys (take (length ys) ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4912
  and h: "length (a # ys) \<le> length ap" "(ap::'a list) = aa # (list::'a list)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4913
  from h show "zip (a # ys) ap = zip (a # ys) (take (length (a # ys)) ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4914
    using ind[of list]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4915
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4916
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4917
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4918
 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4919
lemma start_of_listsum: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4920
  "\<lbrakk>k \<le> length ap; length ss = k\<rbrakk> \<Longrightarrow> start_of (layout_of ap) k = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4921
        Suc (listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) (zip ss ap)) div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4922
proof(induct k arbitrary: ss, simp add: start_of.simps, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4923
  fix k ss
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4924
  assume ind: "\<And>ss. length ss = k \<Longrightarrow> start_of (layout_of ap) k = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4925
            Suc (listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) (zip ss ap)) div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4926
  and h: "Suc k \<le>  length ap" "length (ss::nat list) = Suc k"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4927
  have "\<exists> ys y. ss = ys @ [y]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4928
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4929
    apply(rule_tac x = "butlast ss" in exI,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4930
          rule_tac x = "last ss" in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4931
    apply(case_tac "ss = []", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4932
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4933
  from this obtain ys y where k1: "ss = (ys::nat list) @ [y]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4934
    by blast
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4935
  from h and this have k2: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4936
    "start_of (layout_of ap) k = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4937
    Suc (listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) (zip ys ap)) div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4938
    apply(rule_tac ind, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4939
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4940
  have k3: "zip ys ap = zip ys (take k ap)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4941
    using zip_pre[of ys ap] k1 h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4942
    apply(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
  have k4: "(zip [y] (drop (length ys) ap)) = [(y, ap ! length ys)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4945
    using k1 h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4946
    apply(case_tac "drop (length ys) ap", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4947
    apply(subgoal_tac "hd (drop (length ys) ap) = ap ! length ys")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4948
    apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4949
    apply(rule_tac hd_drop_conv_nth, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4950
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4951
  from k1 and h k2 k3 k4 show "start_of (layout_of ap) (Suc k) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4952
    Suc (listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) (zip ss ap)) div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4953
    apply(simp add: zip_append1 start_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4954
    apply(subgoal_tac 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4955
      "listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) (zip ys (take k ap))) mod 2 = 0 \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4956
      length (ci ly y (ap!k)) mod 2 = 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4957
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4958
    apply(rule_tac length_ci, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4959
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4960
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4961
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4962
lemma length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap)  div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4963
apply(simp add: tm_of.simps length_concat tms_of.simps tpairs_of.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4964
apply(rule_tac start_of_listsum, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4965
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4966
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4967
lemma tm_even: "length (tm_of ap) mod 2 = 0" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4968
apply(subgoal_tac "t_ncorrect (tm_of ap)", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4969
apply(simp add: t_ncorrect.simps)
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4972
lemma [elim]: "list_all (\<lambda>(acn, s). s \<le> Suc q) xs
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4973
        \<Longrightarrow> list_all (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) xs"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4974
apply(simp add: list_all_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4975
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4976
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4977
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4978
lemma [simp]: "length mp_up = 12"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4979
apply(simp add: mp_up_def)
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
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4982
lemma [elim]: "\<lbrakk>na < 4 * n; tshift (mop_bef n) q ! na = (a, b)\<rbrakk> \<Longrightarrow> b \<le> q + (2 * n + 6)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4983
apply(induct n, simp, simp add: mop_bef.simps nth_append tshift_append shift_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4984
apply(case_tac "na < 4*n", simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4985
apply(subgoal_tac "na = 4*n \<or> na = 1 + 4*n \<or> na = 2 + 4*n \<or> na = 3 + 4*n",
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4986
  auto simp: shift_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4987
apply(simp_all add: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4988
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4989
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4990
lemma mp_up_all_le: "list_all  (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4991
  [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4992
  (L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4993
  (W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4994
  (W0, Suc (Suc (Suc (2 * n + q)))), (L, 5 + 2 * n + q),
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4995
  (L, 6 + 2 * n + q), (R, 0),  (L, 6 + 2 * n + q)]"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4996
apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4997
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4998
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  4999
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5000
lemma [intro]: "list_all (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) (tMp n q)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5001
apply(auto simp: list_all_length tMp.simps tshift_append nth_append shift_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5002
apply(auto simp: tshift.simps mp_up_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5003
apply(subgoal_tac "na - 4*n \<ge> 0 \<and> na - 4 *n < 12", auto split: nat.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5004
apply(insert mp_up_all_le[of q n])
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5005
apply(simp add: list_all_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5006
apply(erule_tac x = "na - 4 * n" in allE, simp add: numeral_3_eq_3)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5007
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5008
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5009
lemma t_compiled_correct: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5010
  "\<lbrakk>tp = tm_of ap; ly = layout_of ap; mop_ss = start_of ly (length ap)\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5011
       t_correct (tp @ tMp n (mop_ss - Suc 0))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5012
  using tm_even[of ap] length_start_of_tm[of ap] all_le_start_of[of ap]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5013
apply(auto simp: t_correct.simps iseven_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5014
apply(rule_tac x = "q + 2*n + 6" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5015
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5016
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5017
end
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5018
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5019
    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5020
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5021
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5022
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5023
  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5024