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