|
1 (* Title: Turing machines |
|
2 Author: Xu Jian <xujian817@hotmail.com> |
|
3 Maintainer: Xu Jian |
|
4 *) |
|
5 |
|
6 theory turing_basic |
|
7 imports Main |
|
8 begin |
|
9 |
|
10 section {* Basic definitions of Turing machine *} |
|
11 |
|
12 definition |
|
13 "iseven n \<equiv> \<exists>x. n = 2 * x" |
|
14 |
|
15 datatype action = W0 | W1 | L | R | Nop |
|
16 |
|
17 datatype cell = Bk | Oc |
|
18 |
|
19 type_synonym tape = "cell list \<times> cell list" |
|
20 |
|
21 type_synonym state = nat |
|
22 |
|
23 type_synonym instr = "action \<times> state" |
|
24 |
|
25 type_synonym tprog = "instr list" |
|
26 |
|
27 type_synonym config = "state \<times> tape" |
|
28 |
|
29 fun nth_of where |
|
30 "nth_of [] _ = None" |
|
31 | "nth_of (x # xs) 0 = Some x" |
|
32 | "nth_of (x # xs) (Suc n) = nth_of xs n" |
|
33 |
|
34 fun |
|
35 fetch :: "tprog \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr" |
|
36 where |
|
37 "fetch p 0 b = (Nop, 0)" |
|
38 | "fetch p (Suc s) Bk = |
|
39 (case nth_of p (2 * s) of |
|
40 Some i \<Rightarrow> i |
|
41 | None \<Rightarrow> (Nop, 0))" |
|
42 |"fetch p (Suc s) Oc = |
|
43 (case nth_of p ((2 * s) + 1) of |
|
44 Some i \<Rightarrow> i |
|
45 | None \<Rightarrow> (Nop, 0))" |
|
46 |
|
47 fun |
|
48 update :: "action \<Rightarrow> tape \<Rightarrow> tape" |
|
49 where |
|
50 "update W0 (l, r) = (l, Bk # (tl r))" |
|
51 | "update W1 (l, r) = (l, Oc # (tl r))" |
|
52 | "update L (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" |
|
53 | "update R (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" |
|
54 | "update Nop (l, r) = (l, r)" |
|
55 |
|
56 abbreviation |
|
57 "read r == if (r = []) then Bk else hd r" |
|
58 |
|
59 |
|
60 fun |
|
61 step :: "config \<Rightarrow> tprog \<Rightarrow> config" |
|
62 where |
|
63 "step (s, l, r) p = (let (a, s') = fetch p s (read r) in (s', update a (l, r)))" |
|
64 |
|
65 fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config" |
|
66 where |
|
67 "steps c p 0 = c" | |
|
68 "steps c p (Suc n) = steps (step c p) p n" |
|
69 |
|
70 lemma step_red [simp]: |
|
71 shows "steps c p (Suc n) = step (steps c p n) p" |
|
72 by (induct n arbitrary: c) (auto) |
|
73 |
|
74 lemma steps_add [simp]: |
|
75 shows "steps c p (m + n) = steps (steps c p m) p n" |
|
76 by (induct m arbitrary: c) (auto) |
|
77 |
|
78 definition |
|
79 tm_wf :: "tprog \<Rightarrow> bool" |
|
80 where |
|
81 "tm_wf p = (length p \<ge> 2 \<and> iseven (length p) \<and> (\<forall>(a, s) \<in> set p. s \<le> length p div 2))" |
|
82 |
|
83 (* FIXME: needed? *) |
|
84 lemma halt_lemma: |
|
85 "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)" |
|
86 by (metis wf_iff_no_infinite_down_chain) |
|
87 |
|
88 abbreviation exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_ \<up> _" [100, 99] 100) |
|
89 where "x \<up> n == replicate n x" |
|
90 |
|
91 definition tinres :: "cell list \<Rightarrow> cell list \<Rightarrow> bool" |
|
92 where |
|
93 "tinres xs ys = (\<exists>n. xs = ys @ Bk \<up> n \<or> ys = xs @ Bk \<up> n)" |
|
94 |
|
95 fun |
|
96 shift :: "tprog \<Rightarrow> nat \<Rightarrow> tprog" |
|
97 where |
|
98 "shift p n = (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)" |
|
99 |
|
100 |
|
101 lemma [simp]: |
|
102 "length (shift p n) = length p" |
|
103 by (simp) |
|
104 |
|
105 fun |
|
106 adjust :: "tprog \<Rightarrow> tprog" |
|
107 where |
|
108 "adjust p = (map (\<lambda> (a, s). (a, if s = 0 then ((length p) div 2) + 1 else s)) p)" |
|
109 |
|
110 lemma [simp]: |
|
111 shows "length (adjust p) = length p" |
|
112 by (simp) |
|
113 |
|
114 definition |
|
115 tm_comp :: "tprog \<Rightarrow> tprog \<Rightarrow> tprog" ("_ |+| _" [0, 0] 100) |
|
116 where |
|
117 "tm_comp p1 p2 = ((adjust p1) @ (shift p2 ((length p1) div 2)))" |
|
118 |
|
119 fun |
|
120 is_final :: "config \<Rightarrow> bool" |
|
121 where |
|
122 "is_final (s, l, r) = (s = 0)" |
|
123 |
|
124 fun |
|
125 holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100) |
|
126 where |
|
127 "P holds_for (s, l, r) = P (l, r)" |
|
128 |
|
129 lemma step_0 [simp]: |
|
130 shows "step (0, (l, r)) p = (0, (l, r))" |
|
131 by simp |
|
132 |
|
133 lemma steps_0 [simp]: |
|
134 shows "steps (0, (l, r)) p n = (0, (l, r))" |
|
135 by (induct n) (simp_all) |
|
136 |
|
137 type_synonym assert = "tape \<Rightarrow> bool" |
|
138 |
|
139 definition assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100) |
|
140 where |
|
141 "assert_imp P Q = (\<forall>l r. P (l, r) \<longrightarrow> Q (l, r))" |
|
142 |
|
143 definition |
|
144 Hoare :: "assert \<Rightarrow> tprog \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50) |
|
145 where |
|
146 "{P} p {Q} \<equiv> |
|
147 (\<forall>l r. P (l, r) \<longrightarrow> (\<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)))" |
|
148 |
|
149 lemma HoareI: |
|
150 assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)" |
|
151 shows "{P} p {Q}" |
|
152 unfolding Hoare_def using assms by auto |
|
153 |
|
154 lemma HoareI2: |
|
155 assumes "\<And>l r n. P (l, r) \<and> is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)" |
|
156 shows "{P} p {Q}" |
|
157 unfolding Hoare_def using assms by auto |
|
158 |
|
159 text {* |
|
160 {P1} A {Q1} {P2} B {Q2} Q1 \<mapsto> P2 |
|
161 ----------------------------------- |
|
162 {P1} A |+| B {Q2} |
|
163 *} |
|
164 |
|
165 |
|
166 lemma Hoare_plus: |
|
167 assumes aimpb: "Q1 \<mapsto> P2" |
|
168 and A_wf : "tm_wf A" |
|
169 and B_wf : "tm_wf B" |
|
170 and A_halt : "{P1} A {Q1}" |
|
171 and B_halt : "{P2} B {Q2}" |
|
172 shows "{P1} A |+| B {Q2}" |
|
173 proof(rule HoareI) |
|
174 fix a b |
|
175 assume h: "P1 (a, b)" |
|
176 hence "\<exists>n. let (s, tp') = steps (Suc 0, a, b) A n in s = 0 \<and> Q1 tp'" |
|
177 using A_halt unfolding hoare_def by simp |
|
178 from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \<and> Q1 tp'" .. |
|
179 then show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'" |
|
180 proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE) |
|
181 fix aa ba c |
|
182 assume g1: "Q1 (ba, c)" |
|
183 and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)" |
|
184 hence "P2 (ba, c)" |
|
185 using aimpb apply(simp add: assert_imp_def) |
|
186 done |
|
187 hence "\<exists> stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \<and> Q2 tp'" |
|
188 using B_halt unfolding hoare_def by simp |
|
189 from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \<and> Q2 tp'" .. |
|
190 thus "?thesis" |
|
191 proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE) |
|
192 fix aa bb ca |
|
193 assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)" |
|
194 have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)" |
|
195 using g2 A_wf B_wf |
|
196 sorry |
|
197 moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" |
|
198 using g3 A_wf B_wf |
|
199 sorry |
|
200 ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'" |
|
201 apply(erule_tac exE, erule_tac exE) |
|
202 apply(rule_tac x = "stp + stpa" in exI, simp) |
|
203 using g3 by simp |
|
204 qed |
|
205 qed |
|
206 qed |
|
207 |
|
208 lemma hoare_plus: |
|
209 assumes A_wf : "tm_wf A" |
|
210 and B_wf : "tm_wf B" |
|
211 and A_halt : "{P1} A {Q1}" |
|
212 and B_halt : "{P2} B {Q2}" |
|
213 and aimpb: "Q1 \<mapsto> P2" |
|
214 shows "{P1} A |+| B {Q2}" |
|
215 unfolding hoare_def |
|
216 proof(auto split: ) |
|
217 fix a b |
|
218 assume h: "P1 (a, b)" |
|
219 hence "\<exists>n. let (s, tp') = steps (Suc 0, a, b) A n in s = 0 \<and> Q1 tp'" |
|
220 using A_halt unfolding hoare_def by simp |
|
221 from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \<and> Q1 tp'" .. |
|
222 then show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'" |
|
223 proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE) |
|
224 fix aa ba c |
|
225 assume g1: "Q1 (ba, c)" |
|
226 and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)" |
|
227 hence "P2 (ba, c)" |
|
228 using aimpb apply(simp add: assert_imp_def) |
|
229 done |
|
230 hence "\<exists> stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \<and> Q2 tp'" |
|
231 using B_halt unfolding hoare_def by simp |
|
232 from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \<and> Q2 tp'" .. |
|
233 thus "?thesis" |
|
234 proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE) |
|
235 fix aa bb ca |
|
236 assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)" |
|
237 have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)" |
|
238 using g2 A_wf B_wf |
|
239 sorry |
|
240 moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" |
|
241 using g3 A_wf B_wf |
|
242 sorry |
|
243 ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'" |
|
244 apply(erule_tac exE, erule_tac exE) |
|
245 apply(rule_tac x = "stp + stpa" in exI, simp) |
|
246 using g3 by simp |
|
247 qed |
|
248 qed |
|
249 qed |
|
250 |
|
251 |
|
252 |
|
253 locale turing_merge = |
|
254 fixes A :: "tprog" and B :: "tprog" and P1 :: "assert" |
|
255 and P2 :: "assert" |
|
256 and P3 :: "assert" |
|
257 and P4 :: "assert" |
|
258 and Q1:: "assert" |
|
259 and Q2 :: "assert" |
|
260 assumes |
|
261 A_wf : "tm_wf A" |
|
262 and B_wf : "tm_wf B" |
|
263 and A_halt : "P1 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'" |
|
264 and B_halt : "P2 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) B stp in s = 0 \<and> Q2 tp'" |
|
265 and A_uhalt : "P3 tp \<Longrightarrow> \<not> (\<exists> stp. is_final (steps (Suc 0, tp) A stp))" |
|
266 and B_uhalt: "P4 tp \<Longrightarrow> \<not> (\<exists> stp. is_final (steps (Suc 0, tp) B stp))" |
|
267 begin |
|
268 |
|
269 |
|
270 text {* |
|
271 The following lemma tries to derive the Hoare logic rule for sequentially combined TMs. |
|
272 It deals with the situtation when both @{text "A"} and @{text "B"} are terminated. |
|
273 *} |
|
274 |
|
275 |
|
276 |
|
277 lemma t_merge_uhalt_tmp: |
|
278 assumes B_uh: "\<forall>stp. \<not> is_final (steps (Suc 0, b, c) B stp)" |
|
279 and merge_ah: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" |
|
280 shows "\<forall> stp. \<not> is_final (steps (Suc 0, tp) (A |+| B) stp)" |
|
281 using B_uh merge_ah |
|
282 apply(rule_tac allI) |
|
283 apply(case_tac "stp > stpa") |
|
284 apply(erule_tac x = "stp - stpa" in allE) |
|
285 apply(case_tac "(steps (Suc 0, b, c) B (stp - stpa))", simp) |
|
286 proof - |
|
287 fix stp a ba ca |
|
288 assume h1: "\<not> is_final (a, ba, ca)" "stpa < stp" |
|
289 and h2: "steps (Suc 0, b, c) B (stp - stpa) = (a, ba, ca)" |
|
290 have "steps (Suc 0 + length A div 2, b, c) (A |+| B) (stp - stpa) = |
|
291 (if a = 0 then 0 else a + length A div 2, ba, ca)" |
|
292 using A_wf B_wf h2 |
|
293 by(rule_tac t_merge_snd_eq_steps, auto) |
|
294 moreover have "a > 0" using h1 by(simp add: is_final_def) |
|
295 moreover have "\<exists> stpb. stp = stpa + stpb" |
|
296 using h1 by(rule_tac x = "stp - stpa" in exI, simp) |
|
297 ultimately show "\<not> is_final (steps (Suc 0, tp) (A |+| B) stp)" |
|
298 using merge_ah |
|
299 by(auto simp: steps_add is_final_def) |
|
300 next |
|
301 fix stp |
|
302 assume h: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" "\<not> stpa < stp" |
|
303 hence "\<exists> stpb. stpa = stp + stpb" apply(rule_tac x = "stpa - stp" in exI, auto) done |
|
304 thus "\<not> is_final (steps (Suc 0, tp) (A |+| B) stp)" |
|
305 using h |
|
306 apply(auto) |
|
307 apply(cases "steps (Suc 0, tp) (A |+| B) stp", simp add: steps_add is_final_def steps_0) |
|
308 done |
|
309 qed |
|
310 |
|
311 text {* |
|
312 The following lemma deals with the situation when TM @{text "B"} can not terminate. |
|
313 *} |
|
314 |
|
315 lemma t_merge_uhalt: |
|
316 assumes aimpb: "Q1 \<mapsto> P4" |
|
317 shows "P1 \<mapsto> \<lambda> tp. \<not> (\<exists> stp. is_final (steps (Suc 0, tp) (A |+| B) stp))" |
|
318 proof(simp only: assert_imp_def, rule_tac allI, rule_tac impI) |
|
319 fix tp |
|
320 assume init_asst: "P1 tp" |
|
321 show "\<not> (\<exists>stp. is_final (steps (Suc 0, tp) (A |+| B) stp))" |
|
322 proof - |
|
323 have "\<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'" |
|
324 using A_halt[of tp] init_asst |
|
325 by(simp) |
|
326 from this obtain stpx where "let (s, tp') = steps (Suc 0, tp) A stpx in s = 0 \<and> Q1 tp'" .. |
|
327 thus "?thesis" |
|
328 proof(cases "steps (Suc 0, tp) A stpx", simp, erule_tac conjE) |
|
329 fix a b c |
|
330 assume "Q1 (b, c)" |
|
331 and h3: "steps (Suc 0, tp) A stpx = (0, b, c)" |
|
332 hence h2: "P4 (b, c)" using aimpb |
|
333 by(simp add: assert_imp_def) |
|
334 have "\<exists> stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), b, c)" |
|
335 using h3 A_wf B_wf |
|
336 apply(rule_tac stp = stpx in t_merge_pre_halt_same, auto) |
|
337 done |
|
338 from this obtain stpa where h4:"steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" .. |
|
339 have " \<not> (\<exists> stp. is_final (steps (Suc 0, b, c) B stp))" |
|
340 using B_uhalt [of "(b, c)"] h2 apply simp |
|
341 done |
|
342 from this and h4 show "\<forall>stp. \<not> is_final (steps (Suc 0, tp) (A |+| B) stp)" |
|
343 by(rule_tac t_merge_uhalt_tmp, auto) |
|
344 qed |
|
345 qed |
|
346 qed |
|
347 end |
|
348 |
|
349 end |
|
350 |