|
1 (* Title: Turing machines |
|
2 Author: Xu Jian <xujian817@hotmail.com> |
|
3 Maintainer: Xu Jian |
|
4 *) |
|
5 |
|
6 theory Turing |
|
7 imports Main |
|
8 begin |
|
9 |
|
10 section {* Basic definitions of Turing machine *} |
|
11 |
|
12 datatype action = W0 | W1 | L | R | Nop |
|
13 |
|
14 datatype cell = Bk | Oc |
|
15 |
|
16 type_synonym tape = "cell list \<times> cell list" |
|
17 |
|
18 type_synonym state = nat |
|
19 |
|
20 type_synonym instr = "action \<times> state" |
|
21 |
|
22 type_synonym tprog = "instr list \<times> nat" |
|
23 |
|
24 type_synonym tprog0 = "instr list" |
|
25 |
|
26 type_synonym config = "state \<times> tape" |
|
27 |
|
28 fun nth_of where |
|
29 "nth_of xs i = (if i \<ge> length xs then None |
|
30 else Some (xs ! i))" |
|
31 |
|
32 lemma nth_of_map [simp]: |
|
33 shows "nth_of (map f p) n = (case (nth_of p n) of None \<Rightarrow> None | Some x \<Rightarrow> Some (f x))" |
|
34 apply(induct p arbitrary: n) |
|
35 apply(auto) |
|
36 apply(case_tac n) |
|
37 apply(auto) |
|
38 done |
|
39 |
|
40 fun |
|
41 fetch :: "instr list \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr" |
|
42 where |
|
43 "fetch p 0 b = (Nop, 0)" |
|
44 | "fetch p (Suc s) Bk = |
|
45 (case nth_of p (2 * s) of |
|
46 Some i \<Rightarrow> i |
|
47 | None \<Rightarrow> (Nop, 0))" |
|
48 |"fetch p (Suc s) Oc = |
|
49 (case nth_of p ((2 * s) + 1) of |
|
50 Some i \<Rightarrow> i |
|
51 | None \<Rightarrow> (Nop, 0))" |
|
52 |
|
53 lemma fetch_Nil [simp]: |
|
54 shows "fetch [] s b = (Nop, 0)" |
|
55 apply(case_tac s) |
|
56 apply(auto) |
|
57 apply(case_tac b) |
|
58 apply(auto) |
|
59 done |
|
60 |
|
61 fun |
|
62 update :: "action \<Rightarrow> tape \<Rightarrow> tape" |
|
63 where |
|
64 "update W0 (l, r) = (l, Bk # (tl r))" |
|
65 | "update W1 (l, r) = (l, Oc # (tl r))" |
|
66 | "update L (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" |
|
67 | "update R (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" |
|
68 | "update Nop (l, r) = (l, r)" |
|
69 |
|
70 abbreviation |
|
71 "read r == if (r = []) then Bk else hd r" |
|
72 |
|
73 fun step :: "config \<Rightarrow> tprog \<Rightarrow> config" |
|
74 where |
|
75 "step (s, l, r) (p, off) = |
|
76 (let (a, s') = fetch p (s - off) (read r) in (s', update a (l, r)))" |
|
77 |
|
78 fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config" |
|
79 where |
|
80 "steps c p 0 = c" | |
|
81 "steps c p (Suc n) = steps (step c p) p n" |
|
82 |
|
83 |
|
84 abbreviation |
|
85 "step0 c p \<equiv> step c (p, 0)" |
|
86 |
|
87 abbreviation |
|
88 "steps0 c p n \<equiv> steps c (p, 0) n" |
|
89 |
|
90 lemma step_red [simp]: |
|
91 shows "steps c p (Suc n) = step (steps c p n) p" |
|
92 by (induct n arbitrary: c) (auto) |
|
93 |
|
94 lemma steps_add [simp]: |
|
95 shows "steps c p (m + n) = steps (steps c p m) p n" |
|
96 by (induct m arbitrary: c) (auto) |
|
97 |
|
98 lemma step_0 [simp]: |
|
99 shows "step (0, (l, r)) p = (0, (l, r))" |
|
100 by (case_tac p, simp) |
|
101 |
|
102 lemma steps_0 [simp]: |
|
103 shows "steps (0, (l, r)) p n = (0, (l, r))" |
|
104 by (induct n) (simp_all) |
|
105 |
|
106 |
|
107 |
|
108 fun |
|
109 is_final :: "config \<Rightarrow> bool" |
|
110 where |
|
111 "is_final (s, l, r) = (s = 0)" |
|
112 |
|
113 lemma is_final_eq: |
|
114 shows "is_final (s, tp) = (s = 0)" |
|
115 by (case_tac tp) (auto) |
|
116 |
|
117 lemma after_is_final: |
|
118 assumes "is_final c" |
|
119 shows "is_final (steps c p n)" |
|
120 using assms |
|
121 apply(induct n) |
|
122 apply(case_tac [!] c) |
|
123 apply(auto) |
|
124 done |
|
125 |
|
126 lemma not_is_final: |
|
127 assumes a: "\<not> is_final (steps c p n1)" |
|
128 and b: "n2 \<le> n1" |
|
129 shows "\<not> is_final (steps c p n2)" |
|
130 proof (rule notI) |
|
131 obtain n3 where eq: "n1 = n2 + n3" using b by (metis le_iff_add) |
|
132 assume "is_final (steps c p n2)" |
|
133 then have "is_final (steps c p n1)" unfolding eq |
|
134 by (simp add: after_is_final) |
|
135 with a show "False" by simp |
|
136 qed |
|
137 |
|
138 (* if the machine is in the halting state, there must have |
|
139 been a state just before the halting state *) |
|
140 lemma before_final: |
|
141 assumes "steps0 (1, tp) A n = (0, tp')" |
|
142 shows "\<exists> n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" |
|
143 using assms |
|
144 proof(induct n arbitrary: tp') |
|
145 case (0 tp') |
|
146 have asm: "steps0 (1, tp) A 0 = (0, tp')" by fact |
|
147 then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" |
|
148 by simp |
|
149 next |
|
150 case (Suc n tp') |
|
151 have ih: "\<And>tp'. steps0 (1, tp) A n = (0, tp') \<Longrightarrow> |
|
152 \<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" by fact |
|
153 have asm: "steps0 (1, tp) A (Suc n) = (0, tp')" by fact |
|
154 obtain s l r where cases: "steps0 (1, tp) A n = (s, l, r)" |
|
155 by (auto intro: is_final.cases) |
|
156 then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" |
|
157 proof (cases "s = 0") |
|
158 case True (* in halting state *) |
|
159 then have "steps0 (1, tp) A n = (0, tp')" |
|
160 using asm cases by (simp del: steps.simps) |
|
161 then show ?thesis using ih by simp |
|
162 next |
|
163 case False (* not in halting state *) |
|
164 then have "\<not> is_final (steps0 (1, tp) A n) \<and> steps0 (1, tp) A (Suc n) = (0, tp')" |
|
165 using asm cases by simp |
|
166 then show ?thesis by auto |
|
167 qed |
|
168 qed |
|
169 |
|
170 (* well-formedness of Turing machine programs *) |
|
171 abbreviation "is_even n \<equiv> (n::nat) mod 2 = 0" |
|
172 |
|
173 fun |
|
174 tm_wf :: "tprog \<Rightarrow> bool" |
|
175 where |
|
176 "tm_wf (p, off) = (length p \<ge> 2 \<and> is_even (length p) \<and> |
|
177 (\<forall>(a, s) \<in> set p. s \<le> length p div 2 + off \<and> s \<ge> off))" |
|
178 |
|
179 abbreviation |
|
180 "tm_wf0 p \<equiv> tm_wf (p, 0)" |
|
181 |
|
182 abbreviation exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_ \<up> _" [100, 99] 100) |
|
183 where "x \<up> n == replicate n x" |
|
184 |
|
185 consts tape_of :: "'a \<Rightarrow> cell list" ("<_>" 100) |
|
186 |
|
187 defs (overloaded) |
|
188 tape_of_nat_abv: "<(n::nat)> \<equiv> Oc \<up> (Suc n)" |
|
189 |
|
190 fun tape_of_nat_list :: "'a list \<Rightarrow> cell list" |
|
191 where |
|
192 "tape_of_nat_list [] = []" | |
|
193 "tape_of_nat_list [n] = <n>" | |
|
194 "tape_of_nat_list (n#ns) = <n> @ Bk # (tape_of_nat_list ns)" |
|
195 |
|
196 fun tape_of_nat_pair :: "'a \<times> 'b \<Rightarrow> cell list" |
|
197 where |
|
198 "tape_of_nat_pair (n, m) = <n> @ [Bk] @ <m>" |
|
199 |
|
200 |
|
201 defs (overloaded) |
|
202 tape_of_nl_abv: "<(ns::'a list)> \<equiv> tape_of_nat_list ns" |
|
203 tape_of_nat_pair: "<(np::'a\<times>'b)> \<equiv> tape_of_nat_pair np" |
|
204 |
|
205 fun |
|
206 shift :: "instr list \<Rightarrow> nat \<Rightarrow> instr list" |
|
207 where |
|
208 "shift p n = (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)" |
|
209 |
|
210 fun |
|
211 adjust :: "instr list \<Rightarrow> instr list" |
|
212 where |
|
213 "adjust p = map (\<lambda> (a, s). (a, if s = 0 then (Suc (length p div 2)) else s)) p" |
|
214 |
|
215 lemma length_shift [simp]: |
|
216 shows "length (shift p n) = length p" |
|
217 by simp |
|
218 |
|
219 lemma length_adjust [simp]: |
|
220 shows "length (adjust p) = length p" |
|
221 by (induct p) (auto) |
|
222 |
|
223 |
|
224 (* composition of two Turing machines *) |
|
225 fun |
|
226 tm_comp :: "instr list \<Rightarrow> instr list \<Rightarrow> instr list" ("_ |+| _" [0, 0] 100) |
|
227 where |
|
228 "tm_comp p1 p2 = ((adjust p1) @ (shift p2 (length p1 div 2)))" |
|
229 |
|
230 lemma tm_comp_length: |
|
231 shows "length (A |+| B) = length A + length B" |
|
232 by auto |
|
233 |
|
234 lemma tm_comp_wf[intro]: |
|
235 "\<lbrakk>tm_wf (A, 0); tm_wf (B, 0)\<rbrakk> \<Longrightarrow> tm_wf (A |+| B, 0)" |
|
236 by (auto simp: tm_wf.simps shift.simps adjust.simps tm_comp_length tm_comp.simps) |
|
237 |
|
238 |
|
239 lemma tm_comp_step: |
|
240 assumes unfinal: "\<not> is_final (step0 c A)" |
|
241 shows "step0 c (A |+| B) = step0 c A" |
|
242 proof - |
|
243 obtain s l r where eq: "c = (s, l, r)" by (metis is_final.cases) |
|
244 have "\<not> is_final (step0 (s, l, r) A)" using unfinal eq by simp |
|
245 then have "case (fetch A s (read r)) of (a, s) \<Rightarrow> s \<noteq> 0" |
|
246 by (auto simp add: is_final_eq) |
|
247 then have "fetch (A |+| B) s (read r) = fetch A s (read r)" |
|
248 apply(case_tac [!] "read r") |
|
249 apply(case_tac [!] s) |
|
250 apply(auto simp: tm_comp_length nth_append) |
|
251 done |
|
252 then show "step0 c (A |+| B) = step0 c A" by (simp add: eq) |
|
253 qed |
|
254 |
|
255 lemma tm_comp_steps: |
|
256 assumes "\<not> is_final (steps0 c A n)" |
|
257 shows "steps0 c (A |+| B) n = steps0 c A n" |
|
258 using assms |
|
259 proof(induct n) |
|
260 case 0 |
|
261 then show "steps0 c (A |+| B) 0 = steps0 c A 0" by auto |
|
262 next |
|
263 case (Suc n) |
|
264 have ih: "\<not> is_final (steps0 c A n) \<Longrightarrow> steps0 c (A |+| B) n = steps0 c A n" by fact |
|
265 have fin: "\<not> is_final (steps0 c A (Suc n))" by fact |
|
266 then have fin1: "\<not> is_final (step0 (steps0 c A n) A)" |
|
267 by (auto simp only: step_red) |
|
268 then have fin2: "\<not> is_final (steps0 c A n)" |
|
269 by (metis is_final_eq step_0 surj_pair) |
|
270 |
|
271 have "steps0 c (A |+| B) (Suc n) = step0 (steps0 c (A |+| B) n) (A |+| B)" |
|
272 by (simp only: step_red) |
|
273 also have "... = step0 (steps0 c A n) (A |+| B)" by (simp only: ih[OF fin2]) |
|
274 also have "... = step0 (steps0 c A n) A" by (simp only: tm_comp_step[OF fin1]) |
|
275 finally show "steps0 c (A |+| B) (Suc n) = steps0 c A (Suc n)" |
|
276 by (simp only: step_red) |
|
277 qed |
|
278 |
|
279 lemma tm_comp_fetch_in_A: |
|
280 assumes h1: "fetch A s x = (a, 0)" |
|
281 and h2: "s \<le> length A div 2" |
|
282 and h3: "s \<noteq> 0" |
|
283 shows "fetch (A |+| B) s x = (a, Suc (length A div 2))" |
|
284 using h1 h2 h3 |
|
285 apply(case_tac s) |
|
286 apply(case_tac [!] x) |
|
287 apply(auto simp: tm_comp_length nth_append) |
|
288 done |
|
289 |
|
290 lemma tm_comp_exec_after_first: |
|
291 assumes h1: "\<not> is_final c" |
|
292 and h2: "step0 c A = (0, tp)" |
|
293 and h3: "fst c \<le> length A div 2" |
|
294 shows "step0 c (A |+| B) = (Suc (length A div 2), tp)" |
|
295 using h1 h2 h3 |
|
296 apply(case_tac c) |
|
297 apply(auto simp del: tm_comp.simps) |
|
298 apply(case_tac "fetch A a Bk") |
|
299 apply(simp del: tm_comp.simps) |
|
300 apply(subst tm_comp_fetch_in_A) |
|
301 apply(auto)[4] |
|
302 apply(case_tac "fetch A a (hd c)") |
|
303 apply(simp del: tm_comp.simps) |
|
304 apply(subst tm_comp_fetch_in_A) |
|
305 apply(auto)[4] |
|
306 done |
|
307 |
|
308 lemma step_in_range: |
|
309 assumes h1: "\<not> is_final (step0 c A)" |
|
310 and h2: "tm_wf (A, 0)" |
|
311 shows "fst (step0 c A) \<le> length A div 2" |
|
312 using h1 h2 |
|
313 apply(case_tac c) |
|
314 apply(case_tac a) |
|
315 apply(auto simp add: prod_case_unfold Let_def) |
|
316 apply(case_tac "hd c") |
|
317 apply(auto simp add: prod_case_unfold) |
|
318 done |
|
319 |
|
320 lemma steps_in_range: |
|
321 assumes h1: "\<not> is_final (steps0 (1, tp) A stp)" |
|
322 and h2: "tm_wf (A, 0)" |
|
323 shows "fst (steps0 (1, tp) A stp) \<le> length A div 2" |
|
324 using h1 |
|
325 proof(induct stp) |
|
326 case 0 |
|
327 then show "fst (steps0 (1, tp) A 0) \<le> length A div 2" using h2 |
|
328 by (auto simp add: steps.simps tm_wf.simps) |
|
329 next |
|
330 case (Suc stp) |
|
331 have ih: "\<not> is_final (steps0 (1, tp) A stp) \<Longrightarrow> fst (steps0 (1, tp) A stp) \<le> length A div 2" by fact |
|
332 have h: "\<not> is_final (steps0 (1, tp) A (Suc stp))" by fact |
|
333 from ih h h2 show "fst (steps0 (1, tp) A (Suc stp)) \<le> length A div 2" |
|
334 by (metis step_in_range step_red) |
|
335 qed |
|
336 |
|
337 lemma tm_comp_pre_halt_same: |
|
338 assumes a_ht: "steps0 (1, tp) A n = (0, tp')" |
|
339 and a_wf: "tm_wf (A, 0)" |
|
340 obtains n' where "steps0 (1, tp) (A |+| B) n' = (Suc (length A div 2), tp')" |
|
341 proof - |
|
342 assume a: "\<And>n. steps (1, tp) (A |+| B, 0) n = (Suc (length A div 2), tp') \<Longrightarrow> thesis" |
|
343 obtain stp' where fin: "\<not> is_final (steps0 (1, tp) A stp')" and h: "steps0 (1, tp) A (Suc stp') = (0, tp')" |
|
344 using before_final[OF a_ht] by blast |
|
345 from fin have h1:"steps0 (1, tp) (A |+| B) stp' = steps0 (1, tp) A stp'" |
|
346 by (rule tm_comp_steps) |
|
347 from h have h2: "step0 (steps0 (1, tp) A stp') A = (0, tp')" |
|
348 by (simp only: step_red) |
|
349 |
|
350 have "steps0 (1, tp) (A |+| B) (Suc stp') = step0 (steps0 (1, tp) (A |+| B) stp') (A |+| B)" |
|
351 by (simp only: step_red) |
|
352 also have "... = step0 (steps0 (1, tp) A stp') (A |+| B)" using h1 by simp |
|
353 also have "... = (Suc (length A div 2), tp')" |
|
354 by (rule tm_comp_exec_after_first[OF fin h2 steps_in_range[OF fin a_wf]]) |
|
355 finally show thesis using a by blast |
|
356 qed |
|
357 |
|
358 lemma tm_comp_fetch_second_zero: |
|
359 assumes h1: "fetch B s x = (a, 0)" |
|
360 and hs: "tm_wf (A, 0)" "s \<noteq> 0" |
|
361 shows "fetch (A |+| B) (s + (length A div 2)) x = (a, 0)" |
|
362 using h1 hs |
|
363 apply(case_tac x) |
|
364 apply(case_tac [!] s) |
|
365 apply(auto simp: tm_comp_length nth_append) |
|
366 done |
|
367 |
|
368 lemma tm_comp_fetch_second_inst: |
|
369 assumes h1: "fetch B sa x = (a, s)" |
|
370 and hs: "tm_wf (A, 0)" "sa \<noteq> 0" "s \<noteq> 0" |
|
371 shows "fetch (A |+| B) (sa + length A div 2) x = (a, s + length A div 2)" |
|
372 using h1 hs |
|
373 apply(case_tac x) |
|
374 apply(case_tac [!] sa) |
|
375 apply(auto simp: tm_comp_length nth_append) |
|
376 done |
|
377 |
|
378 |
|
379 lemma tm_comp_second_same: |
|
380 assumes a_wf: "tm_wf (A, 0)" |
|
381 and steps: "steps0 (1, l, r) B stp = (s', l', r')" |
|
382 shows "steps0 (Suc (length A div 2), l, r) (A |+| B) stp |
|
383 = (if s' = 0 then 0 else s' + length A div 2, l', r')" |
|
384 using steps |
|
385 proof(induct stp arbitrary: s' l' r') |
|
386 case 0 |
|
387 then show ?case by (simp add: steps.simps) |
|
388 next |
|
389 case (Suc stp s' l' r') |
|
390 obtain s'' l'' r'' where a: "steps0 (1, l, r) B stp = (s'', l'', r'')" |
|
391 by (metis is_final.cases) |
|
392 then have ih1: "s'' = 0 \<Longrightarrow> steps0 (Suc (length A div 2), l, r) (A |+| B) stp = (0, l'', r'')" |
|
393 and ih2: "s'' \<noteq> 0 \<Longrightarrow> steps0 (Suc (length A div 2), l, r) (A |+| B) stp = (s'' + length A div 2, l'', r'')" |
|
394 using Suc by (auto) |
|
395 have h: "steps0 (1, l, r) B (Suc stp) = (s', l', r')" by fact |
|
396 |
|
397 { assume "s'' = 0" |
|
398 then have ?case using a h ih1 by (simp del: steps.simps) |
|
399 } moreover |
|
400 { assume as: "s'' \<noteq> 0" "s' = 0" |
|
401 from as a h |
|
402 have "step0 (s'', l'', r'') B = (0, l', r')" by (simp del: steps.simps) |
|
403 with as have ?case |
|
404 apply(simp add: ih2[OF as(1)] step.simps del: tm_comp.simps steps.simps) |
|
405 apply(case_tac "fetch B s'' (read r'')") |
|
406 apply(auto simp add: tm_comp_fetch_second_zero[OF _ a_wf] simp del: tm_comp.simps) |
|
407 done |
|
408 } moreover |
|
409 { assume as: "s'' \<noteq> 0" "s' \<noteq> 0" |
|
410 from as a h |
|
411 have "step0 (s'', l'', r'') B = (s', l', r')" by (simp del: steps.simps) |
|
412 with as have ?case |
|
413 apply(simp add: ih2[OF as(1)] step.simps del: tm_comp.simps steps.simps) |
|
414 apply(case_tac "fetch B s'' (read r'')") |
|
415 apply(auto simp add: tm_comp_fetch_second_inst[OF _ a_wf as] simp del: tm_comp.simps) |
|
416 done |
|
417 } |
|
418 ultimately show ?case by blast |
|
419 qed |
|
420 |
|
421 lemma tm_comp_second_halt_same: |
|
422 assumes "tm_wf (A, 0)" |
|
423 and "steps0 (1, l, r) B stp = (0, l', r')" |
|
424 shows "steps0 (Suc (length A div 2), l, r) (A |+| B) stp = (0, l', r')" |
|
425 using tm_comp_second_same[OF assms] by (simp) |
|
426 |
|
427 end |
|
428 |