|
1 theory UF |
|
2 imports Main rec_def turing_basic GCD abacus |
|
3 begin |
|
4 |
|
5 text {* |
|
6 This theory file constructs the Universal Function @{text "rec_F"}, which is the UTM defined |
|
7 in terms of recursive functions. This @{text "rec_F"} is essentially an |
|
8 interpreter of Turing Machines. Once the correctness of @{text "rec_F"} is established, |
|
9 UTM can easil be obtained by compling @{text "rec_F"} into the corresponding Turing Machine. |
|
10 *} |
|
11 |
|
12 section {* Univeral Function *} |
|
13 |
|
14 subsection {* The construction of component functions *} |
|
15 |
|
16 text {* |
|
17 This section constructs a set of component functions used to construct @{text "rec_F"}. |
|
18 *} |
|
19 |
|
20 text {* |
|
21 The recursive function used to do arithmatic addition. |
|
22 *} |
|
23 definition rec_add :: "recf" |
|
24 where |
|
25 "rec_add \<equiv> Pr 1 (id 1 0) (Cn 3 s [id 3 2])" |
|
26 |
|
27 text {* |
|
28 The recursive function used to do arithmatic multiplication. |
|
29 *} |
|
30 definition rec_mult :: "recf" |
|
31 where |
|
32 "rec_mult = Pr 1 z (Cn 3 rec_add [id 3 0, id 3 2])" |
|
33 |
|
34 text {* |
|
35 The recursive function used to do arithmatic precede. |
|
36 *} |
|
37 definition rec_pred :: "recf" |
|
38 where |
|
39 "rec_pred = Cn 1 (Pr 1 z (id 3 1)) [id 1 0, id 1 0]" |
|
40 |
|
41 text {* |
|
42 The recursive function used to do arithmatic subtraction. |
|
43 *} |
|
44 definition rec_minus :: "recf" |
|
45 where |
|
46 "rec_minus = Pr 1 (id 1 0) (Cn 3 rec_pred [id 3 2])" |
|
47 |
|
48 text {* |
|
49 @{text "constn n"} is the recursive function which computes |
|
50 nature number @{text "n"}. |
|
51 *} |
|
52 fun constn :: "nat \<Rightarrow> recf" |
|
53 where |
|
54 "constn 0 = z" | |
|
55 "constn (Suc n) = Cn 1 s [constn n]" |
|
56 |
|
57 |
|
58 text {* |
|
59 Signal function, which returns 1 when the input argument is greater than @{text "0"}. |
|
60 *} |
|
61 definition rec_sg :: "recf" |
|
62 where |
|
63 "rec_sg = Cn 1 rec_minus [constn 1, |
|
64 Cn 1 rec_minus [constn 1, id 1 0]]" |
|
65 |
|
66 text {* |
|
67 @{text "rec_less"} compares its two arguments, returns @{text "1"} if |
|
68 the first is less than the second; otherwise returns @{text "0"}. |
|
69 *} |
|
70 definition rec_less :: "recf" |
|
71 where |
|
72 "rec_less = Cn 2 rec_sg [Cn 2 rec_minus [id 2 1, id 2 0]]" |
|
73 |
|
74 text {* |
|
75 @{text "rec_not"} inverse its argument: returns @{text "1"} when the |
|
76 argument is @{text "0"}; returns @{text "0"} otherwise. |
|
77 *} |
|
78 definition rec_not :: "recf" |
|
79 where |
|
80 "rec_not = Cn 1 rec_minus [constn 1, id 1 0]" |
|
81 |
|
82 text {* |
|
83 @{text "rec_eq"} compares its two arguments: returns @{text "1"} |
|
84 if they are equal; return @{text "0"} otherwise. |
|
85 *} |
|
86 definition rec_eq :: "recf" |
|
87 where |
|
88 "rec_eq = Cn 2 rec_minus [Cn 2 (constn 1) [id 2 0], |
|
89 Cn 2 rec_add [Cn 2 rec_minus [id 2 0, id 2 1], |
|
90 Cn 2 rec_minus [id 2 1, id 2 0]]]" |
|
91 |
|
92 text {* |
|
93 @{text "rec_conj"} computes the conjunction of its two arguments, |
|
94 returns @{text "1"} if both of them are non-zero; returns @{text "0"} |
|
95 otherwise. |
|
96 *} |
|
97 definition rec_conj :: "recf" |
|
98 where |
|
99 "rec_conj = Cn 2 rec_sg [Cn 2 rec_mult [id 2 0, id 2 1]] " |
|
100 |
|
101 text {* |
|
102 @{text "rec_disj"} computes the disjunction of its two arguments, |
|
103 returns @{text "0"} if both of them are zero; returns @{text "0"} |
|
104 otherwise. |
|
105 *} |
|
106 definition rec_disj :: "recf" |
|
107 where |
|
108 "rec_disj = Cn 2 rec_sg [Cn 2 rec_add [id 2 0, id 2 1]]" |
|
109 |
|
110 |
|
111 text {* |
|
112 Computes the arity of recursive function. |
|
113 *} |
|
114 |
|
115 fun arity :: "recf \<Rightarrow> nat" |
|
116 where |
|
117 "arity z = 1" |
|
118 | "arity s = 1" |
|
119 | "arity (id m n) = m" |
|
120 | "arity (Cn n f gs) = n" |
|
121 | "arity (Pr n f g) = Suc n" |
|
122 | "arity (Mn n f) = n" |
|
123 |
|
124 text {* |
|
125 @{text "get_fstn_args n (Suc k)"} returns |
|
126 @{text "[id n 0, id n 1, id n 2, \<dots>, id n k]"}, |
|
127 the effect of which is to take out the first @{text "Suc k"} |
|
128 arguments out of the @{text "n"} input arguments. |
|
129 *} |
|
130 |
|
131 fun get_fstn_args :: "nat \<Rightarrow> nat \<Rightarrow> recf list" |
|
132 where |
|
133 "get_fstn_args n 0 = []" |
|
134 | "get_fstn_args n (Suc y) = get_fstn_args n y @ [id n y]" |
|
135 |
|
136 text {* |
|
137 @{text "rec_sigma f"} returns the recursive functions which |
|
138 sums up the results of @{text "f"}: |
|
139 \[ |
|
140 (rec\_sigma f)(x, y) = f(x, 0) + f(x, 1) + \cdots + f(x, y) |
|
141 \] |
|
142 *} |
|
143 fun rec_sigma :: "recf \<Rightarrow> recf" |
|
144 where |
|
145 "rec_sigma rf = |
|
146 (let vl = arity rf in |
|
147 Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ |
|
148 [Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) |
|
149 (Cn (Suc vl) rec_add [id (Suc vl) vl, |
|
150 Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) |
|
151 @ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" |
|
152 |
|
153 text {* |
|
154 @{text "rec_exec"} is the interpreter function for |
|
155 reursive functions. The function is defined such that |
|
156 it always returns meaningful results for primitive recursive |
|
157 functions. |
|
158 *} |
|
159 function rec_exec :: "recf \<Rightarrow> nat list \<Rightarrow> nat" |
|
160 where |
|
161 "rec_exec z xs = 0" | |
|
162 "rec_exec s xs = (Suc (xs ! 0))" | |
|
163 "rec_exec (id m n) xs = (xs ! n)" | |
|
164 "rec_exec (Cn n f gs) xs = |
|
165 (let ys = (map (\<lambda> a. rec_exec a xs) gs) in |
|
166 rec_exec f ys)" | |
|
167 "rec_exec (Pr n f g) xs = |
|
168 (if last xs = 0 then |
|
169 rec_exec f (butlast xs) |
|
170 else rec_exec g (butlast xs @ [last xs - 1] @ |
|
171 [rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]))" | |
|
172 "rec_exec (Mn n f) xs = (LEAST x. rec_exec f (xs @ [x]) = 0)" |
|
173 by pat_completeness auto |
|
174 termination |
|
175 proof |
|
176 show "wf (measures [\<lambda> (r, xs). size r, (\<lambda> (r, xs). last xs)])" |
|
177 by auto |
|
178 next |
|
179 fix n f gs xs x |
|
180 assume "(x::recf) \<in> set gs" |
|
181 thus "((x, xs), Cn n f gs, xs) \<in> |
|
182 measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
|
183 by(induct gs, auto) |
|
184 next |
|
185 fix n f gs xs x |
|
186 assume "x = map (\<lambda>a. rec_exec a xs) gs" |
|
187 "\<And>x. x \<in> set gs \<Longrightarrow> rec_exec_dom (x, xs)" |
|
188 thus "((f, x), Cn n f gs, xs) \<in> |
|
189 measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
|
190 by(auto) |
|
191 next |
|
192 fix n f g xs |
|
193 show "((f, butlast xs), Pr n f g, xs) \<in> |
|
194 measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
|
195 by auto |
|
196 next |
|
197 fix n f g xs |
|
198 assume "last xs \<noteq> (0::nat)" thus |
|
199 "((Pr n f g, butlast xs @ [last xs - 1]), Pr n f g, xs) |
|
200 \<in> measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
|
201 by auto |
|
202 next |
|
203 fix n f g xs |
|
204 show "((g, butlast xs @ [last xs - 1] @ [rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]), |
|
205 Pr n f g, xs) \<in> measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
|
206 by auto |
|
207 next |
|
208 fix n f xs x |
|
209 show "((f, xs @ [x]), Mn n f, xs) \<in> |
|
210 measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
|
211 by auto |
|
212 qed |
|
213 |
|
214 declare rec_exec.simps[simp del] constn.simps[simp del] |
|
215 |
|
216 text {* |
|
217 Correctness of @{text "rec_add"}. |
|
218 *} |
|
219 lemma add_lemma: "\<And> x y. rec_exec rec_add [x, y] = x + y" |
|
220 by(induct_tac y, auto simp: rec_add_def rec_exec.simps) |
|
221 |
|
222 text {* |
|
223 Correctness of @{text "rec_mult"}. |
|
224 *} |
|
225 lemma mult_lemma: "\<And> x y. rec_exec rec_mult [x, y] = x * y" |
|
226 by(induct_tac y, auto simp: rec_mult_def rec_exec.simps add_lemma) |
|
227 |
|
228 text {* |
|
229 Correctness of @{text "rec_pred"}. |
|
230 *} |
|
231 lemma pred_lemma: "\<And> x. rec_exec rec_pred [x] = x - 1" |
|
232 by(induct_tac x, auto simp: rec_pred_def rec_exec.simps) |
|
233 |
|
234 text {* |
|
235 Correctness of @{text "rec_minus"}. |
|
236 *} |
|
237 lemma minus_lemma: "\<And> x y. rec_exec rec_minus [x, y] = x - y" |
|
238 by(induct_tac y, auto simp: rec_exec.simps rec_minus_def pred_lemma) |
|
239 |
|
240 text {* |
|
241 Correctness of @{text "rec_sg"}. |
|
242 *} |
|
243 lemma sg_lemma: "\<And> x. rec_exec rec_sg [x] = (if x = 0 then 0 else 1)" |
|
244 by(auto simp: rec_sg_def minus_lemma rec_exec.simps constn.simps) |
|
245 |
|
246 text {* |
|
247 Correctness of @{text "constn"}. |
|
248 *} |
|
249 lemma constn_lemma: "rec_exec (constn n) [x] = n" |
|
250 by(induct n, auto simp: rec_exec.simps constn.simps) |
|
251 |
|
252 text {* |
|
253 Correctness of @{text "rec_less"}. |
|
254 *} |
|
255 lemma less_lemma: "\<And> x y. rec_exec rec_less [x, y] = |
|
256 (if x < y then 1 else 0)" |
|
257 by(induct_tac y, auto simp: rec_exec.simps |
|
258 rec_less_def minus_lemma sg_lemma) |
|
259 |
|
260 text {* |
|
261 Correctness of @{text "rec_not"}. |
|
262 *} |
|
263 lemma not_lemma: |
|
264 "\<And> x. rec_exec rec_not [x] = (if x = 0 then 1 else 0)" |
|
265 by(induct_tac x, auto simp: rec_exec.simps rec_not_def |
|
266 constn_lemma minus_lemma) |
|
267 |
|
268 text {* |
|
269 Correctness of @{text "rec_eq"}. |
|
270 *} |
|
271 lemma eq_lemma: "\<And> x y. rec_exec rec_eq [x, y] = (if x = y then 1 else 0)" |
|
272 by(induct_tac y, auto simp: rec_exec.simps rec_eq_def constn_lemma add_lemma minus_lemma) |
|
273 |
|
274 text {* |
|
275 Correctness of @{text "rec_conj"}. |
|
276 *} |
|
277 lemma conj_lemma: "\<And> x y. rec_exec rec_conj [x, y] = (if x = 0 \<or> y = 0 then 0 |
|
278 else 1)" |
|
279 by(induct_tac y, auto simp: rec_exec.simps sg_lemma rec_conj_def mult_lemma) |
|
280 |
|
281 |
|
282 text {* |
|
283 Correctness of @{text "rec_disj"}. |
|
284 *} |
|
285 lemma disj_lemma: "\<And> x y. rec_exec rec_disj [x, y] = (if x = 0 \<and> y = 0 then 0 |
|
286 else 1)" |
|
287 by(induct_tac y, auto simp: rec_disj_def sg_lemma add_lemma rec_exec.simps) |
|
288 |
|
289 |
|
290 text {* |
|
291 @{text "primrec recf n"} is true iff |
|
292 @{text "recf"} is a primitive recursive function |
|
293 with arity @{text "n"}. |
|
294 *} |
|
295 inductive primerec :: "recf \<Rightarrow> nat \<Rightarrow> bool" |
|
296 where |
|
297 prime_z[intro]: "primerec z (Suc 0)" | |
|
298 prime_s[intro]: "primerec s (Suc 0)" | |
|
299 prime_id[intro!]: "\<lbrakk>n < m\<rbrakk> \<Longrightarrow> primerec (id m n) m" | |
|
300 prime_cn[intro!]: "\<lbrakk>primerec f k; length gs = k; |
|
301 \<forall> i < length gs. primerec (gs ! i) m; m = n\<rbrakk> |
|
302 \<Longrightarrow> primerec (Cn n f gs) m" | |
|
303 prime_pr[intro!]: "\<lbrakk>primerec f n; |
|
304 primerec g (Suc (Suc n)); m = Suc n\<rbrakk> |
|
305 \<Longrightarrow> primerec (Pr n f g) m" |
|
306 |
|
307 inductive_cases prime_cn_reverse'[elim]: "primerec (Cn n f gs) n" |
|
308 inductive_cases prime_mn_reverse: "primerec (Mn n f) m" |
|
309 inductive_cases prime_z_reverse[elim]: "primerec z n" |
|
310 inductive_cases prime_s_reverse[elim]: "primerec s n" |
|
311 inductive_cases prime_id_reverse[elim]: "primerec (id m n) k" |
|
312 inductive_cases prime_cn_reverse[elim]: "primerec (Cn n f gs) m" |
|
313 inductive_cases prime_pr_reverse[elim]: "primerec (Pr n f g) m" |
|
314 |
|
315 declare mult_lemma[simp] add_lemma[simp] pred_lemma[simp] |
|
316 minus_lemma[simp] sg_lemma[simp] constn_lemma[simp] |
|
317 less_lemma[simp] not_lemma[simp] eq_lemma[simp] |
|
318 conj_lemma[simp] disj_lemma[simp] |
|
319 |
|
320 text {* |
|
321 @{text "Sigma"} is the logical specification of |
|
322 the recursive function @{text "rec_sigma"}. |
|
323 *} |
|
324 function Sigma :: "(nat list \<Rightarrow> nat) \<Rightarrow> nat list \<Rightarrow> nat" |
|
325 where |
|
326 "Sigma g xs = (if last xs = 0 then g xs |
|
327 else (Sigma g (butlast xs @ [last xs - 1]) + |
|
328 g xs)) " |
|
329 by pat_completeness auto |
|
330 termination |
|
331 proof |
|
332 show "wf (measure (\<lambda> (f, xs). last xs))" by auto |
|
333 next |
|
334 fix g xs |
|
335 assume "last (xs::nat list) \<noteq> 0" |
|
336 thus "((g, butlast xs @ [last xs - 1]), g, xs) |
|
337 \<in> measure (\<lambda>(f, xs). last xs)" |
|
338 by auto |
|
339 qed |
|
340 |
|
341 declare rec_exec.simps[simp del] get_fstn_args.simps[simp del] |
|
342 arity.simps[simp del] Sigma.simps[simp del] |
|
343 rec_sigma.simps[simp del] |
|
344 |
|
345 lemma [simp]: "arity z = 1" |
|
346 by(simp add: arity.simps) |
|
347 |
|
348 lemma rec_pr_0_simp_rewrite: " |
|
349 rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs" |
|
350 by(simp add: rec_exec.simps) |
|
351 |
|
352 lemma rec_pr_0_simp_rewrite_single_param: " |
|
353 rec_exec (Pr n f g) [0] = rec_exec f []" |
|
354 by(simp add: rec_exec.simps) |
|
355 |
|
356 lemma rec_pr_Suc_simp_rewrite: |
|
357 "rec_exec (Pr n f g) (xs @ [Suc x]) = |
|
358 rec_exec g (xs @ [x] @ |
|
359 [rec_exec (Pr n f g) (xs @ [x])])" |
|
360 by(simp add: rec_exec.simps) |
|
361 |
|
362 lemma rec_pr_Suc_simp_rewrite_single_param: |
|
363 "rec_exec (Pr n f g) ([Suc x]) = |
|
364 rec_exec g ([x] @ [rec_exec (Pr n f g) ([x])])" |
|
365 by(simp add: rec_exec.simps) |
|
366 |
|
367 lemma Sigma_0_simp_rewrite_single_param: |
|
368 "Sigma f [0] = f [0]" |
|
369 by(simp add: Sigma.simps) |
|
370 |
|
371 lemma Sigma_0_simp_rewrite: |
|
372 "Sigma f (xs @ [0]) = f (xs @ [0])" |
|
373 by(simp add: Sigma.simps) |
|
374 |
|
375 lemma Sigma_Suc_simp_rewrite: |
|
376 "Sigma f (xs @ [Suc x]) = Sigma f (xs @ [x]) + f (xs @ [Suc x])" |
|
377 by(simp add: Sigma.simps) |
|
378 |
|
379 lemma Sigma_Suc_simp_rewrite_single: |
|
380 "Sigma f ([Suc x]) = Sigma f ([x]) + f ([Suc x])" |
|
381 by(simp add: Sigma.simps) |
|
382 |
|
383 lemma [simp]: "(xs @ ys) ! (Suc (length xs)) = ys ! 1" |
|
384 by(simp add: nth_append) |
|
385 |
|
386 lemma get_fstn_args_take: "\<lbrakk>length xs = m; n \<le> m\<rbrakk> \<Longrightarrow> |
|
387 map (\<lambda> f. rec_exec f xs) (get_fstn_args m n)= take n xs" |
|
388 proof(induct n) |
|
389 case 0 thus "?case" |
|
390 by(simp add: get_fstn_args.simps) |
|
391 next |
|
392 case (Suc n) thus "?case" |
|
393 by(simp add: get_fstn_args.simps rec_exec.simps |
|
394 take_Suc_conv_app_nth) |
|
395 qed |
|
396 |
|
397 lemma [simp]: "primerec f n \<Longrightarrow> arity f = n" |
|
398 apply(case_tac f) |
|
399 apply(auto simp: arity.simps ) |
|
400 apply(erule_tac prime_mn_reverse) |
|
401 done |
|
402 |
|
403 lemma rec_sigma_Suc_simp_rewrite: |
|
404 "primerec f (Suc (length xs)) |
|
405 \<Longrightarrow> rec_exec (rec_sigma f) (xs @ [Suc x]) = |
|
406 rec_exec (rec_sigma f) (xs @ [x]) + rec_exec f (xs @ [Suc x])" |
|
407 apply(induct x) |
|
408 apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite |
|
409 rec_exec.simps get_fstn_args_take) |
|
410 done |
|
411 |
|
412 text {* |
|
413 The correctness of @{text "rec_sigma"} with respect to its specification. |
|
414 *} |
|
415 lemma sigma_lemma: |
|
416 "primerec rg (Suc (length xs)) |
|
417 \<Longrightarrow> rec_exec (rec_sigma rg) (xs @ [x]) = Sigma (rec_exec rg) (xs @ [x])" |
|
418 apply(induct x) |
|
419 apply(auto simp: rec_exec.simps rec_sigma.simps Let_def |
|
420 get_fstn_args_take Sigma_0_simp_rewrite |
|
421 Sigma_Suc_simp_rewrite) |
|
422 done |
|
423 |
|
424 text {* |
|
425 @{text "rec_accum f (x1, x2, \<dots>, xn, k) = |
|
426 f(x1, x2, \<dots>, xn, 0) * |
|
427 f(x1, x2, \<dots>, xn, 1) * |
|
428 \<dots> |
|
429 f(x1, x2, \<dots>, xn, k)"} |
|
430 *} |
|
431 fun rec_accum :: "recf \<Rightarrow> recf" |
|
432 where |
|
433 "rec_accum rf = |
|
434 (let vl = arity rf in |
|
435 Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ |
|
436 [Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) |
|
437 (Cn (Suc vl) rec_mult [id (Suc vl) (vl), |
|
438 Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) |
|
439 @ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" |
|
440 |
|
441 text {* |
|
442 @{text "Accum"} is the formal specification of @{text "rec_accum"}. |
|
443 *} |
|
444 function Accum :: "(nat list \<Rightarrow> nat) \<Rightarrow> nat list \<Rightarrow> nat" |
|
445 where |
|
446 "Accum f xs = (if last xs = 0 then f xs |
|
447 else (Accum f (butlast xs @ [last xs - 1]) * |
|
448 f xs))" |
|
449 by pat_completeness auto |
|
450 termination |
|
451 proof |
|
452 show "wf (measure (\<lambda> (f, xs). last xs))" |
|
453 by auto |
|
454 next |
|
455 fix f xs |
|
456 assume "last xs \<noteq> (0::nat)" |
|
457 thus "((f, butlast xs @ [last xs - 1]), f, xs) \<in> |
|
458 measure (\<lambda>(f, xs). last xs)" |
|
459 by auto |
|
460 qed |
|
461 |
|
462 lemma rec_accum_Suc_simp_rewrite: |
|
463 "primerec f (Suc (length xs)) |
|
464 \<Longrightarrow> rec_exec (rec_accum f) (xs @ [Suc x]) = |
|
465 rec_exec (rec_accum f) (xs @ [x]) * rec_exec f (xs @ [Suc x])" |
|
466 apply(induct x) |
|
467 apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite |
|
468 rec_exec.simps get_fstn_args_take) |
|
469 done |
|
470 |
|
471 text {* |
|
472 The correctness of @{text "rec_accum"} with respect to its specification. |
|
473 *} |
|
474 lemma accum_lemma : |
|
475 "primerec rg (Suc (length xs)) |
|
476 \<Longrightarrow> rec_exec (rec_accum rg) (xs @ [x]) = Accum (rec_exec rg) (xs @ [x])" |
|
477 apply(induct x) |
|
478 apply(auto simp: rec_exec.simps rec_sigma.simps Let_def |
|
479 get_fstn_args_take) |
|
480 done |
|
481 |
|
482 declare rec_accum.simps [simp del] |
|
483 |
|
484 text {* |
|
485 @{text "rec_all t f (x1, x2, \<dots>, xn)"} |
|
486 computes the charactrization function of the following FOL formula: |
|
487 @{text "(\<forall> x \<le> t(x1, x2, \<dots>, xn). (f(x1, x2, \<dots>, xn, x) > 0))"} |
|
488 *} |
|
489 fun rec_all :: "recf \<Rightarrow> recf \<Rightarrow> recf" |
|
490 where |
|
491 "rec_all rt rf = |
|
492 (let vl = arity rf in |
|
493 Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_accum rf) |
|
494 (get_fstn_args (vl - 1) (vl - 1) @ [rt])])" |
|
495 |
|
496 lemma rec_accum_ex: "primerec rf (Suc (length xs)) \<Longrightarrow> |
|
497 (rec_exec (rec_accum rf) (xs @ [x]) = 0) = |
|
498 (\<exists> t \<le> x. rec_exec rf (xs @ [t]) = 0)" |
|
499 apply(induct x, simp_all add: rec_accum_Suc_simp_rewrite) |
|
500 apply(simp add: rec_exec.simps rec_accum.simps get_fstn_args_take, |
|
501 auto) |
|
502 apply(rule_tac x = ta in exI, simp) |
|
503 apply(case_tac "t = Suc x", simp_all) |
|
504 apply(rule_tac x = t in exI, simp) |
|
505 done |
|
506 |
|
507 text {* |
|
508 The correctness of @{text "rec_all"}. |
|
509 *} |
|
510 lemma all_lemma: |
|
511 "\<lbrakk>primerec rf (Suc (length xs)); |
|
512 primerec rt (length xs)\<rbrakk> |
|
513 \<Longrightarrow> rec_exec (rec_all rt rf) xs = (if (\<forall> x \<le> (rec_exec rt xs). 0 < rec_exec rf (xs @ [x])) then 1 |
|
514 else 0)" |
|
515 apply(auto simp: rec_all.simps) |
|
516 apply(simp add: rec_exec.simps map_append get_fstn_args_take split: if_splits) |
|
517 apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) |
|
518 apply(case_tac "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0", simp_all) |
|
519 apply(erule_tac exE, erule_tac x = t in allE, simp) |
|
520 apply(simp add: rec_exec.simps map_append get_fstn_args_take) |
|
521 apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) |
|
522 apply(case_tac "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0", simp, simp) |
|
523 apply(erule_tac x = x in allE, simp) |
|
524 done |
|
525 |
|
526 text {* |
|
527 @{text "rec_ex t f (x1, x2, \<dots>, xn)"} |
|
528 computes the charactrization function of the following FOL formula: |
|
529 @{text "(\<exists> x \<le> t(x1, x2, \<dots>, xn). (f(x1, x2, \<dots>, xn, x) > 0))"} |
|
530 *} |
|
531 fun rec_ex :: "recf \<Rightarrow> recf \<Rightarrow> recf" |
|
532 where |
|
533 "rec_ex rt rf = |
|
534 (let vl = arity rf in |
|
535 Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_sigma rf) |
|
536 (get_fstn_args (vl - 1) (vl - 1) @ [rt])])" |
|
537 |
|
538 lemma rec_sigma_ex: "primerec rf (Suc (length xs)) |
|
539 \<Longrightarrow> (rec_exec (rec_sigma rf) (xs @ [x]) = 0) = |
|
540 (\<forall> t \<le> x. rec_exec rf (xs @ [t]) = 0)" |
|
541 apply(induct x, simp_all add: rec_sigma_Suc_simp_rewrite) |
|
542 apply(simp add: rec_exec.simps rec_sigma.simps |
|
543 get_fstn_args_take, auto) |
|
544 apply(case_tac "t = Suc x", simp_all) |
|
545 done |
|
546 |
|
547 text {* |
|
548 The correctness of @{text "ex_lemma"}. |
|
549 *} |
|
550 lemma ex_lemma:" |
|
551 \<lbrakk>primerec rf (Suc (length xs)); |
|
552 primerec rt (length xs)\<rbrakk> |
|
553 \<Longrightarrow> (rec_exec (rec_ex rt rf) xs = |
|
554 (if (\<exists> x \<le> (rec_exec rt xs). 0 <rec_exec rf (xs @ [x])) then 1 |
|
555 else 0))" |
|
556 apply(auto simp: rec_ex.simps rec_exec.simps map_append get_fstn_args_take |
|
557 split: if_splits) |
|
558 apply(drule_tac x = "rec_exec rt xs" in rec_sigma_ex, simp) |
|
559 apply(drule_tac x = "rec_exec rt xs" in rec_sigma_ex, simp) |
|
560 done |
|
561 |
|
562 text {* |
|
563 Defintiion of @{text "Min[R]"} on page 77 of Boolos's book. |
|
564 *} |
|
565 |
|
566 fun Minr :: "(nat list \<Rightarrow> bool) \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> nat" |
|
567 where "Minr Rr xs w = (let setx = {y | y. (y \<le> w) \<and> Rr (xs @ [y])} in |
|
568 if (setx = {}) then (Suc w) |
|
569 else (Min setx))" |
|
570 |
|
571 declare Minr.simps[simp del] rec_all.simps[simp del] |
|
572 |
|
573 text {* |
|
574 The following is a set of auxilliary lemmas about @{text "Minr"}. |
|
575 *} |
|
576 lemma Minr_range: "Minr Rr xs w \<le> w \<or> Minr Rr xs w = Suc w" |
|
577 apply(auto simp: Minr.simps) |
|
578 apply(subgoal_tac "Min {x. x \<le> w \<and> Rr (xs @ [x])} \<le> x") |
|
579 apply(erule_tac order_trans, simp) |
|
580 apply(rule_tac Min_le, auto) |
|
581 done |
|
582 |
|
583 lemma [simp]: "{x. x \<le> Suc w \<and> Rr (xs @ [x])} |
|
584 = (if Rr (xs @ [Suc w]) then insert (Suc w) |
|
585 {x. x \<le> w \<and> Rr (xs @ [x])} |
|
586 else {x. x \<le> w \<and> Rr (xs @ [x])})" |
|
587 by(auto, case_tac "x = Suc w", auto) |
|
588 |
|
589 lemma [simp]: "Minr Rr xs w \<le> w \<Longrightarrow> Minr Rr xs (Suc w) = Minr Rr xs w" |
|
590 apply(simp add: Minr.simps, auto) |
|
591 apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto) |
|
592 done |
|
593 |
|
594 lemma [simp]: "\<forall>x\<le>w. \<not> Rr (xs @ [x]) \<Longrightarrow> |
|
595 {x. x \<le> w \<and> Rr (xs @ [x])} = {} " |
|
596 by auto |
|
597 |
|
598 lemma [simp]: "\<lbrakk>Minr Rr xs w = Suc w; Rr (xs @ [Suc w])\<rbrakk> \<Longrightarrow> |
|
599 Minr Rr xs (Suc w) = Suc w" |
|
600 apply(simp add: Minr.simps) |
|
601 apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto) |
|
602 done |
|
603 |
|
604 lemma [simp]: "\<lbrakk>Minr Rr xs w = Suc w; \<not> Rr (xs @ [Suc w])\<rbrakk> \<Longrightarrow> |
|
605 Minr Rr xs (Suc w) = Suc (Suc w)" |
|
606 apply(simp add: Minr.simps) |
|
607 apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto) |
|
608 apply(subgoal_tac "Min {x. x \<le> w \<and> Rr (xs @ [x])} \<in> |
|
609 {x. x \<le> w \<and> Rr (xs @ [x])}", simp) |
|
610 apply(rule_tac Min_in, auto) |
|
611 done |
|
612 |
|
613 lemma Minr_Suc_simp: |
|
614 "Minr Rr xs (Suc w) = |
|
615 (if Minr Rr xs w \<le> w then Minr Rr xs w |
|
616 else if (Rr (xs @ [Suc w])) then (Suc w) |
|
617 else Suc (Suc w))" |
|
618 by(insert Minr_range[of Rr xs w], auto) |
|
619 |
|
620 text {* |
|
621 @{text "rec_Minr"} is the recursive function |
|
622 used to implement @{text "Minr"}: |
|
623 if @{text "Rr"} is implemented by a recursive function @{text "recf"}, |
|
624 then @{text "rec_Minr recf"} is the recursive function used to |
|
625 implement @{text "Minr Rr"} |
|
626 *} |
|
627 fun rec_Minr :: "recf \<Rightarrow> recf" |
|
628 where |
|
629 "rec_Minr rf = |
|
630 (let vl = arity rf |
|
631 in let rq = rec_all (id vl (vl - 1)) (Cn (Suc vl) |
|
632 rec_not [Cn (Suc vl) rf |
|
633 (get_fstn_args (Suc vl) (vl - 1) @ |
|
634 [id (Suc vl) (vl)])]) |
|
635 in rec_sigma rq)" |
|
636 |
|
637 lemma length_getpren_params[simp]: "length (get_fstn_args m n) = n" |
|
638 by(induct n, auto simp: get_fstn_args.simps) |
|
639 |
|
640 lemma length_app: |
|
641 "(length (get_fstn_args (arity rf - Suc 0) |
|
642 (arity rf - Suc 0) |
|
643 @ [Cn (arity rf - Suc 0) (constn 0) |
|
644 [recf.id (arity rf - Suc 0) 0]])) |
|
645 = (Suc (arity rf - Suc 0))" |
|
646 apply(simp) |
|
647 done |
|
648 |
|
649 lemma primerec_accum: "primerec (rec_accum rf) n \<Longrightarrow> primerec rf n" |
|
650 apply(auto simp: rec_accum.simps Let_def) |
|
651 apply(erule_tac prime_pr_reverse, simp) |
|
652 apply(erule_tac prime_cn_reverse, simp only: length_app) |
|
653 done |
|
654 |
|
655 lemma primerec_all: "primerec (rec_all rt rf) n \<Longrightarrow> |
|
656 primerec rt n \<and> primerec rf (Suc n)" |
|
657 apply(simp add: rec_all.simps Let_def) |
|
658 apply(erule_tac prime_cn_reverse, simp) |
|
659 apply(erule_tac prime_cn_reverse, simp) |
|
660 apply(erule_tac x = n in allE, simp add: nth_append primerec_accum) |
|
661 done |
|
662 |
|
663 lemma min_Suc_Suc[simp]: "min (Suc (Suc x)) x = x" |
|
664 by auto |
|
665 |
|
666 declare numeral_3_eq_3[simp] |
|
667 |
|
668 lemma [intro]: "primerec rec_pred (Suc 0)" |
|
669 apply(simp add: rec_pred_def) |
|
670 apply(rule_tac prime_cn, auto) |
|
671 apply(case_tac i, auto intro: prime_id) |
|
672 done |
|
673 |
|
674 lemma [intro]: "primerec rec_minus (Suc (Suc 0))" |
|
675 apply(auto simp: rec_minus_def) |
|
676 done |
|
677 |
|
678 lemma [intro]: "primerec (constn n) (Suc 0)" |
|
679 apply(induct n) |
|
680 apply(auto simp: constn.simps intro: prime_z prime_cn prime_s) |
|
681 done |
|
682 |
|
683 lemma [intro]: "primerec rec_sg (Suc 0)" |
|
684 apply(simp add: rec_sg_def) |
|
685 apply(rule_tac k = "Suc (Suc 0)" in prime_cn, auto) |
|
686 apply(case_tac i, auto) |
|
687 apply(case_tac ia, auto intro: prime_id) |
|
688 done |
|
689 |
|
690 lemma [simp]: "length (get_fstn_args m n) = n" |
|
691 apply(induct n) |
|
692 apply(auto simp: get_fstn_args.simps) |
|
693 done |
|
694 |
|
695 lemma primerec_getpren[elim]: "\<lbrakk>i < n; n \<le> m\<rbrakk> \<Longrightarrow> primerec (get_fstn_args m n ! i) m" |
|
696 apply(induct n, auto simp: get_fstn_args.simps) |
|
697 apply(case_tac "i = n", auto simp: nth_append intro: prime_id) |
|
698 done |
|
699 |
|
700 lemma [intro]: "primerec rec_add (Suc (Suc 0))" |
|
701 apply(simp add: rec_add_def) |
|
702 apply(rule_tac prime_pr, auto) |
|
703 done |
|
704 |
|
705 lemma [intro]:"primerec rec_mult (Suc (Suc 0))" |
|
706 apply(simp add: rec_mult_def ) |
|
707 apply(rule_tac prime_pr, auto intro: prime_z) |
|
708 apply(case_tac i, auto intro: prime_id) |
|
709 done |
|
710 |
|
711 lemma [elim]: "\<lbrakk>primerec rf n; n \<ge> Suc (Suc 0)\<rbrakk> \<Longrightarrow> |
|
712 primerec (rec_accum rf) n" |
|
713 apply(auto simp: rec_accum.simps) |
|
714 apply(simp add: nth_append, auto) |
|
715 apply(case_tac i, auto intro: prime_id) |
|
716 apply(auto simp: nth_append) |
|
717 done |
|
718 |
|
719 lemma primerec_all_iff: |
|
720 "\<lbrakk>primerec rt n; primerec rf (Suc n); n > 0\<rbrakk> \<Longrightarrow> |
|
721 primerec (rec_all rt rf) n" |
|
722 apply(simp add: rec_all.simps, auto) |
|
723 apply(auto, simp add: nth_append, auto) |
|
724 done |
|
725 |
|
726 lemma [simp]: "Rr (xs @ [0]) \<Longrightarrow> |
|
727 Min {x. x = (0::nat) \<and> Rr (xs @ [x])} = 0" |
|
728 by(rule_tac Min_eqI, simp, simp, simp) |
|
729 |
|
730 lemma [intro]: "primerec rec_not (Suc 0)" |
|
731 apply(simp add: rec_not_def) |
|
732 apply(rule prime_cn, auto) |
|
733 apply(case_tac i, auto intro: prime_id) |
|
734 done |
|
735 |
|
736 lemma Min_false1[simp]: "\<lbrakk>\<not> Min {uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])} \<le> w; |
|
737 x \<le> w; 0 < rec_exec rf (xs @ [x])\<rbrakk> |
|
738 \<Longrightarrow> False" |
|
739 apply(subgoal_tac "finite {uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])}") |
|
740 apply(subgoal_tac "{uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])} \<noteq> {}") |
|
741 apply(simp add: Min_le_iff, simp) |
|
742 apply(rule_tac x = x in exI, simp) |
|
743 apply(simp) |
|
744 done |
|
745 |
|
746 lemma sigma_minr_lemma: |
|
747 assumes prrf: "primerec rf (Suc (length xs))" |
|
748 shows "UF.Sigma (rec_exec (rec_all (recf.id (Suc (length xs)) (length xs)) |
|
749 (Cn (Suc (Suc (length xs))) rec_not |
|
750 [Cn (Suc (Suc (length xs))) rf (get_fstn_args (Suc (Suc (length xs))) |
|
751 (length xs) @ [recf.id (Suc (Suc (length xs))) (Suc (length xs))])]))) |
|
752 (xs @ [w]) = |
|
753 Minr (\<lambda>args. 0 < rec_exec rf args) xs w" |
|
754 proof(induct w) |
|
755 let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" |
|
756 let ?rf = "(Cn (Suc (Suc (length xs))) |
|
757 rec_not [Cn (Suc (Suc (length xs))) rf |
|
758 (get_fstn_args (Suc (Suc (length xs))) (length xs) @ |
|
759 [recf.id (Suc (Suc (length xs))) |
|
760 (Suc ((length xs)))])])" |
|
761 let ?rq = "(rec_all ?rt ?rf)" |
|
762 have prrf: "primerec ?rf (Suc (length (xs @ [0]))) \<and> |
|
763 primerec ?rt (length (xs @ [0]))" |
|
764 apply(auto simp: prrf nth_append)+ |
|
765 done |
|
766 show "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [0]) |
|
767 = Minr (\<lambda>args. 0 < rec_exec rf args) xs 0" |
|
768 apply(simp add: Sigma.simps) |
|
769 apply(simp only: prrf all_lemma, |
|
770 auto simp: rec_exec.simps get_fstn_args_take Minr.simps) |
|
771 apply(rule_tac Min_eqI, auto) |
|
772 done |
|
773 next |
|
774 fix w |
|
775 let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" |
|
776 let ?rf = "(Cn (Suc (Suc (length xs))) |
|
777 rec_not [Cn (Suc (Suc (length xs))) rf |
|
778 (get_fstn_args (Suc (Suc (length xs))) (length xs) @ |
|
779 [recf.id (Suc (Suc (length xs))) |
|
780 (Suc ((length xs)))])])" |
|
781 let ?rq = "(rec_all ?rt ?rf)" |
|
782 assume ind: |
|
783 "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [w]) = Minr (\<lambda>args. 0 < rec_exec rf args) xs w" |
|
784 have prrf: "primerec ?rf (Suc (length (xs @ [Suc w]))) \<and> |
|
785 primerec ?rt (length (xs @ [Suc w]))" |
|
786 apply(auto simp: prrf nth_append)+ |
|
787 done |
|
788 show "UF.Sigma (rec_exec (rec_all ?rt ?rf)) |
|
789 (xs @ [Suc w]) = |
|
790 Minr (\<lambda>args. 0 < rec_exec rf args) xs (Suc w)" |
|
791 apply(auto simp: Sigma_Suc_simp_rewrite ind Minr_Suc_simp) |
|
792 apply(simp_all only: prrf all_lemma) |
|
793 apply(auto simp: rec_exec.simps get_fstn_args_take Let_def Minr.simps split: if_splits) |
|
794 apply(drule_tac Min_false1, simp, simp, simp) |
|
795 apply(case_tac "x = Suc w", simp, simp) |
|
796 apply(drule_tac Min_false1, simp, simp, simp) |
|
797 apply(drule_tac Min_false1, simp, simp, simp) |
|
798 done |
|
799 qed |
|
800 |
|
801 text {* |
|
802 The correctness of @{text "rec_Minr"}. |
|
803 *} |
|
804 lemma Minr_lemma: " |
|
805 \<lbrakk>primerec rf (Suc (length xs))\<rbrakk> |
|
806 \<Longrightarrow> rec_exec (rec_Minr rf) (xs @ [w]) = |
|
807 Minr (\<lambda> args. (0 < rec_exec rf args)) xs w" |
|
808 proof - |
|
809 let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" |
|
810 let ?rf = "(Cn (Suc (Suc (length xs))) |
|
811 rec_not [Cn (Suc (Suc (length xs))) rf |
|
812 (get_fstn_args (Suc (Suc (length xs))) (length xs) @ |
|
813 [recf.id (Suc (Suc (length xs))) |
|
814 (Suc ((length xs)))])])" |
|
815 let ?rq = "(rec_all ?rt ?rf)" |
|
816 assume h: "primerec rf (Suc (length xs))" |
|
817 have h1: "primerec ?rq (Suc (length xs))" |
|
818 apply(rule_tac primerec_all_iff) |
|
819 apply(auto simp: h nth_append)+ |
|
820 done |
|
821 moreover have "arity rf = Suc (length xs)" |
|
822 using h by auto |
|
823 ultimately show "rec_exec (rec_Minr rf) (xs @ [w]) = |
|
824 Minr (\<lambda> args. (0 < rec_exec rf args)) xs w" |
|
825 apply(simp add: rec_exec.simps rec_Minr.simps arity.simps Let_def |
|
826 sigma_lemma all_lemma) |
|
827 apply(rule_tac sigma_minr_lemma) |
|
828 apply(simp add: h) |
|
829 done |
|
830 qed |
|
831 |
|
832 text {* |
|
833 @{text "rec_le"} is the comparasion function |
|
834 which compares its two arguments, testing whether the |
|
835 first is less or equal to the second. |
|
836 *} |
|
837 definition rec_le :: "recf" |
|
838 where |
|
839 "rec_le = Cn (Suc (Suc 0)) rec_disj [rec_less, rec_eq]" |
|
840 |
|
841 text {* |
|
842 The correctness of @{text "rec_le"}. |
|
843 *} |
|
844 lemma le_lemma: |
|
845 "\<And>x y. rec_exec rec_le [x, y] = (if (x \<le> y) then 1 else 0)" |
|
846 by(auto simp: rec_le_def rec_exec.simps) |
|
847 |
|
848 text {* |
|
849 Defintiion of @{text "Max[Rr]"} on page 77 of Boolos's book. |
|
850 *} |
|
851 |
|
852 fun Maxr :: "(nat list \<Rightarrow> bool) \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> nat" |
|
853 where |
|
854 "Maxr Rr xs w = (let setx = {y. y \<le> w \<and> Rr (xs @[y])} in |
|
855 if setx = {} then 0 |
|
856 else Max setx)" |
|
857 |
|
858 text {* |
|
859 @{text "rec_maxr"} is the recursive function |
|
860 used to implementation @{text "Maxr"}. |
|
861 *} |
|
862 fun rec_maxr :: "recf \<Rightarrow> recf" |
|
863 where |
|
864 "rec_maxr rr = (let vl = arity rr in |
|
865 let rt = id (Suc vl) (vl - 1) in |
|
866 let rf1 = Cn (Suc (Suc vl)) rec_le |
|
867 [id (Suc (Suc vl)) |
|
868 ((Suc vl)), id (Suc (Suc vl)) (vl)] in |
|
869 let rf2 = Cn (Suc (Suc vl)) rec_not |
|
870 [Cn (Suc (Suc vl)) |
|
871 rr (get_fstn_args (Suc (Suc vl)) |
|
872 (vl - 1) @ |
|
873 [id (Suc (Suc vl)) ((Suc vl))])] in |
|
874 let rf = Cn (Suc (Suc vl)) rec_disj [rf1, rf2] in |
|
875 let rq = rec_all rt rf in |
|
876 let Qf = Cn (Suc vl) rec_not [rec_all rt rf] |
|
877 in Cn vl (rec_sigma Qf) (get_fstn_args vl vl @ |
|
878 [id vl (vl - 1)]))" |
|
879 |
|
880 declare rec_maxr.simps[simp del] Maxr.simps[simp del] |
|
881 declare le_lemma[simp] |
|
882 lemma [simp]: "(min (Suc (Suc (Suc (x)))) (x)) = x" |
|
883 by simp |
|
884 |
|
885 declare numeral_2_eq_2[simp] |
|
886 |
|
887 lemma [intro]: "primerec rec_disj (Suc (Suc 0))" |
|
888 apply(simp add: rec_disj_def, auto) |
|
889 apply(auto) |
|
890 apply(case_tac ia, auto intro: prime_id) |
|
891 done |
|
892 |
|
893 lemma [intro]: "primerec rec_less (Suc (Suc 0))" |
|
894 apply(simp add: rec_less_def, auto) |
|
895 apply(auto) |
|
896 apply(case_tac ia , auto intro: prime_id) |
|
897 done |
|
898 |
|
899 lemma [intro]: "primerec rec_eq (Suc (Suc 0))" |
|
900 apply(simp add: rec_eq_def) |
|
901 apply(rule_tac prime_cn, auto) |
|
902 apply(case_tac i, auto) |
|
903 apply(case_tac ia, auto) |
|
904 apply(case_tac [!] i, auto intro: prime_id) |
|
905 done |
|
906 |
|
907 lemma [intro]: "primerec rec_le (Suc (Suc 0))" |
|
908 apply(simp add: rec_le_def) |
|
909 apply(rule_tac prime_cn, auto) |
|
910 apply(case_tac i, auto) |
|
911 done |
|
912 |
|
913 lemma [simp]: |
|
914 "length ys = Suc n \<Longrightarrow> (take n ys @ [ys ! n, ys ! n]) = |
|
915 ys @ [ys ! n]" |
|
916 apply(simp) |
|
917 apply(subgoal_tac "\<exists> xs y. ys = xs @ [y]", auto) |
|
918 apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI) |
|
919 apply(case_tac "ys = []", simp_all) |
|
920 done |
|
921 |
|
922 lemma Maxr_Suc_simp: |
|
923 "Maxr Rr xs (Suc w) =(if Rr (xs @ [Suc w]) then Suc w |
|
924 else Maxr Rr xs w)" |
|
925 apply(auto simp: Maxr.simps) |
|
926 apply(rule_tac max_absorb1) |
|
927 apply(subgoal_tac "(Max {y. y \<le> w \<and> Rr (xs @ [y])} \<le> (Suc w)) = |
|
928 (\<forall>a\<in>{y. y \<le> w \<and> Rr (xs @ [y])}. a \<le> (Suc w))", simp) |
|
929 apply(rule_tac Max_le_iff, auto) |
|
930 done |
|
931 |
|
932 |
|
933 lemma [simp]: "min (Suc n) n = n" by simp |
|
934 |
|
935 lemma Sigma_0: "\<forall> i \<le> n. (f (xs @ [i]) = 0) \<Longrightarrow> |
|
936 Sigma f (xs @ [n]) = 0" |
|
937 apply(induct n, simp add: Sigma.simps) |
|
938 apply(simp add: Sigma_Suc_simp_rewrite) |
|
939 done |
|
940 |
|
941 lemma [elim]: "\<forall>k<Suc w. f (xs @ [k]) = Suc 0 |
|
942 \<Longrightarrow> Sigma f (xs @ [w]) = Suc w" |
|
943 apply(induct w) |
|
944 apply(simp add: Sigma.simps, simp) |
|
945 apply(simp add: Sigma.simps) |
|
946 done |
|
947 |
|
948 lemma Sigma_max_point: "\<lbrakk>\<forall> k < ma. f (xs @ [k]) = 1; |
|
949 \<forall> k \<ge> ma. f (xs @ [k]) = 0; ma \<le> w\<rbrakk> |
|
950 \<Longrightarrow> Sigma f (xs @ [w]) = ma" |
|
951 apply(induct w, auto) |
|
952 apply(rule_tac Sigma_0, simp) |
|
953 apply(simp add: Sigma_Suc_simp_rewrite) |
|
954 apply(case_tac "ma = Suc w", auto) |
|
955 done |
|
956 |
|
957 lemma Sigma_Max_lemma: |
|
958 assumes prrf: "primerec rf (Suc (length xs))" |
|
959 shows "UF.Sigma (rec_exec (Cn (Suc (Suc (length xs))) rec_not |
|
960 [rec_all (recf.id (Suc (Suc (length xs))) (length xs)) |
|
961 (Cn (Suc (Suc (Suc (length xs)))) rec_disj |
|
962 [Cn (Suc (Suc (Suc (length xs)))) rec_le |
|
963 [recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs))), |
|
964 recf.id (Suc (Suc (Suc (length xs)))) (Suc (length xs))], |
|
965 Cn (Suc (Suc (Suc (length xs)))) rec_not |
|
966 [Cn (Suc (Suc (Suc (length xs)))) rf |
|
967 (get_fstn_args (Suc (Suc (Suc (length xs)))) (length xs) @ |
|
968 [recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs)))])]])])) |
|
969 ((xs @ [w]) @ [w]) = |
|
970 Maxr (\<lambda>args. 0 < rec_exec rf args) xs w" |
|
971 proof - |
|
972 let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" |
|
973 let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) |
|
974 rec_le [recf.id (Suc (Suc (Suc (length xs)))) |
|
975 ((Suc (Suc (length xs)))), recf.id |
|
976 (Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" |
|
977 let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf |
|
978 (get_fstn_args (Suc (Suc (Suc (length xs)))) |
|
979 (length xs) @ |
|
980 [recf.id (Suc (Suc (Suc (length xs)))) |
|
981 ((Suc (Suc (length xs))))])" |
|
982 let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" |
|
983 let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" |
|
984 let ?rq = "rec_all ?rt ?rf" |
|
985 let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" |
|
986 show "?thesis" |
|
987 proof(auto simp: Maxr.simps) |
|
988 assume h: "\<forall>x\<le>w. rec_exec rf (xs @ [x]) = 0" |
|
989 have "primerec ?rf (Suc (length (xs @ [w, i]))) \<and> |
|
990 primerec ?rt (length (xs @ [w, i]))" |
|
991 using prrf |
|
992 apply(auto) |
|
993 apply(case_tac i, auto) |
|
994 apply(case_tac ia, auto simp: h nth_append) |
|
995 done |
|
996 hence "Sigma (rec_exec ?notrq) ((xs@[w])@[w]) = 0" |
|
997 apply(rule_tac Sigma_0) |
|
998 apply(auto simp: rec_exec.simps all_lemma |
|
999 get_fstn_args_take nth_append h) |
|
1000 done |
|
1001 thus "UF.Sigma (rec_exec ?notrq) |
|
1002 (xs @ [w, w]) = 0" |
|
1003 by simp |
|
1004 next |
|
1005 fix x |
|
1006 assume h: "x \<le> w" "0 < rec_exec rf (xs @ [x])" |
|
1007 hence "\<exists> ma. Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} = ma" |
|
1008 by auto |
|
1009 from this obtain ma where k1: |
|
1010 "Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} = ma" .. |
|
1011 hence k2: "ma \<le> w \<and> 0 < rec_exec rf (xs @ [ma])" |
|
1012 using h |
|
1013 apply(subgoal_tac |
|
1014 "Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} \<in> {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])}") |
|
1015 apply(erule_tac CollectE, simp) |
|
1016 apply(rule_tac Max_in, auto) |
|
1017 done |
|
1018 hence k3: "\<forall> k < ma. (rec_exec ?notrq (xs @ [w, k]) = 1)" |
|
1019 apply(auto simp: nth_append) |
|
1020 apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \<and> |
|
1021 primerec ?rt (length (xs @ [w, k]))") |
|
1022 apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append) |
|
1023 using prrf |
|
1024 apply(case_tac i, auto) |
|
1025 apply(case_tac ia, auto simp: h nth_append) |
|
1026 done |
|
1027 have k4: "\<forall> k \<ge> ma. (rec_exec ?notrq (xs @ [w, k]) = 0)" |
|
1028 apply(auto) |
|
1029 apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \<and> |
|
1030 primerec ?rt (length (xs @ [w, k]))") |
|
1031 apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append) |
|
1032 apply(subgoal_tac "x \<le> Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])}", |
|
1033 simp add: k1) |
|
1034 apply(rule_tac Max_ge, auto) |
|
1035 using prrf |
|
1036 apply(case_tac i, auto) |
|
1037 apply(case_tac ia, auto simp: h nth_append) |
|
1038 done |
|
1039 from k3 k4 k1 have "Sigma (rec_exec ?notrq) ((xs @ [w]) @ [w]) = ma" |
|
1040 apply(rule_tac Sigma_max_point, simp, simp, simp add: k2) |
|
1041 done |
|
1042 from k1 and this show "Sigma (rec_exec ?notrq) (xs @ [w, w]) = |
|
1043 Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])}" |
|
1044 by simp |
|
1045 qed |
|
1046 qed |
|
1047 |
|
1048 text {* |
|
1049 The correctness of @{text "rec_maxr"}. |
|
1050 *} |
|
1051 lemma Maxr_lemma: |
|
1052 assumes h: "primerec rf (Suc (length xs))" |
|
1053 shows "rec_exec (rec_maxr rf) (xs @ [w]) = |
|
1054 Maxr (\<lambda> args. 0 < rec_exec rf args) xs w" |
|
1055 proof - |
|
1056 from h have "arity rf = Suc (length xs)" |
|
1057 by auto |
|
1058 thus "?thesis" |
|
1059 proof(simp add: rec_exec.simps rec_maxr.simps nth_append get_fstn_args_take) |
|
1060 let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" |
|
1061 let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) |
|
1062 rec_le [recf.id (Suc (Suc (Suc (length xs)))) |
|
1063 ((Suc (Suc (length xs)))), recf.id |
|
1064 (Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" |
|
1065 let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf |
|
1066 (get_fstn_args (Suc (Suc (Suc (length xs)))) |
|
1067 (length xs) @ |
|
1068 [recf.id (Suc (Suc (Suc (length xs)))) |
|
1069 ((Suc (Suc (length xs))))])" |
|
1070 let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" |
|
1071 let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" |
|
1072 let ?rq = "rec_all ?rt ?rf" |
|
1073 let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" |
|
1074 have prt: "primerec ?rt (Suc (Suc (length xs)))" |
|
1075 by(auto intro: prime_id) |
|
1076 have prrf: "primerec ?rf (Suc (Suc (Suc (length xs))))" |
|
1077 apply(auto) |
|
1078 apply(case_tac i, auto) |
|
1079 apply(case_tac ia, auto intro: prime_id) |
|
1080 apply(simp add: h) |
|
1081 apply(simp add: nth_append, auto intro: prime_id) |
|
1082 done |
|
1083 from prt and prrf have prrq: "primerec ?rq |
|
1084 (Suc (Suc (length xs)))" |
|
1085 by(erule_tac primerec_all_iff, auto) |
|
1086 hence prnotrp: "primerec ?notrq (Suc (length ((xs @ [w]))))" |
|
1087 by(rule_tac prime_cn, auto) |
|
1088 have g1: "rec_exec (rec_sigma ?notrq) ((xs @ [w]) @ [w]) |
|
1089 = Maxr (\<lambda>args. 0 < rec_exec rf args) xs w" |
|
1090 using prnotrp |
|
1091 using sigma_lemma |
|
1092 apply(simp only: sigma_lemma) |
|
1093 apply(rule_tac Sigma_Max_lemma) |
|
1094 apply(simp add: h) |
|
1095 done |
|
1096 thus "rec_exec (rec_sigma ?notrq) |
|
1097 (xs @ [w, w]) = |
|
1098 Maxr (\<lambda>args. 0 < rec_exec rf args) xs w" |
|
1099 apply(simp) |
|
1100 done |
|
1101 qed |
|
1102 qed |
|
1103 |
|
1104 text {* |
|
1105 @{text "quo"} is the formal specification of division. |
|
1106 *} |
|
1107 fun quo :: "nat list \<Rightarrow> nat" |
|
1108 where |
|
1109 "quo [x, y] = (let Rr = |
|
1110 (\<lambda> zs. ((zs ! (Suc 0) * zs ! (Suc (Suc 0)) |
|
1111 \<le> zs ! 0) \<and> zs ! Suc 0 \<noteq> (0::nat))) |
|
1112 in Maxr Rr [x, y] x)" |
|
1113 |
|
1114 declare quo.simps[simp del] |
|
1115 |
|
1116 text {* |
|
1117 The following lemmas shows more directly the menaing of @{text "quo"}: |
|
1118 *} |
|
1119 lemma [elim!]: "y > 0 \<Longrightarrow> quo [x, y] = x div y" |
|
1120 proof(simp add: quo.simps Maxr.simps, auto, |
|
1121 rule_tac Max_eqI, simp, auto) |
|
1122 fix xa ya |
|
1123 assume h: "y * ya \<le> x" "y > 0" |
|
1124 hence "(y * ya) div y \<le> x div y" |
|
1125 by(insert div_le_mono[of "y * ya" x y], simp) |
|
1126 from this and h show "ya \<le> x div y" by simp |
|
1127 next |
|
1128 fix xa |
|
1129 show "y * (x div y) \<le> x" |
|
1130 apply(subgoal_tac "y * (x div y) + x mod y = x") |
|
1131 apply(rule_tac k = "x mod y" in add_leD1, simp) |
|
1132 apply(simp) |
|
1133 done |
|
1134 qed |
|
1135 |
|
1136 lemma [intro]: "quo [x, 0] = 0" |
|
1137 by(simp add: quo.simps Maxr.simps) |
|
1138 |
|
1139 lemma quo_div: "quo [x, y] = x div y" |
|
1140 by(case_tac "y=0", auto) |
|
1141 |
|
1142 text {* |
|
1143 @{text "rec_noteq"} is the recursive function testing whether its |
|
1144 two arguments are not equal. |
|
1145 *} |
|
1146 definition rec_noteq:: "recf" |
|
1147 where |
|
1148 "rec_noteq = Cn (Suc (Suc 0)) rec_not [Cn (Suc (Suc 0)) |
|
1149 rec_eq [id (Suc (Suc 0)) (0), id (Suc (Suc 0)) |
|
1150 ((Suc 0))]]" |
|
1151 |
|
1152 text {* |
|
1153 The correctness of @{text "rec_noteq"}. |
|
1154 *} |
|
1155 lemma noteq_lemma: |
|
1156 "\<And> x y. rec_exec rec_noteq [x, y] = |
|
1157 (if x \<noteq> y then 1 else 0)" |
|
1158 by(simp add: rec_exec.simps rec_noteq_def) |
|
1159 |
|
1160 declare noteq_lemma[simp] |
|
1161 |
|
1162 text {* |
|
1163 @{text "rec_quo"} is the recursive function used to implement @{text "quo"} |
|
1164 *} |
|
1165 definition rec_quo :: "recf" |
|
1166 where |
|
1167 "rec_quo = (let rR = Cn (Suc (Suc (Suc 0))) rec_conj |
|
1168 [Cn (Suc (Suc (Suc 0))) rec_le |
|
1169 [Cn (Suc (Suc (Suc 0))) rec_mult |
|
1170 [id (Suc (Suc (Suc 0))) (Suc 0), |
|
1171 id (Suc (Suc (Suc 0))) ((Suc (Suc 0)))], |
|
1172 id (Suc (Suc (Suc 0))) (0)], |
|
1173 Cn (Suc (Suc (Suc 0))) rec_noteq |
|
1174 [id (Suc (Suc (Suc 0))) (Suc (0)), |
|
1175 Cn (Suc (Suc (Suc 0))) (constn 0) |
|
1176 [id (Suc (Suc (Suc 0))) (0)]]] |
|
1177 in Cn (Suc (Suc 0)) (rec_maxr rR)) [id (Suc (Suc 0)) |
|
1178 (0),id (Suc (Suc 0)) (Suc (0)), |
|
1179 id (Suc (Suc 0)) (0)]" |
|
1180 |
|
1181 lemma [intro]: "primerec rec_conj (Suc (Suc 0))" |
|
1182 apply(simp add: rec_conj_def) |
|
1183 apply(rule_tac prime_cn, auto)+ |
|
1184 apply(case_tac i, auto intro: prime_id) |
|
1185 done |
|
1186 |
|
1187 lemma [intro]: "primerec rec_noteq (Suc (Suc 0))" |
|
1188 apply(simp add: rec_noteq_def) |
|
1189 apply(rule_tac prime_cn, auto)+ |
|
1190 apply(case_tac i, auto intro: prime_id) |
|
1191 done |
|
1192 |
|
1193 |
|
1194 lemma quo_lemma1: "rec_exec rec_quo [x, y] = quo [x, y]" |
|
1195 proof(simp add: rec_exec.simps rec_quo_def) |
|
1196 let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_conj |
|
1197 [Cn (Suc (Suc (Suc 0))) rec_le |
|
1198 [Cn (Suc (Suc (Suc 0))) rec_mult |
|
1199 [recf.id (Suc (Suc (Suc 0))) (Suc (0)), |
|
1200 recf.id (Suc (Suc (Suc 0))) (Suc (Suc (0)))], |
|
1201 recf.id (Suc (Suc (Suc 0))) (0)], |
|
1202 Cn (Suc (Suc (Suc 0))) rec_noteq |
|
1203 [recf.id (Suc (Suc (Suc 0))) |
|
1204 (Suc (0)), Cn (Suc (Suc (Suc 0))) (constn 0) |
|
1205 [recf.id (Suc (Suc (Suc 0))) (0)]]])" |
|
1206 have "rec_exec (rec_maxr ?rR) ([x, y]@ [ x]) = Maxr (\<lambda> args. 0 < rec_exec ?rR args) [x, y] x" |
|
1207 proof(rule_tac Maxr_lemma, simp) |
|
1208 show "primerec ?rR (Suc (Suc (Suc 0)))" |
|
1209 apply(auto) |
|
1210 apply(case_tac i, auto) |
|
1211 apply(case_tac [!] ia, auto) |
|
1212 apply(case_tac i, auto) |
|
1213 done |
|
1214 qed |
|
1215 hence g1: "rec_exec (rec_maxr ?rR) ([x, y, x]) = |
|
1216 Maxr (\<lambda> args. if rec_exec ?rR args = 0 then False |
|
1217 else True) [x, y] x" |
|
1218 by simp |
|
1219 have g2: "Maxr (\<lambda> args. if rec_exec ?rR args = 0 then False |
|
1220 else True) [x, y] x = quo [x, y]" |
|
1221 apply(simp add: rec_exec.simps) |
|
1222 apply(simp add: Maxr.simps quo.simps, auto) |
|
1223 done |
|
1224 from g1 and g2 show |
|
1225 "rec_exec (rec_maxr ?rR) ([x, y, x]) = quo [x, y]" |
|
1226 by simp |
|
1227 qed |
|
1228 |
|
1229 text {* |
|
1230 The correctness of @{text "quo"}. |
|
1231 *} |
|
1232 lemma quo_lemma2: "rec_exec rec_quo [x, y] = x div y" |
|
1233 using quo_lemma1[of x y] quo_div[of x y] |
|
1234 by simp |
|
1235 |
|
1236 text {* |
|
1237 @{text "rec_mod"} is the recursive function used to implement |
|
1238 the reminder function. |
|
1239 *} |
|
1240 definition rec_mod :: "recf" |
|
1241 where |
|
1242 "rec_mod = Cn (Suc (Suc 0)) rec_minus [id (Suc (Suc 0)) (0), |
|
1243 Cn (Suc (Suc 0)) rec_mult [rec_quo, id (Suc (Suc 0)) |
|
1244 (Suc (0))]]" |
|
1245 |
|
1246 text {* |
|
1247 The correctness of @{text "rec_mod"}: |
|
1248 *} |
|
1249 lemma mod_lemma: "\<And> x y. rec_exec rec_mod [x, y] = (x mod y)" |
|
1250 proof(simp add: rec_exec.simps rec_mod_def quo_lemma2) |
|
1251 fix x y |
|
1252 show "x - x div y * y = x mod (y::nat)" |
|
1253 using mod_div_equality2[of y x] |
|
1254 apply(subgoal_tac "y * (x div y) = (x div y ) * y", arith, simp) |
|
1255 done |
|
1256 qed |
|
1257 |
|
1258 text{* lemmas for embranch function*} |
|
1259 type_synonym ftype = "nat list \<Rightarrow> nat" |
|
1260 type_synonym rtype = "nat list \<Rightarrow> bool" |
|
1261 |
|
1262 text {* |
|
1263 The specifation of the mutli-way branching statement on |
|
1264 page 79 of Boolos's book. |
|
1265 *} |
|
1266 fun Embranch :: "(ftype * rtype) list \<Rightarrow> nat list \<Rightarrow> nat" |
|
1267 where |
|
1268 "Embranch [] xs = 0" | |
|
1269 "Embranch (gc # gcs) xs = ( |
|
1270 let (g, c) = gc in |
|
1271 if c xs then g xs else Embranch gcs xs)" |
|
1272 |
|
1273 fun rec_embranch' :: "(recf * recf) list \<Rightarrow> nat \<Rightarrow> recf" |
|
1274 where |
|
1275 "rec_embranch' [] vl = Cn vl z [id vl (vl - 1)]" | |
|
1276 "rec_embranch' ((rg, rc) # rgcs) vl = Cn vl rec_add |
|
1277 [Cn vl rec_mult [rg, rc], rec_embranch' rgcs vl]" |
|
1278 |
|
1279 text {* |
|
1280 @{text "rec_embrach"} is the recursive function used to implement |
|
1281 @{text "Embranch"}. |
|
1282 *} |
|
1283 fun rec_embranch :: "(recf * recf) list \<Rightarrow> recf" |
|
1284 where |
|
1285 "rec_embranch ((rg, rc) # rgcs) = |
|
1286 (let vl = arity rg in |
|
1287 rec_embranch' ((rg, rc) # rgcs) vl)" |
|
1288 |
|
1289 declare Embranch.simps[simp del] rec_embranch.simps[simp del] |
|
1290 |
|
1291 lemma embranch_all0: |
|
1292 "\<lbrakk>\<forall> j < length rcs. rec_exec (rcs ! j) xs = 0; |
|
1293 length rgs = length rcs; |
|
1294 rcs \<noteq> []; |
|
1295 list_all (\<lambda> rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk> \<Longrightarrow> |
|
1296 rec_exec (rec_embranch (zip rgs rcs)) xs = 0" |
|
1297 proof(induct rcs arbitrary: rgs, simp, case_tac rgs, simp) |
|
1298 fix a rcs rgs aa list |
|
1299 assume ind: |
|
1300 "\<And>rgs. \<lbrakk>\<forall>j<length rcs. rec_exec (rcs ! j) xs = 0; |
|
1301 length rgs = length rcs; rcs \<noteq> []; |
|
1302 list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk> \<Longrightarrow> |
|
1303 rec_exec (rec_embranch (zip rgs rcs)) xs = 0" |
|
1304 and h: "\<forall>j<length (a # rcs). rec_exec ((a # rcs) ! j) xs = 0" |
|
1305 "length rgs = length (a # rcs)" |
|
1306 "a # rcs \<noteq> []" |
|
1307 "list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ a # rcs)" |
|
1308 "rgs = aa # list" |
|
1309 have g: "rcs \<noteq> [] \<Longrightarrow> rec_exec (rec_embranch (zip list rcs)) xs = 0" |
|
1310 using h |
|
1311 by(rule_tac ind, auto) |
|
1312 show "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" |
|
1313 proof(case_tac "rcs = []", simp) |
|
1314 show "rec_exec (rec_embranch (zip rgs [a])) xs = 0" |
|
1315 using h |
|
1316 apply(simp add: rec_embranch.simps rec_exec.simps) |
|
1317 apply(erule_tac x = 0 in allE, simp) |
|
1318 done |
|
1319 next |
|
1320 assume "rcs \<noteq> []" |
|
1321 hence "rec_exec (rec_embranch (zip list rcs)) xs = 0" |
|
1322 using g by simp |
|
1323 thus "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" |
|
1324 using h |
|
1325 apply(simp add: rec_embranch.simps rec_exec.simps) |
|
1326 apply(case_tac rcs, |
|
1327 auto simp: rec_exec.simps rec_embranch.simps) |
|
1328 apply(case_tac list, |
|
1329 auto simp: rec_exec.simps rec_embranch.simps) |
|
1330 done |
|
1331 qed |
|
1332 qed |
|
1333 |
|
1334 |
|
1335 lemma embranch_exec_0: "\<lbrakk>rec_exec aa xs = 0; zip rgs list \<noteq> []; |
|
1336 list_all (\<lambda> rf. primerec rf (length xs)) ([a, aa] @ rgs @ list)\<rbrakk> |
|
1337 \<Longrightarrow> rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs |
|
1338 = rec_exec (rec_embranch (zip rgs list)) xs" |
|
1339 apply(simp add: rec_exec.simps rec_embranch.simps) |
|
1340 apply(case_tac "zip rgs list", simp, case_tac ab, |
|
1341 simp add: rec_embranch.simps rec_exec.simps) |
|
1342 apply(subgoal_tac "arity a = length xs", auto) |
|
1343 apply(subgoal_tac "arity aaa = length xs", auto) |
|
1344 apply(case_tac rgs, simp, case_tac list, simp, simp) |
|
1345 done |
|
1346 |
|
1347 lemma zip_null_iff: "\<lbrakk>length xs = k; length ys = k; zip xs ys = []\<rbrakk> \<Longrightarrow> xs = [] \<and> ys = []" |
|
1348 apply(case_tac xs, simp, simp) |
|
1349 apply(case_tac ys, simp, simp) |
|
1350 done |
|
1351 |
|
1352 lemma zip_null_gr: "\<lbrakk>length xs = k; length ys = k; zip xs ys \<noteq> []\<rbrakk> \<Longrightarrow> 0 < k" |
|
1353 apply(case_tac xs, simp, simp) |
|
1354 done |
|
1355 |
|
1356 lemma Embranch_0: |
|
1357 "\<lbrakk>length rgs = k; length rcs = k; k > 0; |
|
1358 \<forall> j < k. rec_exec (rcs ! j) xs = 0\<rbrakk> \<Longrightarrow> |
|
1359 Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs = 0" |
|
1360 proof(induct rgs arbitrary: rcs k, simp, simp) |
|
1361 fix a rgs rcs k |
|
1362 assume ind: |
|
1363 "\<And>rcs k. \<lbrakk>length rgs = k; length rcs = k; 0 < k; \<forall>j<k. rec_exec (rcs ! j) xs = 0\<rbrakk> |
|
1364 \<Longrightarrow> Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs = 0" |
|
1365 and h: "Suc (length rgs) = k" "length rcs = k" |
|
1366 "\<forall>j<k. rec_exec (rcs ! j) xs = 0" |
|
1367 from h show |
|
1368 "Embranch (zip (rec_exec a # map rec_exec rgs) |
|
1369 (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs = 0" |
|
1370 apply(case_tac rcs, simp, case_tac "rgs = []", simp) |
|
1371 apply(simp add: Embranch.simps) |
|
1372 apply(erule_tac x = 0 in allE, simp) |
|
1373 apply(simp add: Embranch.simps) |
|
1374 apply(erule_tac x = 0 in all_dupE, simp) |
|
1375 apply(rule_tac ind, simp, simp, simp, auto) |
|
1376 apply(erule_tac x = "Suc j" in allE, simp) |
|
1377 done |
|
1378 qed |
|
1379 |
|
1380 text {* |
|
1381 The correctness of @{text "rec_embranch"}. |
|
1382 *} |
|
1383 lemma embranch_lemma: |
|
1384 assumes branch_num: |
|
1385 "length rgs = n" "length rcs = n" "n > 0" |
|
1386 and partition: |
|
1387 "(\<exists> i < n. (rec_exec (rcs ! i) xs = 1 \<and> (\<forall> j < n. j \<noteq> i \<longrightarrow> |
|
1388 rec_exec (rcs ! j) xs = 0)))" |
|
1389 and prime_all: "list_all (\<lambda> rf. primerec rf (length xs)) (rgs @ rcs)" |
|
1390 shows "rec_exec (rec_embranch (zip rgs rcs)) xs = |
|
1391 Embranch (zip (map rec_exec rgs) |
|
1392 (map (\<lambda> r args. 0 < rec_exec r args) rcs)) xs" |
|
1393 using branch_num partition prime_all |
|
1394 proof(induct rgs arbitrary: rcs n, simp) |
|
1395 fix a rgs rcs n |
|
1396 assume ind: |
|
1397 "\<And>rcs n. \<lbrakk>length rgs = n; length rcs = n; 0 < n; |
|
1398 \<exists>i<n. rec_exec (rcs ! i) xs = 1 \<and> (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec (rcs ! j) xs = 0); |
|
1399 list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk> |
|
1400 \<Longrightarrow> rec_exec (rec_embranch (zip rgs rcs)) xs = |
|
1401 Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs" |
|
1402 and h: "length (a # rgs) = n" "length (rcs::recf list) = n" "0 < n" |
|
1403 " \<exists>i<n. rec_exec (rcs ! i) xs = 1 \<and> |
|
1404 (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec (rcs ! j) xs = 0)" |
|
1405 "list_all (\<lambda>rf. primerec rf (length xs)) ((a # rgs) @ rcs)" |
|
1406 from h show "rec_exec (rec_embranch (zip (a # rgs) rcs)) xs = |
|
1407 Embranch (zip (map rec_exec (a # rgs)) (map (\<lambda>r args. |
|
1408 0 < rec_exec r args) rcs)) xs" |
|
1409 apply(case_tac rcs, simp, simp) |
|
1410 apply(case_tac "rec_exec aa xs = 0") |
|
1411 apply(case_tac [!] "zip rgs list = []", simp) |
|
1412 apply(subgoal_tac "rgs = [] \<and> list = []", simp add: Embranch.simps rec_exec.simps rec_embranch.simps) |
|
1413 apply(rule_tac zip_null_iff, simp, simp, simp) |
|
1414 proof - |
|
1415 fix aa list |
|
1416 assume g: |
|
1417 "Suc (length rgs) = n" "Suc (length list) = n" |
|
1418 "\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> |
|
1419 (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)" |
|
1420 "primerec a (length xs) \<and> |
|
1421 list_all (\<lambda>rf. primerec rf (length xs)) rgs \<and> |
|
1422 primerec aa (length xs) \<and> |
|
1423 list_all (\<lambda>rf. primerec rf (length xs)) list" |
|
1424 "rec_exec aa xs = 0" "rcs = aa # list" "zip rgs list \<noteq> []" |
|
1425 have "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs |
|
1426 = rec_exec (rec_embranch (zip rgs list)) xs" |
|
1427 apply(rule embranch_exec_0, simp_all add: g) |
|
1428 done |
|
1429 from g and this show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = |
|
1430 Embranch ((rec_exec a, \<lambda>args. 0 < rec_exec aa args) # |
|
1431 zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) list)) xs" |
|
1432 apply(simp add: Embranch.simps) |
|
1433 apply(rule_tac n = "n - Suc 0" in ind) |
|
1434 apply(case_tac n, simp, simp) |
|
1435 apply(case_tac n, simp, simp) |
|
1436 apply(case_tac n, simp, simp add: zip_null_gr ) |
|
1437 apply(auto) |
|
1438 apply(case_tac i, simp, simp) |
|
1439 apply(rule_tac x = nat in exI, simp) |
|
1440 apply(rule_tac allI, erule_tac x = "Suc j" in allE, simp) |
|
1441 done |
|
1442 next |
|
1443 fix aa list |
|
1444 assume g: "Suc (length rgs) = n" "Suc (length list) = n" |
|
1445 "\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> |
|
1446 (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)" |
|
1447 "primerec a (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) rgs \<and> |
|
1448 primerec aa (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) list" |
|
1449 "rcs = aa # list" "rec_exec aa xs \<noteq> 0" "zip rgs list = []" |
|
1450 thus "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = |
|
1451 Embranch ((rec_exec a, \<lambda>args. 0 < rec_exec aa args) # |
|
1452 zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) list)) xs" |
|
1453 apply(subgoal_tac "rgs = [] \<and> list = []", simp) |
|
1454 prefer 2 |
|
1455 apply(rule_tac zip_null_iff, simp, simp, simp) |
|
1456 apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps, auto) |
|
1457 done |
|
1458 next |
|
1459 fix aa list |
|
1460 assume g: "Suc (length rgs) = n" "Suc (length list) = n" |
|
1461 "\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> |
|
1462 (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)" |
|
1463 "primerec a (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) rgs |
|
1464 \<and> primerec aa (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) list" |
|
1465 "rcs = aa # list" "rec_exec aa xs \<noteq> 0" "zip rgs list \<noteq> []" |
|
1466 have "rec_exec aa xs = Suc 0" |
|
1467 using g |
|
1468 apply(case_tac "rec_exec aa xs", simp, auto) |
|
1469 done |
|
1470 moreover have "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" |
|
1471 proof - |
|
1472 have "rec_embranch' (zip rgs list) (length xs) = rec_embranch (zip rgs list)" |
|
1473 using g |
|
1474 apply(case_tac "zip rgs list", simp, case_tac ab) |
|
1475 apply(simp add: rec_embranch.simps) |
|
1476 apply(subgoal_tac "arity aaa = length xs", simp, auto) |
|
1477 apply(case_tac rgs, simp, simp, case_tac list, simp, simp) |
|
1478 done |
|
1479 moreover have "rec_exec (rec_embranch (zip rgs list)) xs = 0" |
|
1480 proof(rule embranch_all0) |
|
1481 show " \<forall>j<length list. rec_exec (list ! j) xs = 0" |
|
1482 using g |
|
1483 apply(auto) |
|
1484 apply(case_tac i, simp) |
|
1485 apply(erule_tac x = "Suc j" in allE, simp) |
|
1486 apply(simp) |
|
1487 apply(erule_tac x = 0 in allE, simp) |
|
1488 done |
|
1489 next |
|
1490 show "length rgs = length list" |
|
1491 using g |
|
1492 apply(case_tac n, simp, simp) |
|
1493 done |
|
1494 next |
|
1495 show "list \<noteq> []" |
|
1496 using g |
|
1497 apply(case_tac list, simp, simp) |
|
1498 done |
|
1499 next |
|
1500 show "list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ list)" |
|
1501 using g |
|
1502 apply auto |
|
1503 done |
|
1504 qed |
|
1505 ultimately show "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" |
|
1506 by simp |
|
1507 qed |
|
1508 moreover have |
|
1509 "Embranch (zip (map rec_exec rgs) |
|
1510 (map (\<lambda>r args. 0 < rec_exec r args) list)) xs = 0" |
|
1511 using g |
|
1512 apply(rule_tac k = "length rgs" in Embranch_0) |
|
1513 apply(simp, case_tac n, simp, simp) |
|
1514 apply(case_tac rgs, simp, simp) |
|
1515 apply(auto) |
|
1516 apply(case_tac i, simp) |
|
1517 apply(erule_tac x = "Suc j" in allE, simp) |
|
1518 apply(simp) |
|
1519 apply(rule_tac x = 0 in allE, auto) |
|
1520 done |
|
1521 moreover have "arity a = length xs" |
|
1522 using g |
|
1523 apply(auto) |
|
1524 done |
|
1525 ultimately show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = |
|
1526 Embranch ((rec_exec a, \<lambda>args. 0 < rec_exec aa args) # |
|
1527 zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) list)) xs" |
|
1528 apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps) |
|
1529 done |
|
1530 qed |
|
1531 qed |
|
1532 |
|
1533 text{* |
|
1534 @{text "prime n"} means @{text "n"} is a prime number. |
|
1535 *} |
|
1536 fun Prime :: "nat \<Rightarrow> bool" |
|
1537 where |
|
1538 "Prime x = (1 < x \<and> (\<forall> u < x. (\<forall> v < x. u * v \<noteq> x)))" |
|
1539 |
|
1540 declare Prime.simps [simp del] |
|
1541 |
|
1542 lemma primerec_all1: |
|
1543 "primerec (rec_all rt rf) n \<Longrightarrow> primerec rt n" |
|
1544 by (simp add: primerec_all) |
|
1545 |
|
1546 lemma primerec_all2: "primerec (rec_all rt rf) n \<Longrightarrow> |
|
1547 primerec rf (Suc n)" |
|
1548 by(insert primerec_all[of rt rf n], simp) |
|
1549 |
|
1550 text {* |
|
1551 @{text "rec_prime"} is the recursive function used to implement |
|
1552 @{text "Prime"}. |
|
1553 *} |
|
1554 definition rec_prime :: "recf" |
|
1555 where |
|
1556 "rec_prime = Cn (Suc 0) rec_conj |
|
1557 [Cn (Suc 0) rec_less [constn 1, id (Suc 0) (0)], |
|
1558 rec_all (Cn 1 rec_minus [id 1 0, constn 1]) |
|
1559 (rec_all (Cn 2 rec_minus [id 2 0, Cn 2 (constn 1) |
|
1560 [id 2 0]]) (Cn 3 rec_noteq |
|
1561 [Cn 3 rec_mult [id 3 1, id 3 2], id 3 0]))]" |
|
1562 |
|
1563 declare numeral_2_eq_2[simp del] numeral_3_eq_3[simp del] |
|
1564 |
|
1565 lemma exec_tmp: |
|
1566 "rec_exec (rec_all (Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]]) |
|
1567 (Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])) [x, k] = |
|
1568 ((if (\<forall>w\<le>rec_exec (Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]]) ([x, k]). |
|
1569 0 < rec_exec (Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0]) |
|
1570 ([x, k] @ [w])) then 1 else 0))" |
|
1571 apply(rule_tac all_lemma) |
|
1572 apply(auto) |
|
1573 apply(case_tac [!] i, auto) |
|
1574 apply(case_tac ia, auto simp: numeral_3_eq_3 numeral_2_eq_2) |
|
1575 done |
|
1576 |
|
1577 text {* |
|
1578 The correctness of @{text "Prime"}. |
|
1579 *} |
|
1580 lemma prime_lemma: "rec_exec rec_prime [x] = (if Prime x then 1 else 0)" |
|
1581 proof(simp add: rec_exec.simps rec_prime_def) |
|
1582 let ?rt1 = "(Cn 2 rec_minus [recf.id 2 0, |
|
1583 Cn 2 (constn (Suc 0)) [recf.id 2 0]])" |
|
1584 let ?rf1 = "(Cn 3 rec_noteq [Cn 3 rec_mult |
|
1585 [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 (0)])" |
|
1586 let ?rt2 = "(Cn (Suc 0) rec_minus |
|
1587 [recf.id (Suc 0) 0, constn (Suc 0)])" |
|
1588 let ?rf2 = "rec_all ?rt1 ?rf1" |
|
1589 have h1: "rec_exec (rec_all ?rt2 ?rf2) ([x]) = |
|
1590 (if (\<forall>k\<le>rec_exec ?rt2 ([x]). 0 < rec_exec ?rf2 ([x] @ [k])) then 1 else 0)" |
|
1591 proof(rule_tac all_lemma, simp_all) |
|
1592 show "primerec ?rf2 (Suc (Suc 0))" |
|
1593 apply(rule_tac primerec_all_iff) |
|
1594 apply(auto) |
|
1595 apply(case_tac [!] i, auto simp: numeral_2_eq_2) |
|
1596 apply(case_tac ia, auto simp: numeral_3_eq_3) |
|
1597 done |
|
1598 next |
|
1599 show "primerec (Cn (Suc 0) rec_minus |
|
1600 [recf.id (Suc 0) 0, constn (Suc 0)]) (Suc 0)" |
|
1601 apply(auto) |
|
1602 apply(case_tac i, auto) |
|
1603 done |
|
1604 qed |
|
1605 from h1 show |
|
1606 "(Suc 0 < x \<longrightarrow> (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 \<longrightarrow> |
|
1607 \<not> Prime x) \<and> |
|
1608 (0 < rec_exec (rec_all ?rt2 ?rf2) [x] \<longrightarrow> Prime x)) \<and> |
|
1609 (\<not> Suc 0 < x \<longrightarrow> \<not> Prime x \<and> (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 |
|
1610 \<longrightarrow> \<not> Prime x))" |
|
1611 apply(auto simp:rec_exec.simps) |
|
1612 apply(simp add: exec_tmp rec_exec.simps) |
|
1613 proof - |
|
1614 assume "\<forall>k\<le>x - Suc 0. (0::nat) < (if \<forall>w\<le>x - Suc 0. |
|
1615 0 < (if k * w \<noteq> x then 1 else (0 :: nat)) then 1 else 0)" "Suc 0 < x" |
|
1616 thus "Prime x" |
|
1617 apply(simp add: rec_exec.simps split: if_splits) |
|
1618 apply(simp add: Prime.simps, auto) |
|
1619 apply(erule_tac x = u in allE, auto) |
|
1620 apply(case_tac u, simp, case_tac nat, simp, simp) |
|
1621 apply(case_tac v, simp, case_tac nat, simp, simp) |
|
1622 done |
|
1623 next |
|
1624 assume "\<not> Suc 0 < x" "Prime x" |
|
1625 thus "False" |
|
1626 apply(simp add: Prime.simps) |
|
1627 done |
|
1628 next |
|
1629 fix k |
|
1630 assume "rec_exec (rec_all ?rt1 ?rf1) |
|
1631 [x, k] = 0" "k \<le> x - Suc 0" "Prime x" |
|
1632 thus "False" |
|
1633 apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) |
|
1634 done |
|
1635 next |
|
1636 fix k |
|
1637 assume "rec_exec (rec_all ?rt1 ?rf1) |
|
1638 [x, k] = 0" "k \<le> x - Suc 0" "Prime x" |
|
1639 thus "False" |
|
1640 apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) |
|
1641 done |
|
1642 qed |
|
1643 qed |
|
1644 |
|
1645 definition rec_dummyfac :: "recf" |
|
1646 where |
|
1647 "rec_dummyfac = Pr 1 (constn 1) |
|
1648 (Cn 3 rec_mult [id 3 2, Cn 3 s [id 3 1]])" |
|
1649 |
|
1650 text {* |
|
1651 The recursive function used to implment factorization. |
|
1652 *} |
|
1653 definition rec_fac :: "recf" |
|
1654 where |
|
1655 "rec_fac = Cn 1 rec_dummyfac [id 1 0, id 1 0]" |
|
1656 |
|
1657 text {* |
|
1658 Formal specification of factorization. |
|
1659 *} |
|
1660 fun fac :: "nat \<Rightarrow> nat" ("_!" [100] 99) |
|
1661 where |
|
1662 "fac 0 = 1" | |
|
1663 "fac (Suc x) = (Suc x) * fac x" |
|
1664 |
|
1665 lemma [simp]: "rec_exec rec_dummyfac [0, 0] = Suc 0" |
|
1666 by(simp add: rec_dummyfac_def rec_exec.simps) |
|
1667 |
|
1668 lemma rec_cn_simp: "rec_exec (Cn n f gs) xs = |
|
1669 (let rgs = map (\<lambda> g. rec_exec g xs) gs in |
|
1670 rec_exec f rgs)" |
|
1671 by(simp add: rec_exec.simps) |
|
1672 |
|
1673 lemma rec_id_simp: "rec_exec (id m n) xs = xs ! n" |
|
1674 by(simp add: rec_exec.simps) |
|
1675 |
|
1676 lemma fac_dummy: "rec_exec rec_dummyfac [x, y] = y !" |
|
1677 apply(induct y) |
|
1678 apply(auto simp: rec_dummyfac_def rec_exec.simps) |
|
1679 done |
|
1680 |
|
1681 text {* |
|
1682 The correctness of @{text "rec_fac"}. |
|
1683 *} |
|
1684 lemma fac_lemma: "rec_exec rec_fac [x] = x!" |
|
1685 apply(simp add: rec_fac_def rec_exec.simps fac_dummy) |
|
1686 done |
|
1687 |
|
1688 declare fac.simps[simp del] |
|
1689 |
|
1690 text {* |
|
1691 @{text "Np x"} returns the first prime number after @{text "x"}. |
|
1692 *} |
|
1693 fun Np ::"nat \<Rightarrow> nat" |
|
1694 where |
|
1695 "Np x = Min {y. y \<le> Suc (x!) \<and> x < y \<and> Prime y}" |
|
1696 |
|
1697 declare Np.simps[simp del] rec_Minr.simps[simp del] |
|
1698 |
|
1699 text {* |
|
1700 @{text "rec_np"} is the recursive function used to implement |
|
1701 @{text "Np"}. |
|
1702 *} |
|
1703 definition rec_np :: "recf" |
|
1704 where |
|
1705 "rec_np = (let Rr = Cn 2 rec_conj [Cn 2 rec_less [id 2 0, id 2 1], |
|
1706 Cn 2 rec_prime [id 2 1]] |
|
1707 in Cn 1 (rec_Minr Rr) [id 1 0, Cn 1 s [rec_fac]])" |
|
1708 |
|
1709 lemma [simp]: "n < Suc (n!)" |
|
1710 apply(induct n, simp) |
|
1711 apply(simp add: fac.simps) |
|
1712 apply(case_tac n, auto simp: fac.simps) |
|
1713 done |
|
1714 |
|
1715 lemma divsor_ex: |
|
1716 "\<lbrakk>\<not> Prime x; x > Suc 0\<rbrakk> \<Longrightarrow> (\<exists> u > Suc 0. (\<exists> v > Suc 0. u * v = x))" |
|
1717 by(auto simp: Prime.simps) |
|
1718 |
|
1719 lemma divsor_prime_ex: "\<lbrakk>\<not> Prime x; x > Suc 0\<rbrakk> \<Longrightarrow> |
|
1720 \<exists> p. Prime p \<and> p dvd x" |
|
1721 apply(induct x rule: wf_induct[where r = "measure (\<lambda> y. y)"], simp) |
|
1722 apply(drule_tac divsor_ex, simp, auto) |
|
1723 apply(erule_tac x = u in allE, simp) |
|
1724 apply(case_tac "Prime u", simp) |
|
1725 apply(rule_tac x = u in exI, simp, auto) |
|
1726 done |
|
1727 |
|
1728 lemma [intro]: "0 < n!" |
|
1729 apply(induct n) |
|
1730 apply(auto simp: fac.simps) |
|
1731 done |
|
1732 |
|
1733 lemma fac_Suc: "Suc n! = (Suc n) * (n!)" by(simp add: fac.simps) |
|
1734 |
|
1735 lemma fac_dvd: "\<lbrakk>0 < q; q \<le> n\<rbrakk> \<Longrightarrow> q dvd n!" |
|
1736 apply(induct n, simp) |
|
1737 apply(case_tac "q \<le> n", simp add: fac_Suc) |
|
1738 apply(subgoal_tac "q = Suc n", simp only: fac_Suc) |
|
1739 apply(rule_tac dvd_mult2, simp, simp) |
|
1740 done |
|
1741 |
|
1742 lemma fac_dvd2: "\<lbrakk>Suc 0 < q; q dvd n!; q \<le> n\<rbrakk> \<Longrightarrow> \<not> q dvd Suc (n!)" |
|
1743 proof(auto simp: dvd_def) |
|
1744 fix k ka |
|
1745 assume h1: "Suc 0 < q" "q \<le> n" |
|
1746 and h2: "Suc (q * k) = q * ka" |
|
1747 have "k < ka" |
|
1748 proof - |
|
1749 have "q * k < q * ka" |
|
1750 using h2 by arith |
|
1751 thus "k < ka" |
|
1752 using h1 |
|
1753 by(auto) |
|
1754 qed |
|
1755 hence "\<exists>d. d > 0 \<and> ka = d + k" |
|
1756 by(rule_tac x = "ka - k" in exI, simp) |
|
1757 from this obtain d where "d > 0 \<and> ka = d + k" .. |
|
1758 from h2 and this and h1 show "False" |
|
1759 by(simp add: add_mult_distrib2) |
|
1760 qed |
|
1761 |
|
1762 lemma prime_ex: "\<exists> p. n < p \<and> p \<le> Suc (n!) \<and> Prime p" |
|
1763 proof(cases "Prime (n! + 1)") |
|
1764 case True thus "?thesis" |
|
1765 by(rule_tac x = "Suc (n!)" in exI, simp) |
|
1766 next |
|
1767 assume h: "\<not> Prime (n! + 1)" |
|
1768 hence "\<exists> p. Prime p \<and> p dvd (n! + 1)" |
|
1769 by(erule_tac divsor_prime_ex, auto) |
|
1770 from this obtain q where k: "Prime q \<and> q dvd (n! + 1)" .. |
|
1771 thus "?thesis" |
|
1772 proof(cases "q > n") |
|
1773 case True thus "?thesis" |
|
1774 using k |
|
1775 apply(rule_tac x = q in exI, auto) |
|
1776 apply(rule_tac dvd_imp_le, auto) |
|
1777 done |
|
1778 next |
|
1779 case False thus "?thesis" |
|
1780 proof - |
|
1781 assume g: "\<not> n < q" |
|
1782 have j: "q > Suc 0" |
|
1783 using k by(case_tac q, auto simp: Prime.simps) |
|
1784 hence "q dvd n!" |
|
1785 using g |
|
1786 apply(rule_tac fac_dvd, auto) |
|
1787 done |
|
1788 hence "\<not> q dvd Suc (n!)" |
|
1789 using g j |
|
1790 by(rule_tac fac_dvd2, auto) |
|
1791 thus "?thesis" |
|
1792 using k by simp |
|
1793 qed |
|
1794 qed |
|
1795 qed |
|
1796 |
|
1797 lemma Suc_Suc_induct[elim!]: "\<lbrakk>i < Suc (Suc 0); |
|
1798 primerec (ys ! 0) n; primerec (ys ! 1) n\<rbrakk> \<Longrightarrow> primerec (ys ! i) n" |
|
1799 by(case_tac i, auto) |
|
1800 |
|
1801 lemma [intro]: "primerec rec_prime (Suc 0)" |
|
1802 apply(auto simp: rec_prime_def, auto) |
|
1803 apply(rule_tac primerec_all_iff, auto, auto) |
|
1804 apply(rule_tac primerec_all_iff, auto, auto simp: |
|
1805 numeral_2_eq_2 numeral_3_eq_3) |
|
1806 done |
|
1807 |
|
1808 text {* |
|
1809 The correctness of @{text "rec_np"}. |
|
1810 *} |
|
1811 lemma np_lemma: "rec_exec rec_np [x] = Np x" |
|
1812 proof(auto simp: rec_np_def rec_exec.simps Let_def fac_lemma) |
|
1813 let ?rr = "(Cn 2 rec_conj [Cn 2 rec_less [recf.id 2 0, |
|
1814 recf.id 2 (Suc 0)], Cn 2 rec_prime [recf.id 2 (Suc 0)]])" |
|
1815 let ?R = "\<lambda> zs. zs ! 0 < zs ! 1 \<and> Prime (zs ! 1)" |
|
1816 have g1: "rec_exec (rec_Minr ?rr) ([x] @ [Suc (x!)]) = |
|
1817 Minr (\<lambda> args. 0 < rec_exec ?rr args) [x] (Suc (x!))" |
|
1818 by(rule_tac Minr_lemma, auto simp: rec_exec.simps |
|
1819 prime_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) |
|
1820 have g2: "Minr (\<lambda> args. 0 < rec_exec ?rr args) [x] (Suc (x!)) = Np x" |
|
1821 using prime_ex[of x] |
|
1822 apply(auto simp: Minr.simps Np.simps rec_exec.simps) |
|
1823 apply(erule_tac x = p in allE, simp add: prime_lemma) |
|
1824 apply(simp add: prime_lemma split: if_splits) |
|
1825 apply(subgoal_tac |
|
1826 "{uu. (Prime uu \<longrightarrow> (x < uu \<longrightarrow> uu \<le> Suc (x!)) \<and> x < uu) \<and> Prime uu} |
|
1827 = {y. y \<le> Suc (x!) \<and> x < y \<and> Prime y}", auto) |
|
1828 done |
|
1829 from g1 and g2 show "rec_exec (rec_Minr ?rr) ([x, Suc (x!)]) = Np x" |
|
1830 by simp |
|
1831 qed |
|
1832 |
|
1833 text {* |
|
1834 @{text "rec_power"} is the recursive function used to implement |
|
1835 power function. |
|
1836 *} |
|
1837 definition rec_power :: "recf" |
|
1838 where |
|
1839 "rec_power = Pr 1 (constn 1) (Cn 3 rec_mult [id 3 0, id 3 2])" |
|
1840 |
|
1841 text {* |
|
1842 The correctness of @{text "rec_power"}. |
|
1843 *} |
|
1844 lemma power_lemma: "rec_exec rec_power [x, y] = x^y" |
|
1845 by(induct y, auto simp: rec_exec.simps rec_power_def) |
|
1846 |
|
1847 text{* |
|
1848 @{text "Pi k"} returns the @{text "k"}-th prime number. |
|
1849 *} |
|
1850 fun Pi :: "nat \<Rightarrow> nat" |
|
1851 where |
|
1852 "Pi 0 = 2" | |
|
1853 "Pi (Suc x) = Np (Pi x)" |
|
1854 |
|
1855 definition rec_dummy_pi :: "recf" |
|
1856 where |
|
1857 "rec_dummy_pi = Pr 1 (constn 2) (Cn 3 rec_np [id 3 2])" |
|
1858 |
|
1859 text {* |
|
1860 @{text "rec_pi"} is the recursive function used to implement |
|
1861 @{text "Pi"}. |
|
1862 *} |
|
1863 definition rec_pi :: "recf" |
|
1864 where |
|
1865 "rec_pi = Cn 1 rec_dummy_pi [id 1 0, id 1 0]" |
|
1866 |
|
1867 lemma pi_dummy_lemma: "rec_exec rec_dummy_pi [x, y] = Pi y" |
|
1868 apply(induct y) |
|
1869 by(auto simp: rec_exec.simps rec_dummy_pi_def Pi.simps np_lemma) |
|
1870 |
|
1871 text {* |
|
1872 The correctness of @{text "rec_pi"}. |
|
1873 *} |
|
1874 lemma pi_lemma: "rec_exec rec_pi [x] = Pi x" |
|
1875 apply(simp add: rec_pi_def rec_exec.simps pi_dummy_lemma) |
|
1876 done |
|
1877 |
|
1878 fun loR :: "nat list \<Rightarrow> bool" |
|
1879 where |
|
1880 "loR [x, y, u] = (x mod (y^u) = 0)" |
|
1881 |
|
1882 declare loR.simps[simp del] |
|
1883 |
|
1884 text {* |
|
1885 @{text "Lo"} specifies the @{text "lo"} function given on page 79 of |
|
1886 Boolos's book. It is one of the two notions of integeral logarithmatic |
|
1887 operation on that page. The other is @{text "lg"}. |
|
1888 *} |
|
1889 fun lo :: " nat \<Rightarrow> nat \<Rightarrow> nat" |
|
1890 where |
|
1891 "lo x y = (if x > 1 \<and> y > 1 \<and> {u. loR [x, y, u]} \<noteq> {} then Max {u. loR [x, y, u]} |
|
1892 else 0)" |
|
1893 |
|
1894 declare lo.simps[simp del] |
|
1895 |
|
1896 lemma [elim]: "primerec rf n \<Longrightarrow> n > 0" |
|
1897 apply(induct rule: primerec.induct, auto) |
|
1898 done |
|
1899 |
|
1900 lemma primerec_sigma[intro!]: |
|
1901 "\<lbrakk>n > Suc 0; primerec rf n\<rbrakk> \<Longrightarrow> |
|
1902 primerec (rec_sigma rf) n" |
|
1903 apply(simp add: rec_sigma.simps) |
|
1904 apply(auto, auto simp: nth_append) |
|
1905 done |
|
1906 |
|
1907 lemma [intro!]: "\<lbrakk>primerec rf n; n > 0\<rbrakk> \<Longrightarrow> primerec (rec_maxr rf) n" |
|
1908 apply(simp add: rec_maxr.simps) |
|
1909 apply(rule_tac prime_cn, auto) |
|
1910 apply(rule_tac primerec_all_iff, auto, auto simp: nth_append) |
|
1911 done |
|
1912 |
|
1913 lemma Suc_Suc_Suc_induct[elim!]: |
|
1914 "\<lbrakk>i < Suc (Suc (Suc (0::nat))); primerec (ys ! 0) n; |
|
1915 primerec (ys ! 1) n; |
|
1916 primerec (ys ! 2) n\<rbrakk> \<Longrightarrow> primerec (ys ! i) n" |
|
1917 apply(case_tac i, auto, case_tac nat, simp, simp add: numeral_2_eq_2) |
|
1918 done |
|
1919 |
|
1920 lemma [intro]: "primerec rec_quo (Suc (Suc 0))" |
|
1921 apply(simp add: rec_quo_def) |
|
1922 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
1923 @{thm prime_id}] 1*}, auto+)+ |
|
1924 done |
|
1925 |
|
1926 lemma [intro]: "primerec rec_mod (Suc (Suc 0))" |
|
1927 apply(simp add: rec_mod_def) |
|
1928 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
1929 @{thm prime_id}] 1*}, auto+)+ |
|
1930 done |
|
1931 |
|
1932 lemma [intro]: "primerec rec_power (Suc (Suc 0))" |
|
1933 apply(simp add: rec_power_def numeral_2_eq_2 numeral_3_eq_3) |
|
1934 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
1935 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
1936 done |
|
1937 |
|
1938 text {* |
|
1939 @{text "rec_lo"} is the recursive function used to implement @{text "Lo"}. |
|
1940 *} |
|
1941 definition rec_lo :: "recf" |
|
1942 where |
|
1943 "rec_lo = (let rR = Cn 3 rec_eq [Cn 3 rec_mod [id 3 0, |
|
1944 Cn 3 rec_power [id 3 1, id 3 2]], |
|
1945 Cn 3 (constn 0) [id 3 1]] in |
|
1946 let rb = Cn 2 (rec_maxr rR) [id 2 0, id 2 1, id 2 0] in |
|
1947 let rcond = Cn 2 rec_conj [Cn 2 rec_less [Cn 2 (constn 1) |
|
1948 [id 2 0], id 2 0], |
|
1949 Cn 2 rec_less [Cn 2 (constn 1) |
|
1950 [id 2 0], id 2 1]] in |
|
1951 let rcond2 = Cn 2 rec_minus |
|
1952 [Cn 2 (constn 1) [id 2 0], rcond] |
|
1953 in Cn 2 rec_add [Cn 2 rec_mult [rb, rcond], |
|
1954 Cn 2 rec_mult [Cn 2 (constn 0) [id 2 0], rcond2]])" |
|
1955 |
|
1956 lemma rec_lo_Maxr_lor: |
|
1957 "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
|
1958 rec_exec rec_lo [x, y] = Maxr loR [x, y] x" |
|
1959 proof(auto simp: rec_exec.simps rec_lo_def Let_def |
|
1960 numeral_2_eq_2 numeral_3_eq_3) |
|
1961 let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_eq |
|
1962 [Cn (Suc (Suc (Suc 0))) rec_mod [recf.id (Suc (Suc (Suc 0))) 0, |
|
1963 Cn (Suc (Suc (Suc 0))) rec_power [recf.id (Suc (Suc (Suc 0))) |
|
1964 (Suc 0), recf.id (Suc (Suc (Suc 0))) (Suc (Suc 0))]], |
|
1965 Cn (Suc (Suc (Suc 0))) (constn 0) [recf.id (Suc (Suc (Suc 0))) (Suc 0)]])" |
|
1966 have h: "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) = |
|
1967 Maxr (\<lambda> args. 0 < rec_exec ?rR args) [x, y] x" |
|
1968 by(rule_tac Maxr_lemma, auto simp: rec_exec.simps |
|
1969 mod_lemma power_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) |
|
1970 have "Maxr loR [x, y] x = Maxr (\<lambda> args. 0 < rec_exec ?rR args) [x, y] x" |
|
1971 apply(simp add: rec_exec.simps mod_lemma power_lemma) |
|
1972 apply(simp add: Maxr.simps loR.simps) |
|
1973 done |
|
1974 from h and this show "rec_exec (rec_maxr ?rR) [x, y, x] = |
|
1975 Maxr loR [x, y] x" |
|
1976 apply(simp) |
|
1977 done |
|
1978 qed |
|
1979 |
|
1980 lemma [simp]: "Max {ya. ya = 0 \<and> loR [0, y, ya]} = 0" |
|
1981 apply(rule_tac Max_eqI, auto simp: loR.simps) |
|
1982 done |
|
1983 |
|
1984 lemma [simp]: "Suc 0 < y \<Longrightarrow> Suc (Suc 0) < y * y" |
|
1985 apply(induct y, simp) |
|
1986 apply(case_tac y, simp, simp) |
|
1987 done |
|
1988 |
|
1989 lemma less_mult: "\<lbrakk>x > 0; y > Suc 0\<rbrakk> \<Longrightarrow> x < y * x" |
|
1990 apply(case_tac y, simp, simp) |
|
1991 done |
|
1992 |
|
1993 lemma x_less_exp: "\<lbrakk>y > Suc 0\<rbrakk> \<Longrightarrow> x < y^x" |
|
1994 apply(induct x, simp, simp) |
|
1995 apply(case_tac x, simp, auto) |
|
1996 apply(rule_tac y = "y* y^nat" in le_less_trans, simp) |
|
1997 apply(rule_tac less_mult, auto) |
|
1998 done |
|
1999 |
|
2000 lemma le_mult: "y \<noteq> (0::nat) \<Longrightarrow> x \<le> x * y" |
|
2001 by(induct y, simp, simp) |
|
2002 |
|
2003 lemma uplimit_loR: "\<lbrakk>Suc 0 < x; Suc 0 < y; loR [x, y, xa]\<rbrakk> \<Longrightarrow> |
|
2004 xa \<le> x" |
|
2005 apply(simp add: loR.simps) |
|
2006 apply(rule_tac classical, auto) |
|
2007 apply(subgoal_tac "xa < y^xa") |
|
2008 apply(subgoal_tac "y^xa \<le> y^xa * q", simp) |
|
2009 apply(rule_tac le_mult, case_tac q, simp, simp) |
|
2010 apply(rule_tac x_less_exp, simp) |
|
2011 done |
|
2012 |
|
2013 lemma [simp]: "\<lbrakk>xa \<le> x; loR [x, y, xa]; Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
|
2014 {u. loR [x, y, u]} = {ya. ya \<le> x \<and> loR [x, y, ya]}" |
|
2015 apply(rule_tac Collect_cong, auto) |
|
2016 apply(erule_tac uplimit_loR, simp, simp) |
|
2017 done |
|
2018 |
|
2019 lemma Maxr_lo: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
|
2020 Maxr loR [x, y] x = lo x y" |
|
2021 apply(simp add: Maxr.simps lo.simps, auto) |
|
2022 apply(erule_tac x = xa in allE, simp, simp add: uplimit_loR) |
|
2023 done |
|
2024 |
|
2025 lemma lo_lemma': "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
|
2026 rec_exec rec_lo [x, y] = lo x y" |
|
2027 by(simp add: Maxr_lo rec_lo_Maxr_lor) |
|
2028 |
|
2029 lemma lo_lemma'': "\<lbrakk>\<not> Suc 0 < x\<rbrakk> \<Longrightarrow> rec_exec rec_lo [x, y] = lo x y" |
|
2030 apply(case_tac x, auto simp: rec_exec.simps rec_lo_def |
|
2031 Let_def lo.simps) |
|
2032 done |
|
2033 |
|
2034 lemma lo_lemma''': "\<lbrakk>\<not> Suc 0 < y\<rbrakk> \<Longrightarrow> rec_exec rec_lo [x, y] = lo x y" |
|
2035 apply(case_tac y, auto simp: rec_exec.simps rec_lo_def |
|
2036 Let_def lo.simps) |
|
2037 done |
|
2038 |
|
2039 text {* |
|
2040 The correctness of @{text "rec_lo"}: |
|
2041 *} |
|
2042 lemma lo_lemma: "rec_exec rec_lo [x, y] = lo x y" |
|
2043 apply(case_tac "Suc 0 < x \<and> Suc 0 < y") |
|
2044 apply(auto simp: lo_lemma' lo_lemma'' lo_lemma''') |
|
2045 done |
|
2046 |
|
2047 fun lgR :: "nat list \<Rightarrow> bool" |
|
2048 where |
|
2049 "lgR [x, y, u] = (y^u \<le> x)" |
|
2050 |
|
2051 text {* |
|
2052 @{text "lg"} specifies the @{text "lg"} function given on page 79 of |
|
2053 Boolos's book. It is one of the two notions of integeral logarithmatic |
|
2054 operation on that page. The other is @{text "lo"}. |
|
2055 *} |
|
2056 fun lg :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2057 where |
|
2058 "lg x y = (if x > 1 \<and> y > 1 \<and> {u. lgR [x, y, u]} \<noteq> {} then |
|
2059 Max {u. lgR [x, y, u]} |
|
2060 else 0)" |
|
2061 |
|
2062 declare lg.simps[simp del] lgR.simps[simp del] |
|
2063 |
|
2064 text {* |
|
2065 @{text "rec_lg"} is the recursive function used to implement @{text "lg"}. |
|
2066 *} |
|
2067 definition rec_lg :: "recf" |
|
2068 where |
|
2069 "rec_lg = (let rec_lgR = Cn 3 rec_le |
|
2070 [Cn 3 rec_power [id 3 1, id 3 2], id 3 0] in |
|
2071 let conR1 = Cn 2 rec_conj [Cn 2 rec_less |
|
2072 [Cn 2 (constn 1) [id 2 0], id 2 0], |
|
2073 Cn 2 rec_less [Cn 2 (constn 1) |
|
2074 [id 2 0], id 2 1]] in |
|
2075 let conR2 = Cn 2 rec_not [conR1] in |
|
2076 Cn 2 rec_add [Cn 2 rec_mult |
|
2077 [conR1, Cn 2 (rec_maxr rec_lgR) |
|
2078 [id 2 0, id 2 1, id 2 0]], |
|
2079 Cn 2 rec_mult [conR2, Cn 2 (constn 0) |
|
2080 [id 2 0]]])" |
|
2081 |
|
2082 lemma lg_maxr: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
|
2083 rec_exec rec_lg [x, y] = Maxr lgR [x, y] x" |
|
2084 proof(simp add: rec_exec.simps rec_lg_def Let_def) |
|
2085 assume h: "Suc 0 < x" "Suc 0 < y" |
|
2086 let ?rR = "(Cn 3 rec_le [Cn 3 rec_power |
|
2087 [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])" |
|
2088 have "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) |
|
2089 = Maxr ((\<lambda> args. 0 < rec_exec ?rR args)) [x, y] x" |
|
2090 proof(rule Maxr_lemma) |
|
2091 show "primerec (Cn 3 rec_le [Cn 3 rec_power |
|
2092 [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0]) (Suc (length [x, y]))" |
|
2093 apply(auto simp: numeral_3_eq_3)+ |
|
2094 done |
|
2095 qed |
|
2096 moreover have "Maxr lgR [x, y] x = Maxr ((\<lambda> args. 0 < rec_exec ?rR args)) [x, y] x" |
|
2097 apply(simp add: rec_exec.simps power_lemma) |
|
2098 apply(simp add: Maxr.simps lgR.simps) |
|
2099 done |
|
2100 ultimately show "rec_exec (rec_maxr ?rR) [x, y, x] = Maxr lgR [x, y] x" |
|
2101 by simp |
|
2102 qed |
|
2103 |
|
2104 lemma [simp]: "\<lbrakk>Suc 0 < y; lgR [x, y, xa]\<rbrakk> \<Longrightarrow> xa \<le> x" |
|
2105 apply(simp add: lgR.simps) |
|
2106 apply(subgoal_tac "y^xa > xa", simp) |
|
2107 apply(erule x_less_exp) |
|
2108 done |
|
2109 |
|
2110 lemma [simp]: "\<lbrakk>Suc 0 < x; Suc 0 < y; lgR [x, y, xa]\<rbrakk> \<Longrightarrow> |
|
2111 {u. lgR [x, y, u]} = {ya. ya \<le> x \<and> lgR [x, y, ya]}" |
|
2112 apply(rule_tac Collect_cong, auto) |
|
2113 done |
|
2114 |
|
2115 lemma maxr_lg: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> Maxr lgR [x, y] x = lg x y" |
|
2116 apply(simp add: lg.simps Maxr.simps, auto) |
|
2117 apply(erule_tac x = xa in allE, simp) |
|
2118 done |
|
2119 |
|
2120 lemma lg_lemma': "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> rec_exec rec_lg [x, y] = lg x y" |
|
2121 apply(simp add: maxr_lg lg_maxr) |
|
2122 done |
|
2123 |
|
2124 lemma lg_lemma'': "\<not> Suc 0 < x \<Longrightarrow> rec_exec rec_lg [x, y] = lg x y" |
|
2125 apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) |
|
2126 done |
|
2127 |
|
2128 lemma lg_lemma''': "\<not> Suc 0 < y \<Longrightarrow> rec_exec rec_lg [x, y] = lg x y" |
|
2129 apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) |
|
2130 done |
|
2131 |
|
2132 text {* |
|
2133 The correctness of @{text "rec_lg"}. |
|
2134 *} |
|
2135 lemma lg_lemma: "rec_exec rec_lg [x, y] = lg x y" |
|
2136 apply(case_tac "Suc 0 < x \<and> Suc 0 < y", auto simp: |
|
2137 lg_lemma' lg_lemma'' lg_lemma''') |
|
2138 done |
|
2139 |
|
2140 text {* |
|
2141 @{text "Entry sr i"} returns the @{text "i"}-th entry of a list of natural |
|
2142 numbers encoded by number @{text "sr"} using Godel's coding. |
|
2143 *} |
|
2144 fun Entry :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2145 where |
|
2146 "Entry sr i = lo sr (Pi (Suc i))" |
|
2147 |
|
2148 text {* |
|
2149 @{text "rec_entry"} is the recursive function used to implement |
|
2150 @{text "Entry"}. |
|
2151 *} |
|
2152 definition rec_entry:: "recf" |
|
2153 where |
|
2154 "rec_entry = Cn 2 rec_lo [id 2 0, Cn 2 rec_pi [Cn 2 s [id 2 1]]]" |
|
2155 |
|
2156 declare Pi.simps[simp del] |
|
2157 |
|
2158 text {* |
|
2159 The correctness of @{text "rec_entry"}. |
|
2160 *} |
|
2161 lemma entry_lemma: "rec_exec rec_entry [str, i] = Entry str i" |
|
2162 by(simp add: rec_entry_def rec_exec.simps lo_lemma pi_lemma) |
|
2163 |
|
2164 |
|
2165 subsection {* The construction of F *} |
|
2166 |
|
2167 text {* |
|
2168 Using the auxilliary functions obtained in last section, |
|
2169 we are going to contruct the function @{text "F"}, |
|
2170 which is an interpreter of Turing Machines. |
|
2171 *} |
|
2172 |
|
2173 fun listsum2 :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
|
2174 where |
|
2175 "listsum2 xs 0 = 0" |
|
2176 | "listsum2 xs (Suc n) = listsum2 xs n + xs ! n" |
|
2177 |
|
2178 fun rec_listsum2 :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
|
2179 where |
|
2180 "rec_listsum2 vl 0 = Cn vl z [id vl 0]" |
|
2181 | "rec_listsum2 vl (Suc n) = Cn vl rec_add |
|
2182 [rec_listsum2 vl n, id vl (n)]" |
|
2183 |
|
2184 declare listsum2.simps[simp del] rec_listsum2.simps[simp del] |
|
2185 |
|
2186 lemma listsum2_lemma: "\<lbrakk>length xs = vl; n \<le> vl\<rbrakk> \<Longrightarrow> |
|
2187 rec_exec (rec_listsum2 vl n) xs = listsum2 xs n" |
|
2188 apply(induct n, simp_all) |
|
2189 apply(simp_all add: rec_exec.simps rec_listsum2.simps listsum2.simps) |
|
2190 done |
|
2191 |
|
2192 fun strt' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
|
2193 where |
|
2194 "strt' xs 0 = 0" |
|
2195 | "strt' xs (Suc n) = (let dbound = listsum2 xs n + n in |
|
2196 strt' xs n + (2^(xs ! n + dbound) - 2^dbound))" |
|
2197 |
|
2198 fun rec_strt' :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
|
2199 where |
|
2200 "rec_strt' vl 0 = Cn vl z [id vl 0]" |
|
2201 | "rec_strt' vl (Suc n) = (let rec_dbound = |
|
2202 Cn vl rec_add [rec_listsum2 vl n, Cn vl (constn n) [id vl 0]] |
|
2203 in Cn vl rec_add [rec_strt' vl n, Cn vl rec_minus |
|
2204 [Cn vl rec_power [Cn vl (constn 2) [id vl 0], Cn vl rec_add |
|
2205 [id vl (n), rec_dbound]], |
|
2206 Cn vl rec_power [Cn vl (constn 2) [id vl 0], rec_dbound]]])" |
|
2207 |
|
2208 declare strt'.simps[simp del] rec_strt'.simps[simp del] |
|
2209 |
|
2210 lemma strt'_lemma: "\<lbrakk>length xs = vl; n \<le> vl\<rbrakk> \<Longrightarrow> |
|
2211 rec_exec (rec_strt' vl n) xs = strt' xs n" |
|
2212 apply(induct n) |
|
2213 apply(simp_all add: rec_exec.simps rec_strt'.simps strt'.simps |
|
2214 Let_def power_lemma listsum2_lemma) |
|
2215 done |
|
2216 |
|
2217 text {* |
|
2218 @{text "strt"} corresponds to the @{text "strt"} function on page 90 of B book, but |
|
2219 this definition generalises the original one to deal with multiple input arguments. |
|
2220 *} |
|
2221 fun strt :: "nat list \<Rightarrow> nat" |
|
2222 where |
|
2223 "strt xs = (let ys = map Suc xs in |
|
2224 strt' ys (length ys))" |
|
2225 |
|
2226 fun rec_map :: "recf \<Rightarrow> nat \<Rightarrow> recf list" |
|
2227 where |
|
2228 "rec_map rf vl = map (\<lambda> i. Cn vl rf [id vl (i)]) [0..<vl]" |
|
2229 |
|
2230 text {* |
|
2231 @{text "rec_strt"} is the recursive function used to implement @{text "strt"}. |
|
2232 *} |
|
2233 fun rec_strt :: "nat \<Rightarrow> recf" |
|
2234 where |
|
2235 "rec_strt vl = Cn vl (rec_strt' vl vl) (rec_map s vl)" |
|
2236 |
|
2237 lemma map_s_lemma: "length xs = vl \<Longrightarrow> |
|
2238 map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn vl s [recf.id vl i])) |
|
2239 [0..<vl] |
|
2240 = map Suc xs" |
|
2241 apply(induct vl arbitrary: xs, simp, auto simp: rec_exec.simps) |
|
2242 apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]", auto) |
|
2243 proof - |
|
2244 fix ys y |
|
2245 assume ind: "\<And>xs. length xs = length (ys::nat list) \<Longrightarrow> |
|
2246 map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn (length ys) s |
|
2247 [recf.id (length ys) (i)])) [0..<length ys] = map Suc xs" |
|
2248 show |
|
2249 "map ((\<lambda>a. rec_exec a (ys @ [y])) \<circ> (\<lambda>i. Cn (Suc (length ys)) s |
|
2250 [recf.id (Suc (length ys)) (i)])) [0..<length ys] = map Suc ys" |
|
2251 proof - |
|
2252 have "map ((\<lambda>a. rec_exec a ys) \<circ> (\<lambda>i. Cn (length ys) s |
|
2253 [recf.id (length ys) (i)])) [0..<length ys] = map Suc ys" |
|
2254 apply(rule_tac ind, simp) |
|
2255 done |
|
2256 moreover have |
|
2257 "map ((\<lambda>a. rec_exec a (ys @ [y])) \<circ> (\<lambda>i. Cn (Suc (length ys)) s |
|
2258 [recf.id (Suc (length ys)) (i)])) [0..<length ys] |
|
2259 = map ((\<lambda>a. rec_exec a ys) \<circ> (\<lambda>i. Cn (length ys) s |
|
2260 [recf.id (length ys) (i)])) [0..<length ys]" |
|
2261 apply(rule_tac map_ext, auto simp: rec_exec.simps nth_append) |
|
2262 done |
|
2263 ultimately show "?thesis" |
|
2264 by simp |
|
2265 qed |
|
2266 next |
|
2267 fix vl xs |
|
2268 assume "length xs = Suc vl" |
|
2269 thus "\<exists>ys y. xs = ys @ [y]" |
|
2270 apply(rule_tac x = "butlast xs" in exI, rule_tac x = "last xs" in exI) |
|
2271 apply(subgoal_tac "xs \<noteq> []", auto) |
|
2272 done |
|
2273 qed |
|
2274 |
|
2275 text {* |
|
2276 The correctness of @{text "rec_strt"}. |
|
2277 *} |
|
2278 lemma strt_lemma: "length xs = vl \<Longrightarrow> |
|
2279 rec_exec (rec_strt vl) xs = strt xs" |
|
2280 apply(simp add: strt.simps rec_exec.simps strt'_lemma) |
|
2281 apply(subgoal_tac "(map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn vl s [recf.id vl (i)])) [0..<vl]) |
|
2282 = map Suc xs", auto) |
|
2283 apply(rule map_s_lemma, simp) |
|
2284 done |
|
2285 |
|
2286 text {* |
|
2287 The @{text "scan"} function on page 90 of B book. |
|
2288 *} |
|
2289 fun scan :: "nat \<Rightarrow> nat" |
|
2290 where |
|
2291 "scan r = r mod 2" |
|
2292 |
|
2293 text {* |
|
2294 @{text "rec_scan"} is the implemention of @{text "scan"}. |
|
2295 *} |
|
2296 definition rec_scan :: "recf" |
|
2297 where "rec_scan = Cn 1 rec_mod [id 1 0, constn 2]" |
|
2298 |
|
2299 text {* |
|
2300 The correctness of @{text "scan"}. |
|
2301 *} |
|
2302 lemma scan_lemma: "rec_exec rec_scan [r] = r mod 2" |
|
2303 by(simp add: rec_exec.simps rec_scan_def mod_lemma) |
|
2304 |
|
2305 fun newleft0 :: "nat list \<Rightarrow> nat" |
|
2306 where |
|
2307 "newleft0 [p, r] = p" |
|
2308 |
|
2309 definition rec_newleft0 :: "recf" |
|
2310 where |
|
2311 "rec_newleft0 = id 2 0" |
|
2312 |
|
2313 fun newrgt0 :: "nat list \<Rightarrow> nat" |
|
2314 where |
|
2315 "newrgt0 [p, r] = r - scan r" |
|
2316 |
|
2317 definition rec_newrgt0 :: "recf" |
|
2318 where |
|
2319 "rec_newrgt0 = Cn 2 rec_minus [id 2 1, Cn 2 rec_scan [id 2 1]]" |
|
2320 |
|
2321 (*newleft1, newrgt1: left rgt number after execute on step*) |
|
2322 fun newleft1 :: "nat list \<Rightarrow> nat" |
|
2323 where |
|
2324 "newleft1 [p, r] = p" |
|
2325 |
|
2326 definition rec_newleft1 :: "recf" |
|
2327 where |
|
2328 "rec_newleft1 = id 2 0" |
|
2329 |
|
2330 fun newrgt1 :: "nat list \<Rightarrow> nat" |
|
2331 where |
|
2332 "newrgt1 [p, r] = r + 1 - scan r" |
|
2333 |
|
2334 definition rec_newrgt1 :: "recf" |
|
2335 where |
|
2336 "rec_newrgt1 = |
|
2337 Cn 2 rec_minus [Cn 2 rec_add [id 2 1, Cn 2 (constn 1) [id 2 0]], |
|
2338 Cn 2 rec_scan [id 2 1]]" |
|
2339 |
|
2340 fun newleft2 :: "nat list \<Rightarrow> nat" |
|
2341 where |
|
2342 "newleft2 [p, r] = p div 2" |
|
2343 |
|
2344 definition rec_newleft2 :: "recf" |
|
2345 where |
|
2346 "rec_newleft2 = Cn 2 rec_quo [id 2 0, Cn 2 (constn 2) [id 2 0]]" |
|
2347 |
|
2348 fun newrgt2 :: "nat list \<Rightarrow> nat" |
|
2349 where |
|
2350 "newrgt2 [p, r] = 2 * r + p mod 2" |
|
2351 |
|
2352 definition rec_newrgt2 :: "recf" |
|
2353 where |
|
2354 "rec_newrgt2 = |
|
2355 Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 1], |
|
2356 Cn 2 rec_mod [id 2 0, Cn 2 (constn 2) [id 2 0]]]" |
|
2357 |
|
2358 fun newleft3 :: "nat list \<Rightarrow> nat" |
|
2359 where |
|
2360 "newleft3 [p, r] = 2 * p + r mod 2" |
|
2361 |
|
2362 definition rec_newleft3 :: "recf" |
|
2363 where |
|
2364 "rec_newleft3 = |
|
2365 Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 0], |
|
2366 Cn 2 rec_mod [id 2 1, Cn 2 (constn 2) [id 2 0]]]" |
|
2367 |
|
2368 fun newrgt3 :: "nat list \<Rightarrow> nat" |
|
2369 where |
|
2370 "newrgt3 [p, r] = r div 2" |
|
2371 |
|
2372 definition rec_newrgt3 :: "recf" |
|
2373 where |
|
2374 "rec_newrgt3 = Cn 2 rec_quo [id 2 1, Cn 2 (constn 2) [id 2 0]]" |
|
2375 |
|
2376 text {* |
|
2377 The @{text "new_left"} function on page 91 of B book. |
|
2378 *} |
|
2379 fun newleft :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2380 where |
|
2381 "newleft p r a = (if a = 0 \<or> a = 1 then newleft0 [p, r] |
|
2382 else if a = 2 then newleft2 [p, r] |
|
2383 else if a = 3 then newleft3 [p, r] |
|
2384 else p)" |
|
2385 |
|
2386 text {* |
|
2387 @{text "rec_newleft"} is the recursive function used to |
|
2388 implement @{text "newleft"}. |
|
2389 *} |
|
2390 definition rec_newleft :: "recf" |
|
2391 where |
|
2392 "rec_newleft = |
|
2393 (let g0 = |
|
2394 Cn 3 rec_newleft0 [id 3 0, id 3 1] in |
|
2395 let g1 = Cn 3 rec_newleft2 [id 3 0, id 3 1] in |
|
2396 let g2 = Cn 3 rec_newleft3 [id 3 0, id 3 1] in |
|
2397 let g3 = id 3 0 in |
|
2398 let r0 = Cn 3 rec_disj |
|
2399 [Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]], |
|
2400 Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]]] in |
|
2401 let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in |
|
2402 let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in |
|
2403 let r3 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in |
|
2404 let gs = [g0, g1, g2, g3] in |
|
2405 let rs = [r0, r1, r2, r3] in |
|
2406 rec_embranch (zip gs rs))" |
|
2407 |
|
2408 declare newleft.simps[simp del] |
|
2409 |
|
2410 |
|
2411 lemma Suc_Suc_Suc_Suc_induct: |
|
2412 "\<lbrakk>i < Suc (Suc (Suc (Suc 0))); i = 0 \<Longrightarrow> P i; |
|
2413 i = 1 \<Longrightarrow> P i; i =2 \<Longrightarrow> P i; |
|
2414 i =3 \<Longrightarrow> P i\<rbrakk> \<Longrightarrow> P i" |
|
2415 apply(case_tac i, simp, case_tac nat, simp, |
|
2416 case_tac nata, simp, case_tac natb, simp, simp) |
|
2417 done |
|
2418 |
|
2419 declare quo_lemma2[simp] mod_lemma[simp] |
|
2420 |
|
2421 text {* |
|
2422 The correctness of @{text "rec_newleft"}. |
|
2423 *} |
|
2424 lemma newleft_lemma: |
|
2425 "rec_exec rec_newleft [p, r, a] = newleft p r a" |
|
2426 proof(simp only: rec_newleft_def Let_def) |
|
2427 let ?rgs = "[Cn 3 rec_newleft0 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newleft2 |
|
2428 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newleft3 [recf.id 3 0, recf.id 3 1], recf.id 3 0]" |
|
2429 let ?rrs = |
|
2430 "[Cn 3 rec_disj [Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 0) |
|
2431 [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 1) [recf.id 3 0]]], |
|
2432 Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 2) [recf.id 3 0]], |
|
2433 Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], |
|
2434 Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" |
|
2435 thm embranch_lemma |
|
2436 have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] |
|
2437 = Embranch (zip (map rec_exec ?rgs) (map (\<lambda>r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" |
|
2438 apply(rule_tac embranch_lemma ) |
|
2439 apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newleft0_def |
|
2440 rec_newleft1_def rec_newleft2_def rec_newleft3_def)+ |
|
2441 apply(case_tac "a = 0 \<or> a = 1", rule_tac x = 0 in exI) |
|
2442 prefer 2 |
|
2443 apply(case_tac "a = 2", rule_tac x = "Suc 0" in exI) |
|
2444 prefer 2 |
|
2445 apply(case_tac "a = 3", rule_tac x = "2" in exI) |
|
2446 prefer 2 |
|
2447 apply(case_tac "a > 3", rule_tac x = "3" in exI, auto) |
|
2448 apply(auto simp: rec_exec.simps) |
|
2449 apply(erule_tac [!] Suc_Suc_Suc_Suc_induct, auto simp: rec_exec.simps) |
|
2450 done |
|
2451 have k2: "Embranch (zip (map rec_exec ?rgs) (map (\<lambda>r args. 0 < rec_exec r args) ?rrs)) [p, r, a] = newleft p r a" |
|
2452 apply(simp add: Embranch.simps) |
|
2453 apply(simp add: rec_exec.simps) |
|
2454 apply(auto simp: newleft.simps rec_newleft0_def rec_exec.simps |
|
2455 rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
|
2456 done |
|
2457 from k1 and k2 show |
|
2458 "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = newleft p r a" |
|
2459 by simp |
|
2460 qed |
|
2461 |
|
2462 text {* |
|
2463 The @{text "newrght"} function is one similar to @{text "newleft"}, but used to |
|
2464 compute the right number. |
|
2465 *} |
|
2466 fun newrght :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2467 where |
|
2468 "newrght p r a = (if a = 0 then newrgt0 [p, r] |
|
2469 else if a = 1 then newrgt1 [p, r] |
|
2470 else if a = 2 then newrgt2 [p, r] |
|
2471 else if a = 3 then newrgt3 [p, r] |
|
2472 else r)" |
|
2473 |
|
2474 text {* |
|
2475 @{text "rec_newrght"} is the recursive function used to implement |
|
2476 @{text "newrgth"}. |
|
2477 *} |
|
2478 definition rec_newrght :: "recf" |
|
2479 where |
|
2480 "rec_newrght = |
|
2481 (let g0 = Cn 3 rec_newrgt0 [id 3 0, id 3 1] in |
|
2482 let g1 = Cn 3 rec_newrgt1 [id 3 0, id 3 1] in |
|
2483 let g2 = Cn 3 rec_newrgt2 [id 3 0, id 3 1] in |
|
2484 let g3 = Cn 3 rec_newrgt3 [id 3 0, id 3 1] in |
|
2485 let g4 = id 3 1 in |
|
2486 let r0 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]] in |
|
2487 let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]] in |
|
2488 let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in |
|
2489 let r3 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in |
|
2490 let r4 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in |
|
2491 let gs = [g0, g1, g2, g3, g4] in |
|
2492 let rs = [r0, r1, r2, r3, r4] in |
|
2493 rec_embranch (zip gs rs))" |
|
2494 declare newrght.simps[simp del] |
|
2495 |
|
2496 lemma numeral_4_eq_4: "4 = Suc 3" |
|
2497 by auto |
|
2498 |
|
2499 lemma Suc_5_induct: |
|
2500 "\<lbrakk>i < Suc (Suc (Suc (Suc (Suc 0)))); i = 0 \<Longrightarrow> P 0; |
|
2501 i = 1 \<Longrightarrow> P 1; i = 2 \<Longrightarrow> P 2; i = 3 \<Longrightarrow> P 3; i = 4 \<Longrightarrow> P 4\<rbrakk> \<Longrightarrow> P i" |
|
2502 apply(case_tac i, auto) |
|
2503 apply(case_tac nat, auto) |
|
2504 apply(case_tac nata, auto simp: numeral_2_eq_2) |
|
2505 apply(case_tac nat, auto simp: numeral_3_eq_3 numeral_4_eq_4) |
|
2506 done |
|
2507 |
|
2508 lemma [intro]: "primerec rec_scan (Suc 0)" |
|
2509 apply(auto simp: rec_scan_def, auto) |
|
2510 done |
|
2511 |
|
2512 text {* |
|
2513 The correctness of @{text "rec_newrght"}. |
|
2514 *} |
|
2515 lemma newrght_lemma: "rec_exec rec_newrght [p, r, a] = newrght p r a" |
|
2516 proof(simp only: rec_newrght_def Let_def) |
|
2517 let ?gs' = "[newrgt0, newrgt1, newrgt2, newrgt3, \<lambda> zs. zs ! 1]" |
|
2518 let ?r0 = "\<lambda> zs. zs ! 2 = 0" |
|
2519 let ?r1 = "\<lambda> zs. zs ! 2 = 1" |
|
2520 let ?r2 = "\<lambda> zs. zs ! 2 = 2" |
|
2521 let ?r3 = "\<lambda> zs. zs ! 2 = 3" |
|
2522 let ?r4 = "\<lambda> zs. zs ! 2 > 3" |
|
2523 let ?gs = "map (\<lambda> g. (\<lambda> zs. g [zs ! 0, zs ! 1])) ?gs'" |
|
2524 let ?rs = "[?r0, ?r1, ?r2, ?r3, ?r4]" |
|
2525 let ?rgs = |
|
2526 "[Cn 3 rec_newrgt0 [recf.id 3 0, recf.id 3 1], |
|
2527 Cn 3 rec_newrgt1 [recf.id 3 0, recf.id 3 1], |
|
2528 Cn 3 rec_newrgt2 [recf.id 3 0, recf.id 3 1], |
|
2529 Cn 3 rec_newrgt3 [recf.id 3 0, recf.id 3 1], recf.id 3 1]" |
|
2530 let ?rrs = |
|
2531 "[Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 0) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, |
|
2532 Cn 3 (constn 1) [recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 2) [recf.id 3 0]], |
|
2533 Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], |
|
2534 Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" |
|
2535 |
|
2536 have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] |
|
2537 = Embranch (zip (map rec_exec ?rgs) (map (\<lambda>r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" |
|
2538 apply(rule_tac embranch_lemma) |
|
2539 apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newrgt0_def |
|
2540 rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def)+ |
|
2541 apply(case_tac "a = 0", rule_tac x = 0 in exI) |
|
2542 prefer 2 |
|
2543 apply(case_tac "a = 1", rule_tac x = "Suc 0" in exI) |
|
2544 prefer 2 |
|
2545 apply(case_tac "a = 2", rule_tac x = "2" in exI) |
|
2546 prefer 2 |
|
2547 apply(case_tac "a = 3", rule_tac x = "3" in exI) |
|
2548 prefer 2 |
|
2549 apply(case_tac "a > 3", rule_tac x = "4" in exI, auto simp: rec_exec.simps) |
|
2550 apply(erule_tac [!] Suc_5_induct, auto simp: rec_exec.simps) |
|
2551 done |
|
2552 have k2: "Embranch (zip (map rec_exec ?rgs) |
|
2553 (map (\<lambda>r args. 0 < rec_exec r args) ?rrs)) [p, r, a] = newrght p r a" |
|
2554 apply(auto simp:Embranch.simps rec_exec.simps) |
|
2555 apply(auto simp: newrght.simps rec_newrgt3_def rec_newrgt2_def |
|
2556 rec_newrgt1_def rec_newrgt0_def rec_exec.simps |
|
2557 scan_lemma) |
|
2558 done |
|
2559 from k1 and k2 show |
|
2560 "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = |
|
2561 newrght p r a" by simp |
|
2562 qed |
|
2563 |
|
2564 declare Entry.simps[simp del] |
|
2565 |
|
2566 text {* |
|
2567 The @{text "actn"} function given on page 92 of B book, which is used to |
|
2568 fetch Turing Machine intructions. |
|
2569 In @{text "actn m q r"}, @{text "m"} is the Godel coding of a Turing Machine, |
|
2570 @{text "q"} is the current state of Turing Machine, @{text "r"} is the |
|
2571 right number of Turing Machine tape. |
|
2572 *} |
|
2573 fun actn :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2574 where |
|
2575 "actn m q r = (if q \<noteq> 0 then Entry m (4*(q - 1) + 2 * scan r) |
|
2576 else 4)" |
|
2577 |
|
2578 text {* |
|
2579 @{text "rec_actn"} is the recursive function used to implement @{text "actn"} |
|
2580 *} |
|
2581 definition rec_actn :: "recf" |
|
2582 where |
|
2583 "rec_actn = |
|
2584 Cn 3 rec_add [Cn 3 rec_mult |
|
2585 [Cn 3 rec_entry [id 3 0, Cn 3 rec_add [Cn 3 rec_mult |
|
2586 [Cn 3 (constn 4) [id 3 0], |
|
2587 Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], |
|
2588 Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], |
|
2589 Cn 3 rec_scan [id 3 2]]]], |
|
2590 Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], |
|
2591 Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], |
|
2592 Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " |
|
2593 |
|
2594 text {* |
|
2595 The correctness of @{text "actn"}. |
|
2596 *} |
|
2597 lemma actn_lemma: "rec_exec rec_actn [m, q, r] = actn m q r" |
|
2598 by(auto simp: rec_actn_def rec_exec.simps entry_lemma scan_lemma) |
|
2599 |
|
2600 fun newstat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2601 where |
|
2602 "newstat m q r = (if q \<noteq> 0 then Entry m (4*(q - 1) + 2*scan r + 1) |
|
2603 else 0)" |
|
2604 |
|
2605 definition rec_newstat :: "recf" |
|
2606 where |
|
2607 "rec_newstat = Cn 3 rec_add |
|
2608 [Cn 3 rec_mult [Cn 3 rec_entry [id 3 0, |
|
2609 Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], |
|
2610 Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], |
|
2611 Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], |
|
2612 Cn 3 rec_scan [id 3 2]], Cn 3 (constn 1) [id 3 0]]]], |
|
2613 Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], |
|
2614 Cn 3 rec_mult [Cn 3 (constn 0) [id 3 0], |
|
2615 Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " |
|
2616 |
|
2617 lemma newstat_lemma: "rec_exec rec_newstat [m, q, r] = newstat m q r" |
|
2618 by(auto simp: rec_exec.simps entry_lemma scan_lemma rec_newstat_def) |
|
2619 |
|
2620 declare newstat.simps[simp del] actn.simps[simp del] |
|
2621 |
|
2622 text{*code the configuration*} |
|
2623 |
|
2624 fun trpl :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2625 where |
|
2626 "trpl p q r = (Pi 0)^p * (Pi 1)^q * (Pi 2)^r" |
|
2627 |
|
2628 definition rec_trpl :: "recf" |
|
2629 where |
|
2630 "rec_trpl = Cn 3 rec_mult [Cn 3 rec_mult |
|
2631 [Cn 3 rec_power [Cn 3 (constn (Pi 0)) [id 3 0], id 3 0], |
|
2632 Cn 3 rec_power [Cn 3 (constn (Pi 1)) [id 3 0], id 3 1]], |
|
2633 Cn 3 rec_power [Cn 3 (constn (Pi 2)) [id 3 0], id 3 2]]" |
|
2634 declare trpl.simps[simp del] |
|
2635 lemma trpl_lemma: "rec_exec rec_trpl [p, q, r] = trpl p q r" |
|
2636 by(auto simp: rec_trpl_def rec_exec.simps power_lemma trpl.simps) |
|
2637 |
|
2638 text{*left, stat, rght: decode func*} |
|
2639 fun left :: "nat \<Rightarrow> nat" |
|
2640 where |
|
2641 "left c = lo c (Pi 0)" |
|
2642 |
|
2643 fun stat :: "nat \<Rightarrow> nat" |
|
2644 where |
|
2645 "stat c = lo c (Pi 1)" |
|
2646 |
|
2647 fun rght :: "nat \<Rightarrow> nat" |
|
2648 where |
|
2649 "rght c = lo c (Pi 2)" |
|
2650 |
|
2651 thm Prime.simps |
|
2652 |
|
2653 fun inpt :: "nat \<Rightarrow> nat list \<Rightarrow> nat" |
|
2654 where |
|
2655 "inpt m xs = trpl 0 1 (strt xs)" |
|
2656 |
|
2657 fun newconf :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2658 where |
|
2659 "newconf m c = trpl (newleft (left c) (rght c) |
|
2660 (actn m (stat c) (rght c))) |
|
2661 (newstat m (stat c) (rght c)) |
|
2662 (newrght (left c) (rght c) |
|
2663 (actn m (stat c) (rght c)))" |
|
2664 |
|
2665 declare left.simps[simp del] stat.simps[simp del] rght.simps[simp del] |
|
2666 inpt.simps[simp del] newconf.simps[simp del] |
|
2667 |
|
2668 definition rec_left :: "recf" |
|
2669 where |
|
2670 "rec_left = Cn 1 rec_lo [id 1 0, constn (Pi 0)]" |
|
2671 |
|
2672 definition rec_right :: "recf" |
|
2673 where |
|
2674 "rec_right = Cn 1 rec_lo [id 1 0, constn (Pi 2)]" |
|
2675 |
|
2676 definition rec_stat :: "recf" |
|
2677 where |
|
2678 "rec_stat = Cn 1 rec_lo [id 1 0, constn (Pi 1)]" |
|
2679 |
|
2680 definition rec_inpt :: "nat \<Rightarrow> recf" |
|
2681 where |
|
2682 "rec_inpt vl = Cn vl rec_trpl |
|
2683 [Cn vl (constn 0) [id vl 0], |
|
2684 Cn vl (constn 1) [id vl 0], |
|
2685 Cn vl (rec_strt (vl - 1)) |
|
2686 (map (\<lambda> i. id vl (i)) [1..<vl])]" |
|
2687 |
|
2688 lemma left_lemma: "rec_exec rec_left [c] = left c" |
|
2689 by(simp add: rec_exec.simps rec_left_def left.simps lo_lemma) |
|
2690 |
|
2691 lemma right_lemma: "rec_exec rec_right [c] = rght c" |
|
2692 by(simp add: rec_exec.simps rec_right_def rght.simps lo_lemma) |
|
2693 |
|
2694 lemma stat_lemma: "rec_exec rec_stat [c] = stat c" |
|
2695 by(simp add: rec_exec.simps rec_stat_def stat.simps lo_lemma) |
|
2696 |
|
2697 declare rec_strt.simps[simp del] strt.simps[simp del] |
|
2698 |
|
2699 lemma map_cons_eq: |
|
2700 "(map ((\<lambda>a. rec_exec a (m # xs)) \<circ> |
|
2701 (\<lambda>i. recf.id (Suc (length xs)) (i))) |
|
2702 [Suc 0..<Suc (length xs)]) |
|
2703 = map (\<lambda> i. xs ! (i - 1)) [Suc 0..<Suc (length xs)]" |
|
2704 apply(rule map_ext, auto) |
|
2705 apply(auto simp: rec_exec.simps nth_append nth_Cons split: nat.split) |
|
2706 done |
|
2707 |
|
2708 lemma list_map_eq: |
|
2709 "vl = length (xs::nat list) \<Longrightarrow> map (\<lambda> i. xs ! (i - 1)) |
|
2710 [Suc 0..<Suc vl] = xs" |
|
2711 apply(induct vl arbitrary: xs, simp) |
|
2712 apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]", auto) |
|
2713 proof - |
|
2714 fix ys y |
|
2715 assume ind: |
|
2716 "\<And>xs. length (ys::nat list) = length (xs::nat list) \<Longrightarrow> |
|
2717 map (\<lambda>i. xs ! (i - Suc 0)) [Suc 0..<length xs] @ |
|
2718 [xs ! (length xs - Suc 0)] = xs" |
|
2719 and h: "Suc 0 \<le> length (ys::nat list)" |
|
2720 have "map (\<lambda>i. ys ! (i - Suc 0)) [Suc 0..<length ys] @ |
|
2721 [ys ! (length ys - Suc 0)] = ys" |
|
2722 apply(rule_tac ind, simp) |
|
2723 done |
|
2724 moreover have |
|
2725 "map (\<lambda>i. (ys @ [y]) ! (i - Suc 0)) [Suc 0..<length ys] |
|
2726 = map (\<lambda>i. ys ! (i - Suc 0)) [Suc 0..<length ys]" |
|
2727 apply(rule map_ext) |
|
2728 using h |
|
2729 apply(auto simp: nth_append) |
|
2730 done |
|
2731 ultimately show "map (\<lambda>i. (ys @ [y]) ! (i - Suc 0)) |
|
2732 [Suc 0..<length ys] @ [(ys @ [y]) ! (length ys - Suc 0)] = ys" |
|
2733 apply(simp del: map_eq_conv add: nth_append, auto) |
|
2734 using h |
|
2735 apply(simp) |
|
2736 done |
|
2737 next |
|
2738 fix vl xs |
|
2739 assume "Suc vl = length (xs::nat list)" |
|
2740 thus "\<exists>ys y. xs = ys @ [y]" |
|
2741 apply(rule_tac x = "butlast xs" in exI, |
|
2742 rule_tac x = "last xs" in exI) |
|
2743 apply(case_tac "xs \<noteq> []", auto) |
|
2744 done |
|
2745 qed |
|
2746 |
|
2747 lemma [elim]: |
|
2748 "Suc 0 \<le> length xs \<Longrightarrow> |
|
2749 (map ((\<lambda>a. rec_exec a (m # xs)) \<circ> |
|
2750 (\<lambda>i. recf.id (Suc (length xs)) (i))) |
|
2751 [Suc 0..<length xs] @ [(m # xs) ! length xs]) = xs" |
|
2752 using map_cons_eq[of m xs] |
|
2753 apply(simp del: map_eq_conv add: rec_exec.simps) |
|
2754 using list_map_eq[of "length xs" xs] |
|
2755 apply(simp) |
|
2756 done |
|
2757 |
|
2758 |
|
2759 lemma inpt_lemma: |
|
2760 "\<lbrakk>Suc (length xs) = vl\<rbrakk> \<Longrightarrow> |
|
2761 rec_exec (rec_inpt vl) (m # xs) = inpt m xs" |
|
2762 apply(auto simp: rec_exec.simps rec_inpt_def |
|
2763 trpl_lemma inpt.simps strt_lemma) |
|
2764 apply(subgoal_tac |
|
2765 "(map ((\<lambda>a. rec_exec a (m # xs)) \<circ> |
|
2766 (\<lambda>i. recf.id (Suc (length xs)) (i))) |
|
2767 [Suc 0..<length xs] @ [(m # xs) ! length xs]) = xs", simp) |
|
2768 apply(auto, case_tac xs, auto) |
|
2769 done |
|
2770 |
|
2771 definition rec_newconf:: "recf" |
|
2772 where |
|
2773 "rec_newconf = |
|
2774 Cn 2 rec_trpl |
|
2775 [Cn 2 rec_newleft [Cn 2 rec_left [id 2 1], |
|
2776 Cn 2 rec_right [id 2 1], |
|
2777 Cn 2 rec_actn [id 2 0, |
|
2778 Cn 2 rec_stat [id 2 1], |
|
2779 Cn 2 rec_right [id 2 1]]], |
|
2780 Cn 2 rec_newstat [id 2 0, |
|
2781 Cn 2 rec_stat [id 2 1], |
|
2782 Cn 2 rec_right [id 2 1]], |
|
2783 Cn 2 rec_newrght [Cn 2 rec_left [id 2 1], |
|
2784 Cn 2 rec_right [id 2 1], |
|
2785 Cn 2 rec_actn [id 2 0, |
|
2786 Cn 2 rec_stat [id 2 1], |
|
2787 Cn 2 rec_right [id 2 1]]]]" |
|
2788 |
|
2789 lemma newconf_lemma: "rec_exec rec_newconf [m ,c] = newconf m c" |
|
2790 by(auto simp: rec_newconf_def rec_exec.simps |
|
2791 trpl_lemma newleft_lemma left_lemma |
|
2792 right_lemma stat_lemma newrght_lemma actn_lemma |
|
2793 newstat_lemma stat_lemma newconf.simps) |
|
2794 |
|
2795 declare newconf_lemma[simp] |
|
2796 |
|
2797 text {* |
|
2798 @{text "conf m r k"} computes the TM configuration after @{text "k"} steps of execution |
|
2799 of TM coded as @{text "m"} starting from the initial configuration where the left number equals @{text "0"}, |
|
2800 right number equals @{text "r"}. |
|
2801 *} |
|
2802 fun conf :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2803 where |
|
2804 "conf m r 0 = trpl 0 (Suc 0) r" |
|
2805 | "conf m r (Suc t) = newconf m (conf m r t)" |
|
2806 |
|
2807 declare conf.simps[simp del] |
|
2808 |
|
2809 text {* |
|
2810 @{text "conf"} is implemented by the following recursive function @{text "rec_conf"}. |
|
2811 *} |
|
2812 definition rec_conf :: "recf" |
|
2813 where |
|
2814 "rec_conf = Pr 2 (Cn 2 rec_trpl [Cn 2 (constn 0) [id 2 0], Cn 2 (constn (Suc 0)) [id 2 0], id 2 1]) |
|
2815 (Cn 4 rec_newconf [id 4 0, id 4 3])" |
|
2816 |
|
2817 lemma conf_step: |
|
2818 "rec_exec rec_conf [m, r, Suc t] = |
|
2819 rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" |
|
2820 proof - |
|
2821 have "rec_exec rec_conf ([m, r] @ [Suc t]) = |
|
2822 rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" |
|
2823 by(simp only: rec_conf_def rec_pr_Suc_simp_rewrite, |
|
2824 simp add: rec_exec.simps) |
|
2825 thus "rec_exec rec_conf [m, r, Suc t] = |
|
2826 rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" |
|
2827 by simp |
|
2828 qed |
|
2829 |
|
2830 text {* |
|
2831 The correctness of @{text "rec_conf"}. |
|
2832 *} |
|
2833 lemma conf_lemma: |
|
2834 "rec_exec rec_conf [m, r, t] = conf m r t" |
|
2835 apply(induct t) |
|
2836 apply(simp add: rec_conf_def rec_exec.simps conf.simps inpt_lemma trpl_lemma) |
|
2837 apply(simp add: conf_step conf.simps) |
|
2838 done |
|
2839 |
|
2840 text {* |
|
2841 @{text "NSTD c"} returns true if the configureation coded by @{text "c"} is no a stardard |
|
2842 final configuration. |
|
2843 *} |
|
2844 fun NSTD :: "nat \<Rightarrow> bool" |
|
2845 where |
|
2846 "NSTD c = (stat c \<noteq> 0 \<or> left c \<noteq> 0 \<or> |
|
2847 rght c \<noteq> 2^(lg (rght c + 1) 2) - 1 \<or> rght c = 0)" |
|
2848 |
|
2849 text {* |
|
2850 @{text "rec_NSTD"} is the recursive function implementing @{text "NSTD"}. |
|
2851 *} |
|
2852 definition rec_NSTD :: "recf" |
|
2853 where |
|
2854 "rec_NSTD = |
|
2855 Cn 1 rec_disj [ |
|
2856 Cn 1 rec_disj [ |
|
2857 Cn 1 rec_disj |
|
2858 [Cn 1 rec_noteq [rec_stat, constn 0], |
|
2859 Cn 1 rec_noteq [rec_left, constn 0]] , |
|
2860 Cn 1 rec_noteq [rec_right, |
|
2861 Cn 1 rec_minus [Cn 1 rec_power |
|
2862 [constn 2, Cn 1 rec_lg |
|
2863 [Cn 1 rec_add |
|
2864 [rec_right, constn 1], |
|
2865 constn 2]], constn 1]]], |
|
2866 Cn 1 rec_eq [rec_right, constn 0]]" |
|
2867 |
|
2868 lemma NSTD_lemma1: "rec_exec rec_NSTD [c] = Suc 0 \<or> |
|
2869 rec_exec rec_NSTD [c] = 0" |
|
2870 by(simp add: rec_exec.simps rec_NSTD_def) |
|
2871 |
|
2872 declare NSTD.simps[simp del] |
|
2873 lemma NSTD_lemma2': "(rec_exec rec_NSTD [c] = Suc 0) \<Longrightarrow> NSTD c" |
|
2874 apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma left_lemma |
|
2875 lg_lemma right_lemma power_lemma NSTD.simps eq_lemma) |
|
2876 apply(auto) |
|
2877 apply(case_tac "0 < left c", simp, simp) |
|
2878 done |
|
2879 |
|
2880 lemma NSTD_lemma2'': |
|
2881 "NSTD c \<Longrightarrow> (rec_exec rec_NSTD [c] = Suc 0)" |
|
2882 apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma |
|
2883 left_lemma lg_lemma right_lemma power_lemma NSTD.simps) |
|
2884 apply(auto split: if_splits) |
|
2885 done |
|
2886 |
|
2887 text {* |
|
2888 The correctness of @{text "NSTD"}. |
|
2889 *} |
|
2890 lemma NSTD_lemma2: "(rec_exec rec_NSTD [c] = Suc 0) = NSTD c" |
|
2891 using NSTD_lemma1 |
|
2892 apply(auto intro: NSTD_lemma2' NSTD_lemma2'') |
|
2893 done |
|
2894 |
|
2895 fun nstd :: "nat \<Rightarrow> nat" |
|
2896 where |
|
2897 "nstd c = (if NSTD c then 1 else 0)" |
|
2898 |
|
2899 lemma nstd_lemma: "rec_exec rec_NSTD [c] = nstd c" |
|
2900 using NSTD_lemma1 |
|
2901 apply(simp add: NSTD_lemma2, auto) |
|
2902 done |
|
2903 |
|
2904 text{* |
|
2905 @{text "nonstep m r t"} means afer @{text "t"} steps of execution, the TM coded by @{text "m"} |
|
2906 is not at a stardard final configuration. |
|
2907 *} |
|
2908 fun nonstop :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2909 where |
|
2910 "nonstop m r t = nstd (conf m r t)" |
|
2911 |
|
2912 text {* |
|
2913 @{text "rec_nonstop"} is the recursive function implementing @{text "nonstop"}. |
|
2914 *} |
|
2915 definition rec_nonstop :: "recf" |
|
2916 where |
|
2917 "rec_nonstop = Cn 3 rec_NSTD [rec_conf]" |
|
2918 |
|
2919 text {* |
|
2920 The correctness of @{text "rec_nonstop"}. |
|
2921 *} |
|
2922 lemma nonstop_lemma: |
|
2923 "rec_exec rec_nonstop [m, r, t] = nonstop m r t" |
|
2924 apply(simp add: rec_exec.simps rec_nonstop_def nstd_lemma conf_lemma) |
|
2925 done |
|
2926 |
|
2927 text{* |
|
2928 @{text "rec_halt"} is the recursive function calculating the steps a TM needs to execute before |
|
2929 to reach a stardard final configuration. This recursive function is the only one |
|
2930 using @{text "Mn"} combinator. So it is the only non-primitive recursive function |
|
2931 needs to be used in the construction of the universal function @{text "F"}. |
|
2932 *} |
|
2933 |
|
2934 definition rec_halt :: "recf" |
|
2935 where |
|
2936 "rec_halt = Mn (Suc (Suc 0)) (rec_nonstop)" |
|
2937 |
|
2938 declare nonstop.simps[simp del] |
|
2939 |
|
2940 lemma primerec_not0: "primerec f n \<Longrightarrow> n > 0" |
|
2941 by(induct f n rule: primerec.induct, auto) |
|
2942 |
|
2943 lemma [elim]: "primerec f 0 \<Longrightarrow> RR" |
|
2944 apply(drule_tac primerec_not0, simp) |
|
2945 done |
|
2946 |
|
2947 lemma [simp]: "length xs = Suc n \<Longrightarrow> length (butlast xs) = n" |
|
2948 apply(subgoal_tac "\<exists> y ys. xs = ys @ [y]", auto) |
|
2949 apply(rule_tac x = "last xs" in exI) |
|
2950 apply(rule_tac x = "butlast xs" in exI) |
|
2951 apply(case_tac "xs = []", auto) |
|
2952 done |
|
2953 |
|
2954 text {* |
|
2955 The lemma relates the interpreter of primitive fucntions with |
|
2956 the calculation relation of general recursive functions. |
|
2957 *} |
|
2958 lemma prime_rel_exec_eq: "primerec r (length xs) |
|
2959 \<Longrightarrow> rec_calc_rel r xs rs = (rec_exec r xs = rs)" |
|
2960 proof(induct r xs arbitrary: rs rule: rec_exec.induct, simp_all) |
|
2961 fix xs rs |
|
2962 assume "primerec z (length (xs::nat list))" |
|
2963 hence "length xs = Suc 0" by(erule_tac prime_z_reverse, simp) |
|
2964 thus "rec_calc_rel z xs rs = (rec_exec z xs = rs)" |
|
2965 apply(case_tac xs, simp, auto) |
|
2966 apply(erule_tac calc_z_reverse, simp add: rec_exec.simps) |
|
2967 apply(simp add: rec_exec.simps, rule_tac calc_z) |
|
2968 done |
|
2969 next |
|
2970 fix xs rs |
|
2971 assume "primerec s (length (xs::nat list))" |
|
2972 hence "length xs = Suc 0" .. |
|
2973 thus "rec_calc_rel s xs rs = (rec_exec s xs = rs)" |
|
2974 by(case_tac xs, auto simp: rec_exec.simps intro: calc_s |
|
2975 elim: calc_s_reverse) |
|
2976 next |
|
2977 fix m n xs rs |
|
2978 assume "primerec (recf.id m n) (length (xs::nat list))" |
|
2979 thus |
|
2980 "rec_calc_rel (recf.id m n) xs rs = |
|
2981 (rec_exec (recf.id m n) xs = rs)" |
|
2982 apply(erule_tac prime_id_reverse) |
|
2983 apply(simp add: rec_exec.simps, auto) |
|
2984 apply(erule_tac calc_id_reverse, simp) |
|
2985 apply(rule_tac calc_id, auto) |
|
2986 done |
|
2987 next |
|
2988 fix n f gs xs rs |
|
2989 assume ind1: |
|
2990 "\<And>x rs. \<lbrakk>x \<in> set gs; primerec x (length xs)\<rbrakk> \<Longrightarrow> |
|
2991 rec_calc_rel x xs rs = (rec_exec x xs = rs)" |
|
2992 and ind2: |
|
2993 "\<And>x rs. \<lbrakk>x = map (\<lambda>a. rec_exec a xs) gs; |
|
2994 primerec f (length gs)\<rbrakk> \<Longrightarrow> |
|
2995 rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) rs = |
|
2996 (rec_exec f (map (\<lambda>a. rec_exec a xs) gs) = rs)" |
|
2997 and h: "primerec (Cn n f gs) (length xs)" |
|
2998 show "rec_calc_rel (Cn n f gs) xs rs = |
|
2999 (rec_exec (Cn n f gs) xs = rs)" |
|
3000 proof(auto simp: rec_exec.simps, erule_tac calc_cn_reverse, auto) |
|
3001 fix ys |
|
3002 assume g1:"\<forall>k<length gs. rec_calc_rel (gs ! k) xs (ys ! k)" |
|
3003 and g2: "length ys = length gs" |
|
3004 and g3: "rec_calc_rel f ys rs" |
|
3005 have "rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) rs = |
|
3006 (rec_exec f (map (\<lambda>a. rec_exec a xs) gs) = rs)" |
|
3007 apply(rule_tac ind2, auto) |
|
3008 using h |
|
3009 apply(erule_tac prime_cn_reverse, simp) |
|
3010 done |
|
3011 moreover have "ys = (map (\<lambda>a. rec_exec a xs) gs)" |
|
3012 proof(rule_tac nth_equalityI, auto simp: g2) |
|
3013 fix i |
|
3014 assume "i < length gs" thus "ys ! i = rec_exec (gs!i) xs" |
|
3015 using ind1[of "gs ! i" "ys ! i"] g1 h |
|
3016 apply(erule_tac prime_cn_reverse, simp) |
|
3017 done |
|
3018 qed |
|
3019 ultimately show "rec_exec f (map (\<lambda>a. rec_exec a xs) gs) = rs" |
|
3020 using g3 |
|
3021 by(simp) |
|
3022 next |
|
3023 from h show |
|
3024 "rec_calc_rel (Cn n f gs) xs |
|
3025 (rec_exec f (map (\<lambda>a. rec_exec a xs) gs))" |
|
3026 apply(rule_tac rs = "(map (\<lambda>a. rec_exec a xs) gs)" in calc_cn, |
|
3027 auto) |
|
3028 apply(erule_tac [!] prime_cn_reverse, auto) |
|
3029 proof - |
|
3030 fix k |
|
3031 assume "k < length gs" "primerec f (length gs)" |
|
3032 "\<forall>i<length gs. primerec (gs ! i) (length xs)" |
|
3033 thus "rec_calc_rel (gs ! k) xs (rec_exec (gs ! k) xs)" |
|
3034 using ind1[of "gs!k" "(rec_exec (gs ! k) xs)"] |
|
3035 by(simp) |
|
3036 next |
|
3037 assume "primerec f (length gs)" |
|
3038 "\<forall>i<length gs. primerec (gs ! i) (length xs)" |
|
3039 thus "rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) |
|
3040 (rec_exec f (map (\<lambda>a. rec_exec a xs) gs))" |
|
3041 using ind2[of "(map (\<lambda>a. rec_exec a xs) gs)" |
|
3042 "(rec_exec f (map (\<lambda>a. rec_exec a xs) gs))"] |
|
3043 by simp |
|
3044 qed |
|
3045 qed |
|
3046 next |
|
3047 fix n f g xs rs |
|
3048 assume ind1: |
|
3049 "\<And>rs. \<lbrakk>last xs = 0; primerec f (length xs - Suc 0)\<rbrakk> |
|
3050 \<Longrightarrow> rec_calc_rel f (butlast xs) rs = |
|
3051 (rec_exec f (butlast xs) = rs)" |
|
3052 and ind2 : |
|
3053 "\<And>rs. \<lbrakk>0 < last xs; |
|
3054 primerec (Pr n f g) (Suc (length xs - Suc 0))\<rbrakk> \<Longrightarrow> |
|
3055 rec_calc_rel (Pr n f g) (butlast xs @ [last xs - Suc 0]) rs |
|
3056 = (rec_exec (Pr n f g) (butlast xs @ [last xs - Suc 0]) = rs)" |
|
3057 and ind3: |
|
3058 "\<And>rs. \<lbrakk>0 < last xs; primerec g (Suc (Suc (length xs - Suc 0)))\<rbrakk> |
|
3059 \<Longrightarrow> rec_calc_rel g (butlast xs @ |
|
3060 [last xs - Suc 0, rec_exec (Pr n f g) |
|
3061 (butlast xs @ [last xs - Suc 0])]) rs = |
|
3062 (rec_exec g (butlast xs @ [last xs - Suc 0, |
|
3063 rec_exec (Pr n f g) |
|
3064 (butlast xs @ [last xs - Suc 0])]) = rs)" |
|
3065 and h: "primerec (Pr n f g) (length (xs::nat list))" |
|
3066 show "rec_calc_rel (Pr n f g) xs rs = (rec_exec (Pr n f g) xs = rs)" |
|
3067 proof(auto) |
|
3068 assume "rec_calc_rel (Pr n f g) xs rs" |
|
3069 thus "rec_exec (Pr n f g) xs = rs" |
|
3070 proof(erule_tac calc_pr_reverse) |
|
3071 fix l |
|
3072 assume g: "xs = l @ [0]" |
|
3073 "rec_calc_rel f l rs" |
|
3074 "n = length l" |
|
3075 thus "rec_exec (Pr n f g) xs = rs" |
|
3076 using ind1[of rs] h |
|
3077 apply(simp add: rec_exec.simps, |
|
3078 erule_tac prime_pr_reverse, simp) |
|
3079 done |
|
3080 next |
|
3081 fix l y ry |
|
3082 assume d:"xs = l @ [Suc y]" |
|
3083 "rec_calc_rel (Pr (length l) f g) (l @ [y]) ry" |
|
3084 "n = length l" |
|
3085 "rec_calc_rel g (l @ [y, ry]) rs" |
|
3086 moreover hence "primerec g (Suc (Suc n))" using h |
|
3087 proof(erule_tac prime_pr_reverse) |
|
3088 assume "primerec g (Suc (Suc n))" "length xs = Suc n" |
|
3089 thus "?thesis" by simp |
|
3090 qed |
|
3091 ultimately show "rec_exec (Pr n f g) xs = rs" |
|
3092 apply(simp) |
|
3093 using ind3[of rs] |
|
3094 apply(simp add: rec_pr_Suc_simp_rewrite) |
|
3095 using ind2[of ry] h |
|
3096 apply(simp) |
|
3097 done |
|
3098 qed |
|
3099 next |
|
3100 show "rec_calc_rel (Pr n f g) xs (rec_exec (Pr n f g) xs)" |
|
3101 proof - |
|
3102 have "rec_calc_rel (Pr n f g) (butlast xs @ [last xs]) |
|
3103 (rec_exec (Pr n f g) (butlast xs @ [last xs]))" |
|
3104 using h |
|
3105 apply(erule_tac prime_pr_reverse, simp) |
|
3106 apply(case_tac "last xs", simp) |
|
3107 apply(rule_tac calc_pr_zero, simp) |
|
3108 using ind1[of "rec_exec (Pr n f g) (butlast xs @ [0])"] |
|
3109 apply(simp add: rec_exec.simps, simp, simp, simp) |
|
3110 thm calc_pr_ind |
|
3111 apply(rule_tac rk = "rec_exec (Pr n f g) |
|
3112 (butlast xs@[last xs - Suc 0])" in calc_pr_ind) |
|
3113 using ind2[of "rec_exec (Pr n f g) |
|
3114 (butlast xs @ [last xs - Suc 0])"] h |
|
3115 apply(simp, simp, simp) |
|
3116 proof - |
|
3117 fix nat |
|
3118 assume "length xs = Suc n" |
|
3119 "primerec g (Suc (Suc n))" |
|
3120 "last xs = Suc nat" |
|
3121 thus |
|
3122 "rec_calc_rel g (butlast xs @ [nat, rec_exec (Pr n f g) |
|
3123 (butlast xs @ [nat])]) (rec_exec (Pr n f g) (butlast xs @ [Suc nat]))" |
|
3124 using ind3[of "rec_exec (Pr n f g) |
|
3125 (butlast xs @ [Suc nat])"] |
|
3126 apply(simp add: rec_exec.simps) |
|
3127 done |
|
3128 qed |
|
3129 thus "rec_calc_rel (Pr n f g) xs (rec_exec (Pr n f g) xs)" |
|
3130 using h |
|
3131 apply(erule_tac prime_pr_reverse, simp) |
|
3132 apply(subgoal_tac "butlast xs @ [last xs] = xs", simp) |
|
3133 apply(case_tac xs, simp, simp) |
|
3134 done |
|
3135 qed |
|
3136 qed |
|
3137 next |
|
3138 fix n f xs rs |
|
3139 assume "primerec (Mn n f) (length (xs::nat list))" |
|
3140 thus "rec_calc_rel (Mn n f) xs rs = (rec_exec (Mn n f) xs = rs)" |
|
3141 by(erule_tac prime_mn_reverse) |
|
3142 qed |
|
3143 |
|
3144 declare numeral_2_eq_2[simp] numeral_3_eq_3[simp] |
|
3145 |
|
3146 lemma [intro]: "primerec rec_right (Suc 0)" |
|
3147 apply(simp add: rec_right_def rec_lo_def Let_def) |
|
3148 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3149 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3150 done |
|
3151 |
|
3152 lemma [simp]: |
|
3153 "rec_calc_rel rec_right [r] rs = (rec_exec rec_right [r] = rs)" |
|
3154 apply(rule_tac prime_rel_exec_eq, auto) |
|
3155 done |
|
3156 |
|
3157 lemma [intro]: "primerec rec_pi (Suc 0)" |
|
3158 apply(simp add: rec_pi_def rec_dummy_pi_def |
|
3159 rec_np_def rec_fac_def rec_prime_def |
|
3160 rec_Minr.simps Let_def get_fstn_args.simps |
|
3161 arity.simps |
|
3162 rec_all.simps rec_sigma.simps rec_accum.simps) |
|
3163 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3164 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3165 apply(simp add: rec_dummyfac_def) |
|
3166 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3167 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3168 done |
|
3169 |
|
3170 lemma [intro]: "primerec rec_trpl (Suc (Suc (Suc 0)))" |
|
3171 apply(simp add: rec_trpl_def) |
|
3172 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3173 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3174 done |
|
3175 |
|
3176 lemma [intro!]: "\<lbrakk>0 < vl; n \<le> vl\<rbrakk> \<Longrightarrow> primerec (rec_listsum2 vl n) vl" |
|
3177 apply(induct n) |
|
3178 apply(simp_all add: rec_strt'.simps Let_def rec_listsum2.simps) |
|
3179 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3180 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3181 done |
|
3182 |
|
3183 lemma [elim]: "\<lbrakk>0 < vl; n \<le> vl\<rbrakk> \<Longrightarrow> primerec (rec_strt' vl n) vl" |
|
3184 apply(induct n) |
|
3185 apply(simp_all add: rec_strt'.simps Let_def) |
|
3186 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3187 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+) |
|
3188 done |
|
3189 |
|
3190 lemma [elim]: "vl > 0 \<Longrightarrow> primerec (rec_strt vl) vl" |
|
3191 apply(simp add: rec_strt.simps rec_strt'.simps) |
|
3192 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3193 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3194 done |
|
3195 |
|
3196 lemma [elim]: |
|
3197 "i < vl \<Longrightarrow> primerec ((map (\<lambda>i. recf.id (Suc vl) (i)) |
|
3198 [Suc 0..<vl] @ [recf.id (Suc vl) (vl)]) ! i) (Suc vl)" |
|
3199 apply(induct i, auto simp: nth_append) |
|
3200 done |
|
3201 |
|
3202 lemma [intro]: "primerec rec_newleft0 ((Suc (Suc 0)))" |
|
3203 apply(simp add: rec_newleft_def rec_embranch.simps |
|
3204 Let_def arity.simps rec_newleft0_def |
|
3205 rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
|
3206 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3207 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3208 done |
|
3209 |
|
3210 lemma [intro]: "primerec rec_newleft1 ((Suc (Suc 0)))" |
|
3211 apply(simp add: rec_newleft_def rec_embranch.simps |
|
3212 Let_def arity.simps rec_newleft0_def |
|
3213 rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
|
3214 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3215 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3216 done |
|
3217 |
|
3218 lemma [intro]: "primerec rec_newleft2 ((Suc (Suc 0)))" |
|
3219 apply(simp add: rec_newleft_def rec_embranch.simps |
|
3220 Let_def arity.simps rec_newleft0_def |
|
3221 rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
|
3222 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3223 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3224 done |
|
3225 |
|
3226 lemma [intro]: "primerec rec_newleft3 ((Suc (Suc 0)))" |
|
3227 apply(simp add: rec_newleft_def rec_embranch.simps |
|
3228 Let_def arity.simps rec_newleft0_def |
|
3229 rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
|
3230 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3231 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3232 done |
|
3233 |
|
3234 lemma [intro]: "primerec rec_newleft (Suc (Suc (Suc 0)))" |
|
3235 apply(simp add: rec_newleft_def rec_embranch.simps |
|
3236 Let_def arity.simps) |
|
3237 apply(rule_tac prime_cn, auto+) |
|
3238 done |
|
3239 |
|
3240 lemma [intro]: "primerec rec_left (Suc 0)" |
|
3241 apply(simp add: rec_left_def rec_lo_def rec_entry_def Let_def) |
|
3242 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3243 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3244 done |
|
3245 |
|
3246 lemma [intro]: "primerec rec_actn (Suc (Suc (Suc 0)))" |
|
3247 apply(simp add: rec_left_def rec_lo_def rec_entry_def |
|
3248 Let_def rec_actn_def) |
|
3249 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3250 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3251 done |
|
3252 |
|
3253 lemma [intro]: "primerec rec_stat (Suc 0)" |
|
3254 apply(simp add: rec_left_def rec_lo_def rec_entry_def Let_def |
|
3255 rec_actn_def rec_stat_def) |
|
3256 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3257 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3258 done |
|
3259 |
|
3260 lemma [intro]: "primerec rec_newstat (Suc (Suc (Suc 0)))" |
|
3261 apply(simp add: rec_left_def rec_lo_def rec_entry_def |
|
3262 Let_def rec_actn_def rec_stat_def rec_newstat_def) |
|
3263 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3264 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3265 done |
|
3266 |
|
3267 lemma [intro]: "primerec rec_newrght (Suc (Suc (Suc 0)))" |
|
3268 apply(simp add: rec_newrght_def rec_embranch.simps |
|
3269 Let_def arity.simps rec_newrgt0_def |
|
3270 rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def) |
|
3271 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3272 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3273 done |
|
3274 |
|
3275 lemma [intro]: "primerec rec_newconf (Suc (Suc 0))" |
|
3276 apply(simp add: rec_newconf_def) |
|
3277 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3278 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3279 done |
|
3280 |
|
3281 lemma [intro]: "0 < vl \<Longrightarrow> primerec (rec_inpt (Suc vl)) (Suc vl)" |
|
3282 apply(simp add: rec_inpt_def) |
|
3283 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3284 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3285 done |
|
3286 |
|
3287 lemma [intro]: "primerec rec_conf (Suc (Suc (Suc 0)))" |
|
3288 apply(simp add: rec_conf_def) |
|
3289 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3290 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3291 apply(auto simp: numeral_4_eq_4) |
|
3292 done |
|
3293 |
|
3294 lemma [simp]: |
|
3295 "rec_calc_rel rec_conf [m, r, t] rs = |
|
3296 (rec_exec rec_conf [m, r, t] = rs)" |
|
3297 apply(rule_tac prime_rel_exec_eq, auto) |
|
3298 done |
|
3299 |
|
3300 lemma [intro]: "primerec rec_lg (Suc (Suc 0))" |
|
3301 apply(simp add: rec_lg_def Let_def) |
|
3302 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3303 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3304 done |
|
3305 |
|
3306 lemma [intro]: "primerec rec_nonstop (Suc (Suc (Suc 0)))" |
|
3307 apply(simp add: rec_nonstop_def rec_NSTD_def rec_stat_def |
|
3308 rec_lo_def Let_def rec_left_def rec_right_def rec_newconf_def |
|
3309 rec_newstat_def) |
|
3310 apply(tactic {* resolve_tac [@{thm prime_cn}, |
|
3311 @{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
|
3312 done |
|
3313 |
|
3314 lemma nonstop_eq[simp]: |
|
3315 "rec_calc_rel rec_nonstop [m, r, t] rs = |
|
3316 (rec_exec rec_nonstop [m, r, t] = rs)" |
|
3317 apply(rule prime_rel_exec_eq, auto) |
|
3318 done |
|
3319 |
|
3320 lemma halt_lemma': |
|
3321 "rec_calc_rel rec_halt [m, r] t = |
|
3322 (rec_calc_rel rec_nonstop [m, r, t] 0 \<and> |
|
3323 (\<forall> t'< t. |
|
3324 (\<exists> y. rec_calc_rel rec_nonstop [m, r, t'] y \<and> |
|
3325 y \<noteq> 0)))" |
|
3326 apply(auto simp: rec_halt_def) |
|
3327 apply(erule calc_mn_reverse, simp) |
|
3328 apply(erule_tac calc_mn_reverse) |
|
3329 apply(erule_tac x = t' in allE, simp) |
|
3330 apply(rule_tac calc_mn, simp_all) |
|
3331 done |
|
3332 |
|
3333 text {* |
|
3334 The following lemma gives the correctness of @{text "rec_halt"}. |
|
3335 It says: if @{text "rec_halt"} calculates that the TM coded by @{text "m"} |
|
3336 will reach a standard final configuration after @{text "t"} steps of execution, then it is indeed so. |
|
3337 *} |
|
3338 lemma halt_lemma: |
|
3339 "rec_calc_rel (rec_halt) [m, r] t = |
|
3340 (rec_exec rec_nonstop [m, r, t] = 0 \<and> |
|
3341 (\<forall> t'< t. (\<exists> y. rec_exec rec_nonstop [m, r, t'] = y |
|
3342 \<and> y \<noteq> 0)))" |
|
3343 using halt_lemma'[of m r t] |
|
3344 by simp |
|
3345 |
|
3346 text {*F: universal machine*} |
|
3347 |
|
3348 text {* |
|
3349 @{text "valu r"} extracts computing result out of the right number @{text "r"}. |
|
3350 *} |
|
3351 fun valu :: "nat \<Rightarrow> nat" |
|
3352 where |
|
3353 "valu r = (lg (r + 1) 2) - 1" |
|
3354 |
|
3355 text {* |
|
3356 @{text "rec_valu"} is the recursive function implementing @{text "valu"}. |
|
3357 *} |
|
3358 definition rec_valu :: "recf" |
|
3359 where |
|
3360 "rec_valu = Cn 1 rec_minus [Cn 1 rec_lg [s, constn 2], constn 1]" |
|
3361 |
|
3362 text {* |
|
3363 The correctness of @{text "rec_valu"}. |
|
3364 *} |
|
3365 lemma value_lemma: "rec_exec rec_valu [r] = valu r" |
|
3366 apply(simp add: rec_exec.simps rec_valu_def lg_lemma) |
|
3367 done |
|
3368 |
|
3369 lemma [intro]: "primerec rec_valu (Suc 0)" |
|
3370 apply(simp add: rec_valu_def) |
|
3371 apply(rule_tac k = "Suc (Suc 0)" in prime_cn) |
|
3372 apply(auto simp: prime_s) |
|
3373 proof - |
|
3374 show "primerec rec_lg (Suc (Suc 0))" by auto |
|
3375 next |
|
3376 show "Suc (Suc 0) = Suc (Suc 0)" by simp |
|
3377 next |
|
3378 show "primerec (constn (Suc (Suc 0))) (Suc 0)" by auto |
|
3379 qed |
|
3380 |
|
3381 lemma [simp]: "rec_calc_rel rec_valu [r] rs = |
|
3382 (rec_exec rec_valu [r] = rs)" |
|
3383 apply(rule_tac prime_rel_exec_eq, auto) |
|
3384 done |
|
3385 |
|
3386 declare valu.simps[simp del] |
|
3387 |
|
3388 text {* |
|
3389 The definition of the universal function @{text "rec_F"}. |
|
3390 *} |
|
3391 definition rec_F :: "recf" |
|
3392 where |
|
3393 "rec_F = Cn (Suc (Suc 0)) rec_valu [Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) |
|
3394 rec_conf ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]]" |
|
3395 |
|
3396 lemma get_fstn_args_nth: |
|
3397 "k < n \<Longrightarrow> (get_fstn_args m n ! k) = id m (k)" |
|
3398 apply(induct n, simp) |
|
3399 apply(case_tac "k = n", simp_all add: get_fstn_args.simps |
|
3400 nth_append) |
|
3401 done |
|
3402 |
|
3403 lemma [simp]: |
|
3404 "\<lbrakk>ys \<noteq> []; k < length ys\<rbrakk> \<Longrightarrow> |
|
3405 (get_fstn_args (length ys) (length ys) ! k) = |
|
3406 id (length ys) (k)" |
|
3407 by(erule_tac get_fstn_args_nth) |
|
3408 |
|
3409 lemma calc_rel_get_pren: |
|
3410 "\<lbrakk>ys \<noteq> []; k < length ys\<rbrakk> \<Longrightarrow> |
|
3411 rec_calc_rel (get_fstn_args (length ys) (length ys) ! k) ys |
|
3412 (ys ! k)" |
|
3413 apply(simp) |
|
3414 apply(rule_tac calc_id, auto) |
|
3415 done |
|
3416 |
|
3417 lemma [elim]: |
|
3418 "\<lbrakk>xs \<noteq> []; k < Suc (length xs)\<rbrakk> \<Longrightarrow> |
|
3419 rec_calc_rel (get_fstn_args (Suc (length xs)) |
|
3420 (Suc (length xs)) ! k) (m # xs) ((m # xs) ! k)" |
|
3421 using calc_rel_get_pren[of "m#xs" k] |
|
3422 apply(simp) |
|
3423 done |
|
3424 |
|
3425 text {* |
|
3426 The correctness of @{text "rec_F"}, halt case. |
|
3427 *} |
|
3428 lemma F_lemma: |
|
3429 "rec_calc_rel rec_halt [m, r] t \<Longrightarrow> |
|
3430 rec_calc_rel rec_F [m, r] (valu (rght (conf m r t)))" |
|
3431 apply(simp add: rec_F_def) |
|
3432 apply(rule_tac rs = "[rght (conf m r t)]" in calc_cn, |
|
3433 auto simp: value_lemma) |
|
3434 apply(rule_tac rs = "[conf m r t]" in calc_cn, |
|
3435 auto simp: right_lemma) |
|
3436 apply(rule_tac rs = "[m, r, t]" in calc_cn, auto) |
|
3437 apply(subgoal_tac " k = 0 \<or> k = Suc 0 \<or> k = Suc (Suc 0)", |
|
3438 auto simp:nth_append) |
|
3439 apply(rule_tac [1-2] calc_id, simp_all add: conf_lemma) |
|
3440 done |
|
3441 |
|
3442 |
|
3443 text {* |
|
3444 The correctness of @{text "rec_F"}, nonhalt case. |
|
3445 *} |
|
3446 lemma F_lemma2: |
|
3447 "\<forall> t. \<not> rec_calc_rel rec_halt [m, r] t \<Longrightarrow> |
|
3448 \<forall> rs. \<not> rec_calc_rel rec_F [m, r] rs" |
|
3449 apply(auto simp: rec_F_def) |
|
3450 apply(erule_tac calc_cn_reverse, simp (no_asm_use))+ |
|
3451 proof - |
|
3452 fix rs rsa rsb rsc |
|
3453 assume h: |
|
3454 "\<forall>t. \<not> rec_calc_rel rec_halt [m, r] t" |
|
3455 "length rsa = Suc 0" |
|
3456 "rec_calc_rel rec_valu rsa rs" |
|
3457 "length rsb = Suc 0" |
|
3458 "rec_calc_rel rec_right rsb (rsa ! 0)" |
|
3459 "length rsc = (Suc (Suc (Suc 0)))" |
|
3460 "rec_calc_rel rec_conf rsc (rsb ! 0)" |
|
3461 and g: "\<forall>k<Suc (Suc (Suc 0)). rec_calc_rel ([recf.id (Suc (Suc 0)) 0, |
|
3462 recf.id (Suc (Suc 0)) (Suc 0), rec_halt] ! k) [m, r] (rsc ! k)" |
|
3463 have "rec_calc_rel (rec_halt ) [m, r] |
|
3464 (rsc ! (Suc (Suc 0)))" |
|
3465 using g |
|
3466 apply(erule_tac x = "(Suc (Suc 0))" in allE) |
|
3467 apply(simp add:nth_append) |
|
3468 done |
|
3469 thus "False" |
|
3470 using h |
|
3471 apply(erule_tac x = "ysb ! (Suc (Suc 0))" in allE, simp) |
|
3472 done |
|
3473 qed |
|
3474 |
|
3475 |
|
3476 subsection {* Coding function of TMs *} |
|
3477 |
|
3478 text {* |
|
3479 The purpose of this section is to get the coding function of Turing Machine, which is |
|
3480 going to be named @{text "code"}. |
|
3481 *} |
|
3482 |
|
3483 fun bl2nat :: "block list \<Rightarrow> nat \<Rightarrow> nat" |
|
3484 where |
|
3485 "bl2nat [] n = 0" |
|
3486 | "bl2nat (Bk#bl) n = bl2nat bl (Suc n)" |
|
3487 | "bl2nat (Oc#bl) n = 2^n + bl2nat bl (Suc n)" |
|
3488 |
|
3489 fun bl2wc :: "block list \<Rightarrow> nat" |
|
3490 where |
|
3491 "bl2wc xs = bl2nat xs 0" |
|
3492 |
|
3493 fun trpl_code :: "t_conf \<Rightarrow> nat" |
|
3494 where |
|
3495 "trpl_code (st, l, r) = trpl (bl2wc l) st (bl2wc r)" |
|
3496 |
|
3497 declare bl2nat.simps[simp del] bl2wc.simps[simp del] |
|
3498 trpl_code.simps[simp del] |
|
3499 |
|
3500 fun action_map :: "taction \<Rightarrow> nat" |
|
3501 where |
|
3502 "action_map W0 = 0" |
|
3503 | "action_map W1 = 1" |
|
3504 | "action_map L = 2" |
|
3505 | "action_map R = 3" |
|
3506 | "action_map Nop = 4" |
|
3507 |
|
3508 fun action_map_iff :: "nat \<Rightarrow> taction" |
|
3509 where |
|
3510 "action_map_iff (0::nat) = W0" |
|
3511 | "action_map_iff (Suc 0) = W1" |
|
3512 | "action_map_iff (Suc (Suc 0)) = L" |
|
3513 | "action_map_iff (Suc (Suc (Suc 0))) = R" |
|
3514 | "action_map_iff n = Nop" |
|
3515 |
|
3516 fun block_map :: "block \<Rightarrow> nat" |
|
3517 where |
|
3518 "block_map Bk = 0" |
|
3519 | "block_map Oc = 1" |
|
3520 |
|
3521 fun godel_code' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
|
3522 where |
|
3523 "godel_code' [] n = 1" |
|
3524 | "godel_code' (x#xs) n = (Pi n)^x * godel_code' xs (Suc n) " |
|
3525 |
|
3526 fun godel_code :: "nat list \<Rightarrow> nat" |
|
3527 where |
|
3528 "godel_code xs = (let lh = length xs in |
|
3529 2^lh * (godel_code' xs (Suc 0)))" |
|
3530 |
|
3531 fun modify_tprog :: "tprog \<Rightarrow> nat list" |
|
3532 where |
|
3533 "modify_tprog [] = []" |
|
3534 | "modify_tprog ((ac, ns)#nl) = action_map ac # ns # modify_tprog nl" |
|
3535 |
|
3536 text {* |
|
3537 @{text "code tp"} gives the Godel coding of TM program @{text "tp"}. |
|
3538 *} |
|
3539 fun code :: "tprog \<Rightarrow> nat" |
|
3540 where |
|
3541 "code tp = (let nl = modify_tprog tp in |
|
3542 godel_code nl)" |
|
3543 |
|
3544 subsection {* Relating interperter functions to the execution of TMs *} |
|
3545 |
|
3546 lemma [simp]: "bl2wc [] = 0" by(simp add: bl2wc.simps bl2nat.simps) |
|
3547 term trpl |
|
3548 |
|
3549 lemma [simp]: "\<lbrakk>fetch tp 0 b = (nact, ns)\<rbrakk> \<Longrightarrow> action_map nact = 4" |
|
3550 apply(simp add: fetch.simps) |
|
3551 done |
|
3552 |
|
3553 lemma Pi_gr_1[simp]: "Pi n > Suc 0" |
|
3554 proof(induct n, auto simp: Pi.simps Np.simps) |
|
3555 fix n |
|
3556 let ?setx = "{y. y \<le> Suc (Pi n!) \<and> Pi n < y \<and> Prime y}" |
|
3557 have "finite ?setx" by auto |
|
3558 moreover have "?setx \<noteq> {}" |
|
3559 using prime_ex[of "Pi n"] |
|
3560 apply(auto) |
|
3561 done |
|
3562 ultimately show "Suc 0 < Min ?setx" |
|
3563 apply(simp add: Min_gr_iff) |
|
3564 apply(auto simp: Prime.simps) |
|
3565 done |
|
3566 qed |
|
3567 |
|
3568 lemma Pi_not_0[simp]: "Pi n > 0" |
|
3569 using Pi_gr_1[of n] |
|
3570 by arith |
|
3571 |
|
3572 declare godel_code.simps[simp del] |
|
3573 |
|
3574 lemma [simp]: "0 < godel_code' nl n" |
|
3575 apply(induct nl arbitrary: n) |
|
3576 apply(auto simp: godel_code'.simps) |
|
3577 done |
|
3578 |
|
3579 lemma godel_code_great: "godel_code nl > 0" |
|
3580 apply(simp add: godel_code.simps) |
|
3581 done |
|
3582 |
|
3583 lemma godel_code_eq_1: "(godel_code nl = 1) = (nl = [])" |
|
3584 apply(auto simp: godel_code.simps) |
|
3585 done |
|
3586 |
|
3587 lemma [elim]: |
|
3588 "\<lbrakk>i < length nl; \<not> Suc 0 < godel_code nl\<rbrakk> \<Longrightarrow> nl ! i = 0" |
|
3589 using godel_code_great[of nl] godel_code_eq_1[of nl] |
|
3590 apply(simp) |
|
3591 done |
|
3592 |
|
3593 term set_of |
|
3594 lemma prime_coprime: "\<lbrakk>Prime x; Prime y; x\<noteq>y\<rbrakk> \<Longrightarrow> coprime x y" |
|
3595 proof(simp only: Prime.simps coprime_nat, auto simp: dvd_def, |
|
3596 rule_tac classical, simp) |
|
3597 fix d k ka |
|
3598 assume case_ka: "\<forall>u<d * ka. \<forall>v<d * ka. u * v \<noteq> d * ka" |
|
3599 and case_k: "\<forall>u<d * k. \<forall>v<d * k. u * v \<noteq> d * k" |
|
3600 and h: "(0::nat) < d" "d \<noteq> Suc 0" "Suc 0 < d * ka" |
|
3601 "ka \<noteq> k" "Suc 0 < d * k" |
|
3602 from h have "k > Suc 0 \<or> ka >Suc 0" |
|
3603 apply(auto) |
|
3604 apply(case_tac ka, simp, simp) |
|
3605 apply(case_tac k, simp, simp) |
|
3606 done |
|
3607 from this show "False" |
|
3608 proof(erule_tac disjE) |
|
3609 assume "(Suc 0::nat) < k" |
|
3610 hence "k < d*k \<and> d < d*k" |
|
3611 using h |
|
3612 by(auto) |
|
3613 thus "?thesis" |
|
3614 using case_k |
|
3615 apply(erule_tac x = d in allE) |
|
3616 apply(simp) |
|
3617 apply(erule_tac x = k in allE) |
|
3618 apply(simp) |
|
3619 done |
|
3620 next |
|
3621 assume "(Suc 0::nat) < ka" |
|
3622 hence "ka < d * ka \<and> d < d*ka" |
|
3623 using h by auto |
|
3624 thus "?thesis" |
|
3625 using case_ka |
|
3626 apply(erule_tac x = d in allE) |
|
3627 apply(simp) |
|
3628 apply(erule_tac x = ka in allE) |
|
3629 apply(simp) |
|
3630 done |
|
3631 qed |
|
3632 qed |
|
3633 |
|
3634 lemma Pi_inc: "Pi (Suc i) > Pi i" |
|
3635 proof(simp add: Pi.simps Np.simps) |
|
3636 let ?setx = "{y. y \<le> Suc (Pi i!) \<and> Pi i < y \<and> Prime y}" |
|
3637 have "finite ?setx" by simp |
|
3638 moreover have "?setx \<noteq> {}" |
|
3639 using prime_ex[of "Pi i"] |
|
3640 apply(auto) |
|
3641 done |
|
3642 ultimately show "Pi i < Min ?setx" |
|
3643 apply(simp add: Min_gr_iff) |
|
3644 done |
|
3645 qed |
|
3646 |
|
3647 lemma Pi_inc_gr: "i < j \<Longrightarrow> Pi i < Pi j" |
|
3648 proof(induct j, simp) |
|
3649 fix j |
|
3650 assume ind: "i < j \<Longrightarrow> Pi i < Pi j" |
|
3651 and h: "i < Suc j" |
|
3652 from h show "Pi i < Pi (Suc j)" |
|
3653 proof(cases "i < j") |
|
3654 case True thus "?thesis" |
|
3655 proof - |
|
3656 assume "i < j" |
|
3657 hence "Pi i < Pi j" by(erule_tac ind) |
|
3658 moreover have "Pi j < Pi (Suc j)" |
|
3659 apply(simp add: Pi_inc) |
|
3660 done |
|
3661 ultimately show "?thesis" |
|
3662 by simp |
|
3663 qed |
|
3664 next |
|
3665 assume "i < Suc j" "\<not> i < j" |
|
3666 hence "i = j" |
|
3667 by arith |
|
3668 thus "Pi i < Pi (Suc j)" |
|
3669 apply(simp add: Pi_inc) |
|
3670 done |
|
3671 qed |
|
3672 qed |
|
3673 |
|
3674 lemma Pi_notEq: "i \<noteq> j \<Longrightarrow> Pi i \<noteq> Pi j" |
|
3675 apply(case_tac "i < j") |
|
3676 using Pi_inc_gr[of i j] |
|
3677 apply(simp) |
|
3678 using Pi_inc_gr[of j i] |
|
3679 apply(simp) |
|
3680 done |
|
3681 |
|
3682 lemma [intro]: "Prime (Suc (Suc 0))" |
|
3683 apply(auto simp: Prime.simps) |
|
3684 apply(case_tac u, simp, case_tac nat, simp, simp) |
|
3685 done |
|
3686 |
|
3687 lemma Prime_Pi[intro]: "Prime (Pi n)" |
|
3688 proof(induct n, auto simp: Pi.simps Np.simps) |
|
3689 fix n |
|
3690 let ?setx = "{y. y \<le> Suc (Pi n!) \<and> Pi n < y \<and> Prime y}" |
|
3691 show "Prime (Min ?setx)" |
|
3692 proof - |
|
3693 have "finite ?setx" by simp |
|
3694 moreover have "?setx \<noteq> {}" |
|
3695 using prime_ex[of "Pi n"] |
|
3696 apply(simp) |
|
3697 done |
|
3698 ultimately show "?thesis" |
|
3699 apply(drule_tac Min_in, simp, simp) |
|
3700 done |
|
3701 qed |
|
3702 qed |
|
3703 |
|
3704 lemma Pi_coprime: "i \<noteq> j \<Longrightarrow> coprime (Pi i) (Pi j)" |
|
3705 using Prime_Pi[of i] |
|
3706 using Prime_Pi[of j] |
|
3707 apply(rule_tac prime_coprime, simp_all add: Pi_notEq) |
|
3708 done |
|
3709 |
|
3710 lemma Pi_power_coprime: "i \<noteq> j \<Longrightarrow> coprime ((Pi i)^m) ((Pi j)^n)" |
|
3711 by(rule_tac coprime_exp2_nat, erule_tac Pi_coprime) |
|
3712 |
|
3713 lemma coprime_dvd_mult_nat2: "\<lbrakk>coprime (k::nat) n; k dvd n * m\<rbrakk> \<Longrightarrow> k dvd m" |
|
3714 apply(erule_tac coprime_dvd_mult_nat) |
|
3715 apply(simp add: dvd_def, auto) |
|
3716 apply(rule_tac x = ka in exI) |
|
3717 apply(subgoal_tac "n * m = m * n", simp) |
|
3718 apply(simp add: nat_mult_commute) |
|
3719 done |
|
3720 |
|
3721 declare godel_code'.simps[simp del] |
|
3722 |
|
3723 lemma godel_code'_butlast_last_id' : |
|
3724 "godel_code' (ys @ [y]) (Suc j) = godel_code' ys (Suc j) * |
|
3725 Pi (Suc (length ys + j)) ^ y" |
|
3726 proof(induct ys arbitrary: j, simp_all add: godel_code'.simps) |
|
3727 qed |
|
3728 |
|
3729 lemma godel_code'_butlast_last_id: |
|
3730 "xs \<noteq> [] \<Longrightarrow> godel_code' xs (Suc j) = |
|
3731 godel_code' (butlast xs) (Suc j) * Pi (length xs + j)^(last xs)" |
|
3732 apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]") |
|
3733 apply(erule_tac exE, erule_tac exE, simp add: |
|
3734 godel_code'_butlast_last_id') |
|
3735 apply(rule_tac x = "butlast xs" in exI) |
|
3736 apply(rule_tac x = "last xs" in exI, auto) |
|
3737 done |
|
3738 |
|
3739 lemma godel_code'_not0: "godel_code' xs n \<noteq> 0" |
|
3740 apply(induct xs, auto simp: godel_code'.simps) |
|
3741 done |
|
3742 |
|
3743 lemma godel_code_append_cons: |
|
3744 "length xs = i \<Longrightarrow> godel_code' (xs@y#ys) (Suc 0) |
|
3745 = godel_code' xs (Suc 0) * Pi (Suc i)^y * godel_code' ys (i + 2)" |
|
3746 proof(induct "length xs" arbitrary: i y ys xs, simp add: godel_code'.simps,simp) |
|
3747 fix x xs i y ys |
|
3748 assume ind: |
|
3749 "\<And>xs i y ys. \<lbrakk>x = i; length xs = i\<rbrakk> \<Longrightarrow> |
|
3750 godel_code' (xs @ y # ys) (Suc 0) |
|
3751 = godel_code' xs (Suc 0) * Pi (Suc i) ^ y * |
|
3752 godel_code' ys (Suc (Suc i))" |
|
3753 and h: "Suc x = i" |
|
3754 "length (xs::nat list) = i" |
|
3755 have |
|
3756 "godel_code' (butlast xs @ last xs # ((y::nat)#ys)) (Suc 0) = |
|
3757 godel_code' (butlast xs) (Suc 0) * Pi (Suc (i - 1))^(last xs) |
|
3758 * godel_code' (y#ys) (Suc (Suc (i - 1)))" |
|
3759 apply(rule_tac ind) |
|
3760 using h |
|
3761 by(auto) |
|
3762 moreover have |
|
3763 "godel_code' xs (Suc 0)= godel_code' (butlast xs) (Suc 0) * |
|
3764 Pi (i)^(last xs)" |
|
3765 using godel_code'_butlast_last_id[of xs] h |
|
3766 apply(case_tac "xs = []", simp, simp) |
|
3767 done |
|
3768 moreover have "butlast xs @ last xs # y # ys = xs @ y # ys" |
|
3769 using h |
|
3770 apply(case_tac xs, auto) |
|
3771 done |
|
3772 ultimately show |
|
3773 "godel_code' (xs @ y # ys) (Suc 0) = |
|
3774 godel_code' xs (Suc 0) * Pi (Suc i) ^ y * |
|
3775 godel_code' ys (Suc (Suc i))" |
|
3776 using h |
|
3777 apply(simp add: godel_code'_not0 Pi_not_0) |
|
3778 apply(simp add: godel_code'.simps) |
|
3779 done |
|
3780 qed |
|
3781 |
|
3782 lemma Pi_coprime_pre: |
|
3783 "length ps \<le> i \<Longrightarrow> coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" |
|
3784 proof(induct "length ps" arbitrary: ps, simp add: godel_code'.simps) |
|
3785 fix x ps |
|
3786 assume ind: |
|
3787 "\<And>ps. \<lbrakk>x = length ps; length ps \<le> i\<rbrakk> \<Longrightarrow> |
|
3788 coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" |
|
3789 and h: "Suc x = length ps" |
|
3790 "length (ps::nat list) \<le> i" |
|
3791 have g: "coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0))" |
|
3792 apply(rule_tac ind) |
|
3793 using h by auto |
|
3794 have k: "godel_code' ps (Suc 0) = |
|
3795 godel_code' (butlast ps) (Suc 0) * Pi (length ps)^(last ps)" |
|
3796 using godel_code'_butlast_last_id[of ps 0] h |
|
3797 by(case_tac ps, simp, simp) |
|
3798 from g have |
|
3799 "coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0) * |
|
3800 Pi (length ps)^(last ps)) " |
|
3801 proof(rule_tac coprime_mult_nat, simp) |
|
3802 show "coprime (Pi (Suc i)) (Pi (length ps) ^ last ps)" |
|
3803 apply(rule_tac coprime_exp_nat, rule prime_coprime, auto) |
|
3804 using Pi_notEq[of "Suc i" "length ps"] h by simp |
|
3805 qed |
|
3806 from this and k show "coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" |
|
3807 by simp |
|
3808 qed |
|
3809 |
|
3810 lemma Pi_coprime_suf: "i < j \<Longrightarrow> coprime (Pi i) (godel_code' ps j)" |
|
3811 proof(induct "length ps" arbitrary: ps, simp add: godel_code'.simps) |
|
3812 fix x ps |
|
3813 assume ind: |
|
3814 "\<And>ps. \<lbrakk>x = length ps; i < j\<rbrakk> \<Longrightarrow> |
|
3815 coprime (Pi i) (godel_code' ps j)" |
|
3816 and h: "Suc x = length (ps::nat list)" "i < j" |
|
3817 have g: "coprime (Pi i) (godel_code' (butlast ps) j)" |
|
3818 apply(rule ind) using h by auto |
|
3819 have k: "(godel_code' ps j) = godel_code' (butlast ps) j * |
|
3820 Pi (length ps + j - 1)^last ps" |
|
3821 using h godel_code'_butlast_last_id[of ps "j - 1"] |
|
3822 apply(case_tac "ps = []", simp, simp) |
|
3823 done |
|
3824 from g have |
|
3825 "coprime (Pi i) (godel_code' (butlast ps) j * |
|
3826 Pi (length ps + j - 1)^last ps)" |
|
3827 apply(rule_tac coprime_mult_nat, simp) |
|
3828 using Pi_power_coprime[of i "length ps + j - 1" 1 "last ps"] h |
|
3829 apply(auto) |
|
3830 done |
|
3831 from k and this show "coprime (Pi i) (godel_code' ps j)" |
|
3832 by auto |
|
3833 qed |
|
3834 |
|
3835 lemma godel_finite: |
|
3836 "finite {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
|
3837 proof(rule_tac n = "godel_code' nl (Suc 0)" in |
|
3838 bounded_nat_set_is_finite, auto, |
|
3839 case_tac "ia < godel_code' nl (Suc 0)", auto) |
|
3840 fix ia |
|
3841 assume g1: "Pi (Suc i) ^ ia dvd godel_code' nl (Suc 0)" |
|
3842 and g2: "\<not> ia < godel_code' nl (Suc 0)" |
|
3843 from g1 have "Pi (Suc i)^ia \<le> godel_code' nl (Suc 0)" |
|
3844 apply(erule_tac dvd_imp_le) |
|
3845 using godel_code'_not0[of nl "Suc 0"] by simp |
|
3846 moreover have "ia < Pi (Suc i)^ia" |
|
3847 apply(rule x_less_exp) |
|
3848 using Pi_gr_1 by auto |
|
3849 ultimately show "False" |
|
3850 using g2 |
|
3851 by(auto) |
|
3852 qed |
|
3853 |
|
3854 |
|
3855 lemma godel_code_in: |
|
3856 "i < length nl \<Longrightarrow> nl ! i \<in> {u. Pi (Suc i) ^ u dvd |
|
3857 godel_code' nl (Suc 0)}" |
|
3858 proof - |
|
3859 assume h: "i<length nl" |
|
3860 hence "godel_code' (take i nl@(nl!i)#drop (Suc i) nl) (Suc 0) |
|
3861 = godel_code' (take i nl) (Suc 0) * Pi (Suc i)^(nl!i) * |
|
3862 godel_code' (drop (Suc i) nl) (i + 2)" |
|
3863 by(rule_tac godel_code_append_cons, simp) |
|
3864 moreover from h have "take i nl @ (nl ! i) # drop (Suc i) nl = nl" |
|
3865 using upd_conv_take_nth_drop[of i nl "nl ! i"] |
|
3866 apply(simp) |
|
3867 done |
|
3868 ultimately show |
|
3869 "nl ! i \<in> {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
|
3870 by(simp) |
|
3871 qed |
|
3872 |
|
3873 lemma godel_code'_get_nth: |
|
3874 "i < length nl \<Longrightarrow> Max {u. Pi (Suc i) ^ u dvd |
|
3875 godel_code' nl (Suc 0)} = nl ! i" |
|
3876 proof(rule_tac Max_eqI) |
|
3877 let ?gc = "godel_code' nl (Suc 0)" |
|
3878 assume h: "i < length nl" thus "finite {u. Pi (Suc i) ^ u dvd ?gc}" |
|
3879 by (simp add: godel_finite) |
|
3880 next |
|
3881 fix y |
|
3882 let ?suf ="godel_code' (drop (Suc i) nl) (i + 2)" |
|
3883 let ?pref = "godel_code' (take i nl) (Suc 0)" |
|
3884 assume h: "i < length nl" |
|
3885 "y \<in> {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
|
3886 moreover hence |
|
3887 "godel_code' (take i nl@(nl!i)#drop (Suc i) nl) (Suc 0) |
|
3888 = ?pref * Pi (Suc i)^(nl!i) * ?suf" |
|
3889 by(rule_tac godel_code_append_cons, simp) |
|
3890 moreover from h have "take i nl @ (nl!i) # drop (Suc i) nl = nl" |
|
3891 using upd_conv_take_nth_drop[of i nl "nl!i"] |
|
3892 by simp |
|
3893 ultimately show "y\<le>nl!i" |
|
3894 proof(simp) |
|
3895 let ?suf' = "godel_code' (drop (Suc i) nl) (Suc (Suc i))" |
|
3896 assume mult_dvd: |
|
3897 "Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i * ?suf'" |
|
3898 hence "Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i" |
|
3899 proof(rule_tac coprime_dvd_mult_nat) |
|
3900 show "coprime (Pi (Suc i)^y) ?suf'" |
|
3901 proof - |
|
3902 have "coprime (Pi (Suc i) ^ y) (?suf'^(Suc 0))" |
|
3903 apply(rule_tac coprime_exp2_nat) |
|
3904 apply(rule_tac Pi_coprime_suf, simp) |
|
3905 done |
|
3906 thus "?thesis" by simp |
|
3907 qed |
|
3908 qed |
|
3909 hence "Pi (Suc i) ^ y dvd Pi (Suc i) ^ nl ! i" |
|
3910 proof(rule_tac coprime_dvd_mult_nat2) |
|
3911 show "coprime (Pi (Suc i) ^ y) ?pref" |
|
3912 proof - |
|
3913 have "coprime (Pi (Suc i)^y) (?pref^Suc 0)" |
|
3914 apply(rule_tac coprime_exp2_nat) |
|
3915 apply(rule_tac Pi_coprime_pre, simp) |
|
3916 done |
|
3917 thus "?thesis" by simp |
|
3918 qed |
|
3919 qed |
|
3920 hence "Pi (Suc i) ^ y \<le> Pi (Suc i) ^ nl ! i " |
|
3921 apply(rule_tac dvd_imp_le, auto) |
|
3922 done |
|
3923 thus "y \<le> nl ! i" |
|
3924 apply(rule_tac power_le_imp_le_exp, auto) |
|
3925 done |
|
3926 qed |
|
3927 next |
|
3928 assume h: "i<length nl" |
|
3929 thus "nl ! i \<in> {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
|
3930 by(rule_tac godel_code_in, simp) |
|
3931 qed |
|
3932 |
|
3933 lemma [simp]: |
|
3934 "{u. Pi (Suc i) ^ u dvd (Suc (Suc 0)) ^ length nl * |
|
3935 godel_code' nl (Suc 0)} = |
|
3936 {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
|
3937 apply(rule_tac Collect_cong, auto) |
|
3938 apply(rule_tac n = " (Suc (Suc 0)) ^ length nl" in |
|
3939 coprime_dvd_mult_nat2) |
|
3940 proof - |
|
3941 fix u |
|
3942 show "coprime (Pi (Suc i) ^ u) ((Suc (Suc 0)) ^ length nl)" |
|
3943 proof(rule_tac coprime_exp2_nat) |
|
3944 have "Pi 0 = (2::nat)" |
|
3945 apply(simp add: Pi.simps) |
|
3946 done |
|
3947 moreover have "coprime (Pi (Suc i)) (Pi 0)" |
|
3948 apply(rule_tac Pi_coprime, simp) |
|
3949 done |
|
3950 ultimately show "coprime (Pi (Suc i)) (Suc (Suc 0))" by simp |
|
3951 qed |
|
3952 qed |
|
3953 |
|
3954 lemma godel_code_get_nth: |
|
3955 "i < length nl \<Longrightarrow> |
|
3956 Max {u. Pi (Suc i) ^ u dvd godel_code nl} = nl ! i" |
|
3957 by(simp add: godel_code.simps godel_code'_get_nth) |
|
3958 |
|
3959 lemma "trpl l st r = godel_code' [l, st, r] 0" |
|
3960 apply(simp add: trpl.simps godel_code'.simps) |
|
3961 done |
|
3962 |
|
3963 lemma mod_dvd_simp: "(x mod y = (0::nat)) = (y dvd x)" |
|
3964 by(simp add: dvd_def, auto) |
|
3965 |
|
3966 lemma dvd_power_le: "\<lbrakk>a > Suc 0; a ^ y dvd a ^ l\<rbrakk> \<Longrightarrow> y \<le> l" |
|
3967 apply(case_tac "y \<le> l", simp, simp) |
|
3968 apply(subgoal_tac "\<exists> d. y = l + d", auto simp: power_add) |
|
3969 apply(rule_tac x = "y - l" in exI, simp) |
|
3970 done |
|
3971 |
|
3972 |
|
3973 lemma [elim]: "Pi n = 0 \<Longrightarrow> RR" |
|
3974 using Pi_not_0[of n] by simp |
|
3975 |
|
3976 lemma [elim]: "Pi n = Suc 0 \<Longrightarrow> RR" |
|
3977 using Pi_gr_1[of n] by simp |
|
3978 |
|
3979 lemma finite_power_dvd: |
|
3980 "\<lbrakk>(a::nat) > Suc 0; y \<noteq> 0\<rbrakk> \<Longrightarrow> finite {u. a^u dvd y}" |
|
3981 apply(auto simp: dvd_def) |
|
3982 apply(rule_tac n = y in bounded_nat_set_is_finite, auto) |
|
3983 apply(case_tac k, simp,simp) |
|
3984 apply(rule_tac trans_less_add1) |
|
3985 apply(erule_tac x_less_exp) |
|
3986 done |
|
3987 |
|
3988 lemma conf_decode1: "\<lbrakk>m \<noteq> n; m \<noteq> k; k \<noteq> n\<rbrakk> \<Longrightarrow> |
|
3989 Max {u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r} = l" |
|
3990 proof - |
|
3991 let ?setx = "{u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r}" |
|
3992 assume g: "m \<noteq> n" "m \<noteq> k" "k \<noteq> n" |
|
3993 show "Max ?setx = l" |
|
3994 proof(rule_tac Max_eqI) |
|
3995 show "finite ?setx" |
|
3996 apply(rule_tac finite_power_dvd, auto simp: Pi_gr_1) |
|
3997 done |
|
3998 next |
|
3999 fix y |
|
4000 assume h: "y \<in> ?setx" |
|
4001 have "Pi m ^ y dvd Pi m ^ l" |
|
4002 proof - |
|
4003 have "Pi m ^ y dvd Pi m ^ l * Pi n ^ st" |
|
4004 using h g |
|
4005 apply(rule_tac n = "Pi k^r" in coprime_dvd_mult_nat) |
|
4006 apply(rule Pi_power_coprime, simp, simp) |
|
4007 done |
|
4008 thus "Pi m^y dvd Pi m^l" |
|
4009 apply(rule_tac n = " Pi n ^ st" in coprime_dvd_mult_nat) |
|
4010 using g |
|
4011 apply(rule_tac Pi_power_coprime, simp, simp) |
|
4012 done |
|
4013 qed |
|
4014 thus "y \<le> (l::nat)" |
|
4015 apply(rule_tac a = "Pi m" in power_le_imp_le_exp) |
|
4016 apply(simp_all add: Pi_gr_1) |
|
4017 apply(rule_tac dvd_power_le, auto) |
|
4018 done |
|
4019 next |
|
4020 show "l \<in> ?setx" by simp |
|
4021 qed |
|
4022 qed |
|
4023 |
|
4024 lemma conf_decode2: |
|
4025 "\<lbrakk>m \<noteq> n; m \<noteq> k; n \<noteq> k; |
|
4026 \<not> Suc 0 < Pi m ^ l * Pi n ^ st * Pi k ^ r\<rbrakk> \<Longrightarrow> l = 0" |
|
4027 apply(case_tac "Pi m ^ l * Pi n ^ st * Pi k ^ r", auto) |
|
4028 done |
|
4029 |
|
4030 lemma [simp]: "left (trpl l st r) = l" |
|
4031 apply(simp add: left.simps trpl.simps lo.simps |
|
4032 loR.simps mod_dvd_simp, auto simp: conf_decode1) |
|
4033 apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", |
|
4034 auto) |
|
4035 apply(erule_tac x = l in allE, auto) |
|
4036 done |
|
4037 |
|
4038 lemma [simp]: "stat (trpl l st r) = st" |
|
4039 apply(simp add: stat.simps trpl.simps lo.simps |
|
4040 loR.simps mod_dvd_simp, auto) |
|
4041 apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r |
|
4042 = Pi (Suc 0)^st * Pi 0 ^ l * Pi (Suc (Suc 0)) ^ r") |
|
4043 apply(simp (no_asm_simp) add: conf_decode1, simp) |
|
4044 apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * |
|
4045 Pi (Suc (Suc 0)) ^ r", auto) |
|
4046 apply(erule_tac x = st in allE, auto) |
|
4047 done |
|
4048 |
|
4049 lemma [simp]: "rght (trpl l st r) = r" |
|
4050 apply(simp add: rght.simps trpl.simps lo.simps |
|
4051 loR.simps mod_dvd_simp, auto) |
|
4052 apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r |
|
4053 = Pi (Suc (Suc 0))^r * Pi 0 ^ l * Pi (Suc 0) ^ st") |
|
4054 apply(simp (no_asm_simp) add: conf_decode1, simp) |
|
4055 apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", |
|
4056 auto) |
|
4057 apply(erule_tac x = r in allE, auto) |
|
4058 done |
|
4059 |
|
4060 lemma max_lor: |
|
4061 "i < length nl \<Longrightarrow> Max {u. loR [godel_code nl, Pi (Suc i), u]} |
|
4062 = nl ! i" |
|
4063 apply(simp add: loR.simps godel_code_get_nth mod_dvd_simp) |
|
4064 done |
|
4065 |
|
4066 lemma godel_decode: |
|
4067 "i < length nl \<Longrightarrow> Entry (godel_code nl) i = nl ! i" |
|
4068 apply(auto simp: Entry.simps lo.simps max_lor) |
|
4069 apply(erule_tac x = "nl!i" in allE) |
|
4070 using max_lor[of i nl] godel_finite[of i nl] |
|
4071 apply(simp) |
|
4072 apply(drule_tac Max_in, auto simp: loR.simps |
|
4073 godel_code.simps mod_dvd_simp) |
|
4074 using godel_code_in[of i nl] |
|
4075 apply(simp) |
|
4076 done |
|
4077 |
|
4078 lemma Four_Suc: "4 = Suc (Suc (Suc (Suc 0)))" |
|
4079 by auto |
|
4080 |
|
4081 declare numeral_2_eq_2[simp del] |
|
4082 |
|
4083 lemma modify_tprog_fetch_even: |
|
4084 "\<lbrakk>st \<le> length tp div 2; st > 0\<rbrakk> \<Longrightarrow> |
|
4085 modify_tprog tp ! (4 * (st - Suc 0) ) = |
|
4086 action_map (fst (tp ! (2 * (st - Suc 0))))" |
|
4087 proof(induct st arbitrary: tp, simp) |
|
4088 fix tp st |
|
4089 assume ind: |
|
4090 "\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st\<rbrakk> \<Longrightarrow> |
|
4091 modify_tprog tp ! (4 * (st - Suc 0)) = |
|
4092 action_map (fst ((tp::tprog) ! (2 * (st - Suc 0))))" |
|
4093 and h: "Suc st \<le> length (tp::tprog) div 2" "0 < Suc st" |
|
4094 thus "modify_tprog tp ! (4 * (Suc st - Suc 0)) = |
|
4095 action_map (fst (tp ! (2 * (Suc st - Suc 0))))" |
|
4096 proof(cases "st = 0") |
|
4097 case True thus "?thesis" |
|
4098 using h |
|
4099 apply(auto) |
|
4100 apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) |
|
4101 done |
|
4102 next |
|
4103 case False |
|
4104 assume g: "st \<noteq> 0" |
|
4105 hence "\<exists> aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" |
|
4106 using h |
|
4107 apply(case_tac tp, simp, case_tac list, simp, simp) |
|
4108 done |
|
4109 from this obtain aa ab ba bb tp' where g1: |
|
4110 "tp = (aa, ab) # (ba, bb) # tp'" by blast |
|
4111 hence g2: |
|
4112 "modify_tprog tp' ! (4 * (st - Suc 0)) = |
|
4113 action_map (fst ((tp'::tprog) ! (2 * (st - Suc 0))))" |
|
4114 apply(rule_tac ind) |
|
4115 using h g by auto |
|
4116 thus "?thesis" |
|
4117 using g1 g |
|
4118 apply(case_tac st, simp, simp add: Four_Suc) |
|
4119 done |
|
4120 qed |
|
4121 qed |
|
4122 |
|
4123 lemma modify_tprog_fetch_odd: |
|
4124 "\<lbrakk>st \<le> length tp div 2; st > 0\<rbrakk> \<Longrightarrow> |
|
4125 modify_tprog tp ! (Suc (Suc (4 * (st - Suc 0)))) = |
|
4126 action_map (fst (tp ! (Suc (2 * (st - Suc 0)))))" |
|
4127 proof(induct st arbitrary: tp, simp) |
|
4128 fix tp st |
|
4129 assume ind: |
|
4130 "\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st\<rbrakk> \<Longrightarrow> |
|
4131 modify_tprog tp ! Suc (Suc (4 * (st - Suc 0))) = |
|
4132 action_map (fst (tp ! Suc (2 * (st - Suc 0))))" |
|
4133 and h: "Suc st \<le> length (tp::tprog) div 2" "0 < Suc st" |
|
4134 thus "modify_tprog tp ! Suc (Suc (4 * (Suc st - Suc 0))) |
|
4135 = action_map (fst (tp ! Suc (2 * (Suc st - Suc 0))))" |
|
4136 proof(cases "st = 0") |
|
4137 case True thus "?thesis" |
|
4138 using h |
|
4139 apply(auto) |
|
4140 apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) |
|
4141 apply(case_tac list, simp, case_tac ab, |
|
4142 simp add: modify_tprog.simps) |
|
4143 done |
|
4144 next |
|
4145 case False |
|
4146 assume g: "st \<noteq> 0" |
|
4147 hence "\<exists> aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" |
|
4148 using h |
|
4149 apply(case_tac tp, simp, case_tac list, simp, simp) |
|
4150 done |
|
4151 from this obtain aa ab ba bb tp' where g1: |
|
4152 "tp = (aa, ab) # (ba, bb) # tp'" by blast |
|
4153 hence g2: "modify_tprog tp' ! Suc (Suc (4 * (st - Suc 0))) = |
|
4154 action_map (fst (tp' ! Suc (2 * (st - Suc 0))))" |
|
4155 apply(rule_tac ind) |
|
4156 using h g by auto |
|
4157 thus "?thesis" |
|
4158 using g1 g |
|
4159 apply(case_tac st, simp, simp add: Four_Suc) |
|
4160 done |
|
4161 qed |
|
4162 qed |
|
4163 |
|
4164 lemma modify_tprog_fetch_action: |
|
4165 "\<lbrakk>st \<le> length tp div 2; st > 0; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> |
|
4166 modify_tprog tp ! (4 * (st - Suc 0) + 2* b) = |
|
4167 action_map (fst (tp ! ((2 * (st - Suc 0)) + b)))" |
|
4168 apply(erule_tac disjE, auto elim: modify_tprog_fetch_odd |
|
4169 modify_tprog_fetch_even) |
|
4170 done |
|
4171 |
|
4172 lemma length_modify: "length (modify_tprog tp) = 2 * length tp" |
|
4173 apply(induct tp, auto) |
|
4174 done |
|
4175 |
|
4176 declare fetch.simps[simp del] |
|
4177 |
|
4178 lemma fetch_action_eq: |
|
4179 "\<lbrakk>block_map b = scan r; fetch tp st b = (nact, ns); |
|
4180 st \<le> length tp div 2\<rbrakk> \<Longrightarrow> actn (code tp) st r = action_map nact" |
|
4181 proof(simp add: actn.simps, auto) |
|
4182 let ?i = "4 * (st - Suc 0) + 2 * (r mod 2)" |
|
4183 assume h: "block_map b = r mod 2" "fetch tp st b = (nact, ns)" |
|
4184 "st \<le> length tp div 2" "0 < st" |
|
4185 have "?i < length (modify_tprog tp)" |
|
4186 proof - |
|
4187 have "length (modify_tprog tp) = 2 * length tp" |
|
4188 by(simp add: length_modify) |
|
4189 thus "?thesis" |
|
4190 using h |
|
4191 by(auto) |
|
4192 qed |
|
4193 hence |
|
4194 "Entry (godel_code (modify_tprog tp))?i = |
|
4195 (modify_tprog tp) ! ?i" |
|
4196 by(erule_tac godel_decode) |
|
4197 moreover have |
|
4198 "modify_tprog tp ! ?i = |
|
4199 action_map (fst (tp ! (2 * (st - Suc 0) + r mod 2)))" |
|
4200 apply(rule_tac modify_tprog_fetch_action) |
|
4201 using h |
|
4202 by(auto) |
|
4203 moreover have "(fst (tp ! (2 * (st - Suc 0) + r mod 2))) = nact" |
|
4204 using h |
|
4205 apply(simp add: fetch.simps nth_of.simps) |
|
4206 apply(case_tac b, auto simp: block_map.simps nth_of.simps split: if_splits) |
|
4207 done |
|
4208 ultimately show |
|
4209 "Entry (godel_code (modify_tprog tp)) |
|
4210 (4 * (st - Suc 0) + 2 * (r mod 2)) |
|
4211 = action_map nact" |
|
4212 by simp |
|
4213 qed |
|
4214 |
|
4215 lemma [simp]: "fetch tp 0 b = (nact, ns) \<Longrightarrow> ns = 0" |
|
4216 by(simp add: fetch.simps) |
|
4217 |
|
4218 lemma Five_Suc: "5 = Suc 4" by simp |
|
4219 |
|
4220 lemma modify_tprog_fetch_state: |
|
4221 "\<lbrakk>st \<le> length tp div 2; st > 0; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> |
|
4222 modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = |
|
4223 (snd (tp ! (2 * (st - Suc 0) + b)))" |
|
4224 proof(induct st arbitrary: tp, simp) |
|
4225 fix st tp |
|
4226 assume ind: |
|
4227 "\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> |
|
4228 modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = |
|
4229 snd (tp ! (2 * (st - Suc 0) + b))" |
|
4230 and h: |
|
4231 "Suc st \<le> length (tp::tprog) div 2" |
|
4232 "0 < Suc st" |
|
4233 "b = 1 \<or> b = 0" |
|
4234 show "modify_tprog tp ! Suc (4 * (Suc st - Suc 0) + 2 * b) = |
|
4235 snd (tp ! (2 * (Suc st - Suc 0) + b))" |
|
4236 proof(cases "st = 0") |
|
4237 case True |
|
4238 thus "?thesis" |
|
4239 using h |
|
4240 apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) |
|
4241 apply(case_tac list, simp, case_tac ab, |
|
4242 simp add: modify_tprog.simps, auto) |
|
4243 done |
|
4244 next |
|
4245 case False |
|
4246 assume g: "st \<noteq> 0" |
|
4247 hence "\<exists> aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" |
|
4248 using h |
|
4249 apply(case_tac tp, simp, case_tac list, simp, simp) |
|
4250 done |
|
4251 from this obtain aa ab ba bb tp' where g1: |
|
4252 "tp = (aa, ab) # (ba, bb) # tp'" by blast |
|
4253 hence g2: |
|
4254 "modify_tprog tp' ! Suc (4 * (st - Suc 0) + 2 * b) = |
|
4255 snd (tp' ! (2 * (st - Suc 0) + b))" |
|
4256 apply(rule_tac ind) |
|
4257 using h g by auto |
|
4258 thus "?thesis" |
|
4259 using g1 g |
|
4260 apply(case_tac st, simp, simp) |
|
4261 done |
|
4262 qed |
|
4263 qed |
|
4264 |
|
4265 lemma fetch_state_eq: |
|
4266 "\<lbrakk>block_map b = scan r; |
|
4267 fetch tp st b = (nact, ns); |
|
4268 st \<le> length tp div 2\<rbrakk> \<Longrightarrow> newstat (code tp) st r = ns" |
|
4269 proof(simp add: newstat.simps, auto) |
|
4270 let ?i = "Suc (4 * (st - Suc 0) + 2 * (r mod 2))" |
|
4271 assume h: "block_map b = r mod 2" "fetch tp st b = |
|
4272 (nact, ns)" "st \<le> length tp div 2" "0 < st" |
|
4273 have "?i < length (modify_tprog tp)" |
|
4274 proof - |
|
4275 have "length (modify_tprog tp) = 2 * length tp" |
|
4276 apply(simp add: length_modify) |
|
4277 done |
|
4278 thus "?thesis" |
|
4279 using h |
|
4280 by(auto) |
|
4281 qed |
|
4282 hence "Entry (godel_code (modify_tprog tp)) (?i) = |
|
4283 (modify_tprog tp) ! ?i" |
|
4284 by(erule_tac godel_decode) |
|
4285 moreover have |
|
4286 "modify_tprog tp ! ?i = |
|
4287 (snd (tp ! (2 * (st - Suc 0) + r mod 2)))" |
|
4288 apply(rule_tac modify_tprog_fetch_state) |
|
4289 using h |
|
4290 by(auto) |
|
4291 moreover have "(snd (tp ! (2 * (st - Suc 0) + r mod 2))) = ns" |
|
4292 using h |
|
4293 apply(simp add: fetch.simps nth_of.simps) |
|
4294 apply(case_tac b, auto simp: block_map.simps nth_of.simps |
|
4295 split: if_splits) |
|
4296 done |
|
4297 ultimately show "Entry (godel_code (modify_tprog tp)) (?i) |
|
4298 = ns" |
|
4299 by simp |
|
4300 qed |
|
4301 |
|
4302 |
|
4303 lemma [intro!]: |
|
4304 "\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> trpl a b c = trpl a' b' c'" |
|
4305 by simp |
|
4306 |
|
4307 lemma [simp]: "bl2wc [Bk] = 0" |
|
4308 by(simp add: bl2wc.simps bl2nat.simps) |
|
4309 |
|
4310 lemma bl2nat_double: "bl2nat xs (Suc n) = 2 * bl2nat xs n" |
|
4311 proof(induct xs arbitrary: n) |
|
4312 case Nil thus "?case" |
|
4313 by(simp add: bl2nat.simps) |
|
4314 next |
|
4315 case (Cons x xs) thus "?case" |
|
4316 proof - |
|
4317 assume ind: "\<And>n. bl2nat xs (Suc n) = 2 * bl2nat xs n " |
|
4318 show "bl2nat (x # xs) (Suc n) = 2 * bl2nat (x # xs) n" |
|
4319 proof(cases x) |
|
4320 case Bk thus "?thesis" |
|
4321 apply(simp add: bl2nat.simps) |
|
4322 using ind[of "Suc n"] by simp |
|
4323 next |
|
4324 case Oc thus "?thesis" |
|
4325 apply(simp add: bl2nat.simps) |
|
4326 using ind[of "Suc n"] by simp |
|
4327 qed |
|
4328 qed |
|
4329 qed |
|
4330 |
|
4331 |
|
4332 lemma [simp]: "c \<noteq> [] \<Longrightarrow> 2 * bl2wc (tl c) = bl2wc c - bl2wc c mod 2 " |
|
4333 apply(case_tac c, simp, case_tac a) |
|
4334 apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4335 done |
|
4336 |
|
4337 lemma [simp]: |
|
4338 "c \<noteq> [] \<Longrightarrow> bl2wc (Oc # tl c) = Suc (bl2wc c) - bl2wc c mod 2 " |
|
4339 apply(case_tac c, simp, case_tac a) |
|
4340 apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4341 done |
|
4342 |
|
4343 lemma [simp]: "bl2wc (Bk # c) = 2*bl2wc (c)" |
|
4344 apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4345 done |
|
4346 |
|
4347 lemma [simp]: "bl2wc [Oc] = Suc 0" |
|
4348 by(simp add: bl2wc.simps bl2nat.simps) |
|
4349 |
|
4350 lemma [simp]: "b \<noteq> [] \<Longrightarrow> bl2wc (tl b) = bl2wc b div 2" |
|
4351 apply(case_tac b, simp, case_tac a) |
|
4352 apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4353 done |
|
4354 |
|
4355 lemma [simp]: "b \<noteq> [] \<Longrightarrow> bl2wc ([hd b]) = bl2wc b mod 2" |
|
4356 apply(case_tac b, simp, case_tac a) |
|
4357 apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4358 done |
|
4359 |
|
4360 lemma [simp]: "\<lbrakk>b \<noteq> []; c \<noteq> []\<rbrakk> \<Longrightarrow> bl2wc (hd b # c) = 2 * bl2wc c + bl2wc b mod 2" |
|
4361 apply(case_tac b, simp, case_tac a) |
|
4362 apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4363 done |
|
4364 |
|
4365 lemma [simp]: " 2 * (bl2wc c div 2) = bl2wc c - bl2wc c mod 2" |
|
4366 by(simp add: mult_div_cancel) |
|
4367 |
|
4368 lemma [simp]: "bl2wc (Oc # list) mod 2 = Suc 0" |
|
4369 by(simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4370 |
|
4371 |
|
4372 declare code.simps[simp del] |
|
4373 declare nth_of.simps[simp del] |
|
4374 declare new_tape.simps[simp del] |
|
4375 |
|
4376 text {* |
|
4377 The lemma relates the one step execution of TMs with the interpreter function @{text "rec_newconf"}. |
|
4378 *} |
|
4379 lemma rec_t_eq_step: |
|
4380 "(\<lambda> (s, l, r). s \<le> length tp div 2) c \<Longrightarrow> |
|
4381 trpl_code (tstep c tp) = |
|
4382 rec_exec rec_newconf [code tp, trpl_code c]" |
|
4383 apply(cases c, auto simp: tstep.simps) |
|
4384 proof(case_tac "fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", |
|
4385 simp add: newconf.simps trpl_code.simps) |
|
4386 fix a b c aa ba |
|
4387 assume h: "(a::nat) \<le> length tp div 2" |
|
4388 "fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, ba)" |
|
4389 moreover hence "actn (code tp) a (bl2wc c) = action_map aa" |
|
4390 apply(rule_tac b = "(case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)" |
|
4391 in fetch_action_eq, auto) |
|
4392 apply(auto split: list.splits) |
|
4393 apply(case_tac ab, auto) |
|
4394 done |
|
4395 moreover from h have "(newstat (code tp) a (bl2wc c)) = ba" |
|
4396 apply(rule_tac b = "(case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)" |
|
4397 in fetch_state_eq, auto split: list.splits) |
|
4398 apply(case_tac ab, auto) |
|
4399 done |
|
4400 ultimately show |
|
4401 "trpl_code (ba, new_tape aa (b, c)) = |
|
4402 trpl (newleft (bl2wc b) (bl2wc c) (actn (code tp) a (bl2wc c))) |
|
4403 (newstat (code tp) a (bl2wc c)) (newrght (bl2wc b) (bl2wc c) |
|
4404 (actn (code tp) a (bl2wc c)))" |
|
4405 by(auto simp: new_tape.simps trpl_code.simps |
|
4406 newleft.simps newrght.simps split: taction.splits) |
|
4407 qed |
|
4408 |
|
4409 lemma [simp]: "a\<^bsup>0\<^esup> = []" |
|
4410 apply(simp add: exp_zero) |
|
4411 done |
|
4412 lemma [simp]: "bl2nat (Oc # Oc\<^bsup>x\<^esup>) 0 = (2 * 2 ^ x - Suc 0)" |
|
4413 apply(induct x) |
|
4414 apply(simp add: bl2nat.simps) |
|
4415 apply(simp add: bl2nat.simps bl2nat_double exp_ind_def) |
|
4416 done |
|
4417 |
|
4418 lemma [simp]: "bl2nat (Oc\<^bsup>y\<^esup>) 0 = 2^y - Suc 0" |
|
4419 apply(induct y, auto simp: bl2nat.simps exp_ind_def bl2nat_double) |
|
4420 apply(case_tac "(2::nat)^y", auto) |
|
4421 done |
|
4422 |
|
4423 lemma [simp]: "bl2nat (Bk\<^bsup>l\<^esup>) n = 0" |
|
4424 apply(induct l, auto simp: bl2nat.simps bl2nat_double exp_ind_def) |
|
4425 done |
|
4426 |
|
4427 lemma bl2nat_cons_bk: "bl2nat (ks @ [Bk]) 0 = bl2nat ks 0" |
|
4428 apply(induct ks, auto simp: bl2nat.simps split: block.splits) |
|
4429 apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
|
4430 done |
|
4431 |
|
4432 lemma bl2nat_cons_oc: |
|
4433 "bl2nat (ks @ [Oc]) 0 = bl2nat ks 0 + 2 ^ length ks" |
|
4434 apply(induct ks, auto simp: bl2nat.simps split: block.splits) |
|
4435 apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
|
4436 done |
|
4437 |
|
4438 lemma bl2nat_append: |
|
4439 "bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs) " |
|
4440 proof(induct "length xs" arbitrary: xs ys, simp add: bl2nat.simps) |
|
4441 fix x xs ys |
|
4442 assume ind: |
|
4443 "\<And>xs ys. x = length xs \<Longrightarrow> |
|
4444 bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs)" |
|
4445 and h: "Suc x = length (xs::block list)" |
|
4446 have "\<exists> ks k. xs = ks @ [k]" |
|
4447 apply(rule_tac x = "butlast xs" in exI, |
|
4448 rule_tac x = "last xs" in exI) |
|
4449 using h |
|
4450 apply(case_tac xs, auto) |
|
4451 done |
|
4452 from this obtain ks k where "xs = ks @ [k]" by blast |
|
4453 moreover hence |
|
4454 "bl2nat (ks @ (k # ys)) 0 = bl2nat ks 0 + |
|
4455 bl2nat (k # ys) (length ks)" |
|
4456 apply(rule_tac ind) using h by simp |
|
4457 ultimately show "bl2nat (xs @ ys) 0 = |
|
4458 bl2nat xs 0 + bl2nat ys (length xs)" |
|
4459 apply(case_tac k, simp_all add: bl2nat.simps) |
|
4460 apply(simp_all only: bl2nat_cons_bk bl2nat_cons_oc) |
|
4461 done |
|
4462 qed |
|
4463 |
|
4464 lemma bl2nat_exp: "n \<noteq> 0 \<Longrightarrow> bl2nat bl n = 2^n * bl2nat bl 0" |
|
4465 apply(induct bl) |
|
4466 apply(auto simp: bl2nat.simps) |
|
4467 apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
|
4468 done |
|
4469 |
|
4470 lemma nat_minus_eq: "\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> a - c = b - d" |
|
4471 by auto |
|
4472 |
|
4473 lemma tape_of_nat_list_butlast_last: |
|
4474 "ys \<noteq> [] \<Longrightarrow> <ys @ [y]> = <ys> @ Bk # Oc\<^bsup>Suc y\<^esup>" |
|
4475 apply(induct ys, simp, simp) |
|
4476 apply(case_tac "ys = []", simp add: tape_of_nl_abv |
|
4477 tape_of_nat_list.simps) |
|
4478 apply(simp) |
|
4479 done |
|
4480 |
|
4481 lemma listsum2_append: |
|
4482 "\<lbrakk>n \<le> length xs\<rbrakk> \<Longrightarrow> listsum2 (xs @ ys) n = listsum2 xs n" |
|
4483 apply(induct n) |
|
4484 apply(auto simp: listsum2.simps nth_append) |
|
4485 done |
|
4486 |
|
4487 lemma strt'_append: |
|
4488 "\<lbrakk>n \<le> length xs\<rbrakk> \<Longrightarrow> strt' xs n = strt' (xs @ ys) n" |
|
4489 proof(induct n arbitrary: xs ys) |
|
4490 fix xs ys |
|
4491 show "strt' xs 0 = strt' (xs @ ys) 0" by(simp add: strt'.simps) |
|
4492 next |
|
4493 fix n xs ys |
|
4494 assume ind: |
|
4495 "\<And> xs ys. n \<le> length xs \<Longrightarrow> strt' xs n = strt' (xs @ ys) n" |
|
4496 and h: "Suc n \<le> length (xs::nat list)" |
|
4497 show "strt' xs (Suc n) = strt' (xs @ ys) (Suc n)" |
|
4498 using ind[of xs ys] h |
|
4499 apply(simp add: strt'.simps nth_append listsum2_append) |
|
4500 done |
|
4501 qed |
|
4502 |
|
4503 lemma length_listsum2_eq: |
|
4504 "\<lbrakk>length (ys::nat list) = k\<rbrakk> |
|
4505 \<Longrightarrow> length (<ys>) = listsum2 (map Suc ys) k + k - 1" |
|
4506 apply(induct k arbitrary: ys, simp_all add: listsum2.simps) |
|
4507 apply(subgoal_tac "\<exists> xs x. ys = xs @ [x]", auto) |
|
4508 proof - |
|
4509 fix xs x |
|
4510 assume ind: "\<And>ys. length ys = length xs \<Longrightarrow> length (<ys>) |
|
4511 = listsum2 (map Suc ys) (length xs) + |
|
4512 length (xs::nat list) - Suc 0" |
|
4513 have "length (<xs>) |
|
4514 = listsum2 (map Suc xs) (length xs) + length xs - Suc 0" |
|
4515 apply(rule_tac ind, simp) |
|
4516 done |
|
4517 thus "length (<xs @ [x]>) = |
|
4518 Suc (listsum2 (map Suc xs @ [Suc x]) (length xs) + x + length xs)" |
|
4519 apply(case_tac "xs = []") |
|
4520 apply(simp add: tape_of_nl_abv listsum2.simps |
|
4521 tape_of_nat_list.simps) |
|
4522 apply(simp add: tape_of_nat_list_butlast_last) |
|
4523 using listsum2_append[of "length xs" "map Suc xs" "[Suc x]"] |
|
4524 apply(simp) |
|
4525 done |
|
4526 next |
|
4527 fix k ys |
|
4528 assume "length ys = Suc k" |
|
4529 thus "\<exists>xs x. ys = xs @ [x]" |
|
4530 apply(rule_tac x = "butlast ys" in exI, |
|
4531 rule_tac x = "last ys" in exI) |
|
4532 apply(case_tac ys, auto) |
|
4533 done |
|
4534 qed |
|
4535 |
|
4536 lemma tape_of_nat_list_length: |
|
4537 "length (<(ys::nat list)>) = |
|
4538 listsum2 (map Suc ys) (length ys) + length ys - 1" |
|
4539 using length_listsum2_eq[of ys "length ys"] |
|
4540 apply(simp) |
|
4541 done |
|
4542 |
|
4543 |
|
4544 |
|
4545 lemma [simp]: |
|
4546 "trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp 0) = |
|
4547 rec_exec rec_conf [code tp, bl2wc (<lm>), 0]" |
|
4548 apply(simp add: steps.simps rec_exec.simps conf_lemma conf.simps |
|
4549 inpt.simps trpl_code.simps bl2wc.simps) |
|
4550 done |
|
4551 |
|
4552 text {* |
|
4553 The following lemma relates the multi-step interpreter function @{text "rec_conf"} |
|
4554 with the multi-step execution of TMs. |
|
4555 *} |
|
4556 lemma rec_t_eq_steps: |
|
4557 "turing_basic.t_correct tp \<Longrightarrow> |
|
4558 trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp) = |
|
4559 rec_exec rec_conf [code tp, bl2wc (<lm>), stp]" |
|
4560 proof(induct stp) |
|
4561 case 0 thus "?case" by(simp) |
|
4562 next |
|
4563 case (Suc n) thus "?case" |
|
4564 proof - |
|
4565 assume ind: |
|
4566 "t_correct tp \<Longrightarrow> trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp n) |
|
4567 = rec_exec rec_conf [code tp, bl2wc (<lm>), n]" |
|
4568 and h: "t_correct tp" |
|
4569 show |
|
4570 "trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp (Suc n)) = |
|
4571 rec_exec rec_conf [code tp, bl2wc (<lm>), Suc n]" |
|
4572 proof(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp n", |
|
4573 simp only: tstep_red conf_lemma conf.simps) |
|
4574 fix a b c |
|
4575 assume g: "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp n = (a, b, c) " |
|
4576 hence "conf (code tp) (bl2wc (<lm>)) n= trpl_code (a, b, c)" |
|
4577 using ind h |
|
4578 apply(simp add: conf_lemma) |
|
4579 done |
|
4580 moreover hence |
|
4581 "trpl_code (tstep (a, b, c) tp) = |
|
4582 rec_exec rec_newconf [code tp, trpl_code (a, b, c)]" |
|
4583 apply(rule_tac rec_t_eq_step) |
|
4584 using h g |
|
4585 apply(simp add: s_keep) |
|
4586 done |
|
4587 ultimately show |
|
4588 "trpl_code (tstep (a, b, c) tp) = |
|
4589 newconf (code tp) (conf (code tp) (bl2wc (<lm>)) n)" |
|
4590 by(simp add: newconf_lemma) |
|
4591 qed |
|
4592 qed |
|
4593 qed |
|
4594 |
|
4595 lemma [simp]: "bl2wc (Bk\<^bsup>m\<^esup>) = 0" |
|
4596 apply(induct m) |
|
4597 apply(simp, simp) |
|
4598 done |
|
4599 |
|
4600 lemma [simp]: "bl2wc (Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>) = bl2wc (Oc\<^bsup>rs\<^esup>)" |
|
4601 apply(induct rs, simp, |
|
4602 simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
|
4603 done |
|
4604 |
|
4605 lemma lg_power: "x > Suc 0 \<Longrightarrow> lg (x ^ rs) x = rs" |
|
4606 proof(simp add: lg.simps, auto) |
|
4607 fix xa |
|
4608 assume h: "Suc 0 < x" |
|
4609 show "Max {ya. ya \<le> x ^ rs \<and> lgR [x ^ rs, x, ya]} = rs" |
|
4610 apply(rule_tac Max_eqI, simp_all add: lgR.simps) |
|
4611 apply(simp add: h) |
|
4612 using x_less_exp[of x rs] h |
|
4613 apply(simp) |
|
4614 done |
|
4615 next |
|
4616 assume "\<not> Suc 0 < x ^ rs" "Suc 0 < x" |
|
4617 thus "rs = 0" |
|
4618 apply(case_tac "x ^ rs", simp, simp) |
|
4619 done |
|
4620 next |
|
4621 assume "Suc 0 < x" "\<forall>xa. \<not> lgR [x ^ rs, x, xa]" |
|
4622 thus "rs = 0" |
|
4623 apply(simp only:lgR.simps) |
|
4624 apply(erule_tac x = rs in allE, simp) |
|
4625 done |
|
4626 qed |
|
4627 |
|
4628 text {* |
|
4629 The following lemma relates execution of TMs with |
|
4630 the multi-step interpreter function @{text "rec_nonstop"}. Note, |
|
4631 @{text "rec_nonstop"} is constructed using @{text "rec_conf"}. |
|
4632 *} |
|
4633 lemma nonstop_t_eq: |
|
4634 "\<lbrakk>steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>); |
|
4635 turing_basic.t_correct tp; |
|
4636 rs > 0\<rbrakk> |
|
4637 \<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = 0" |
|
4638 proof(simp add: nonstop_lemma nonstop.simps nstd.simps) |
|
4639 assume h: "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
|
4640 and tc_t: "turing_basic.t_correct tp" "rs > 0" |
|
4641 have g: "rec_exec rec_conf [code tp, bl2wc (<lm>), stp] = |
|
4642 trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)" |
|
4643 using rec_t_eq_steps[of tp l lm stp] tc_t h |
|
4644 by(simp) |
|
4645 thus "\<not> NSTD (conf (code tp) (bl2wc (<lm>)) stp)" |
|
4646 proof(auto simp: NSTD.simps) |
|
4647 show "stat (conf (code tp) (bl2wc (<lm>)) stp) = 0" |
|
4648 using g |
|
4649 by(auto simp: conf_lemma trpl_code.simps) |
|
4650 next |
|
4651 show "left (conf (code tp) (bl2wc (<lm>)) stp) = 0" |
|
4652 using g |
|
4653 by(simp add: conf_lemma trpl_code.simps) |
|
4654 next |
|
4655 show "rght (conf (code tp) (bl2wc (<lm>)) stp) = |
|
4656 2 ^ lg (Suc (rght (conf (code tp) (bl2wc (<lm>)) stp))) 2 - Suc 0" |
|
4657 using g h |
|
4658 proof(simp add: conf_lemma trpl_code.simps) |
|
4659 have "2 ^ lg (Suc (bl2wc (Oc\<^bsup>rs\<^esup>))) 2 = Suc (bl2wc (Oc\<^bsup>rs\<^esup>))" |
|
4660 apply(simp add: bl2wc.simps lg_power) |
|
4661 done |
|
4662 thus "bl2wc (Oc\<^bsup>rs\<^esup>) = 2 ^ lg (Suc (bl2wc (Oc\<^bsup>rs\<^esup>))) 2 - Suc 0" |
|
4663 apply(simp) |
|
4664 done |
|
4665 qed |
|
4666 next |
|
4667 show "0 < rght (conf (code tp) (bl2wc (<lm>)) stp)" |
|
4668 using g h tc_t |
|
4669 apply(simp add: conf_lemma trpl_code.simps bl2wc.simps |
|
4670 bl2nat.simps) |
|
4671 apply(case_tac rs, simp, simp add: bl2nat.simps) |
|
4672 done |
|
4673 qed |
|
4674 qed |
|
4675 |
|
4676 lemma [simp]: "actn m 0 r = 4" |
|
4677 by(simp add: actn.simps) |
|
4678 |
|
4679 lemma [simp]: "newstat m 0 r = 0" |
|
4680 by(simp add: newstat.simps) |
|
4681 |
|
4682 declare exp_def[simp del] |
|
4683 |
|
4684 lemma halt_least_step: |
|
4685 "\<lbrakk>steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs \<^esup> @ Bk\<^bsup>n\<^esup>); |
|
4686 turing_basic.t_correct tp; |
|
4687 0<rs\<rbrakk> \<Longrightarrow> |
|
4688 \<exists> stp. (nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
|
4689 (\<forall> stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp'))" |
|
4690 proof(induct stp, simp add: steps.simps, simp) |
|
4691 fix stp |
|
4692 assume ind: |
|
4693 "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) \<Longrightarrow> |
|
4694 \<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
|
4695 (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
|
4696 and h: |
|
4697 "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp (Suc stp) = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
|
4698 "turing_basic.t_correct tp" |
|
4699 "0 < rs" |
|
4700 from h show |
|
4701 "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 |
|
4702 \<and> (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
|
4703 proof(simp add: tstep_red, |
|
4704 case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp", simp, |
|
4705 case_tac a, simp add: tstep_0) |
|
4706 assume "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
|
4707 thus "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
|
4708 (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
|
4709 apply(erule_tac ind) |
|
4710 done |
|
4711 next |
|
4712 fix a b c nat |
|
4713 assume "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (a, b, c)" |
|
4714 "a = Suc nat" |
|
4715 thus "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
|
4716 (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
|
4717 using h |
|
4718 apply(rule_tac x = "Suc stp" in exI, auto) |
|
4719 apply(drule_tac nonstop_t_eq, simp_all add: nonstop_lemma) |
|
4720 proof - |
|
4721 fix stp' |
|
4722 assume g:"steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (Suc nat, b, c)" |
|
4723 "nonstop (code tp) (bl2wc (<lm>)) stp' = 0" |
|
4724 thus "Suc stp \<le> stp'" |
|
4725 proof(case_tac "Suc stp \<le> stp'", simp, simp) |
|
4726 assume "\<not> Suc stp \<le> stp'" |
|
4727 hence "stp' \<le> stp" by simp |
|
4728 hence "\<not> isS0 (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp')" |
|
4729 using g |
|
4730 apply(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp'",auto, |
|
4731 simp add: isS0_def) |
|
4732 apply(subgoal_tac "\<exists> n. stp = stp' + n", |
|
4733 auto simp: steps_add steps_0) |
|
4734 apply(rule_tac x = "stp - stp'" in exI, simp) |
|
4735 done |
|
4736 hence "nonstop (code tp) (bl2wc (<lm>)) stp' = 1" |
|
4737 proof(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp'", |
|
4738 simp add: isS0_def nonstop.simps) |
|
4739 fix a b c |
|
4740 assume k: |
|
4741 "0 < a" "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp' = (a, b, c)" |
|
4742 thus " NSTD (conf (code tp) (bl2wc (<lm>)) stp')" |
|
4743 using rec_t_eq_steps[of tp l lm stp'] h |
|
4744 proof(simp add: conf_lemma) |
|
4745 assume "trpl_code (a, b, c) = conf (code tp) (bl2wc (<lm>)) stp'" |
|
4746 moreover have "NSTD (trpl_code (a, b, c))" |
|
4747 using k |
|
4748 apply(auto simp: trpl_code.simps NSTD.simps) |
|
4749 done |
|
4750 ultimately show "NSTD (conf (code tp) (bl2wc (<lm>)) stp')" by simp |
|
4751 qed |
|
4752 qed |
|
4753 thus "False" using g by simp |
|
4754 qed |
|
4755 qed |
|
4756 qed |
|
4757 qed |
|
4758 |
|
4759 lemma conf_trpl_ex: "\<exists> p q r. conf m (bl2wc (<lm>)) stp = trpl p q r" |
|
4760 apply(induct stp, auto simp: conf.simps inpt.simps trpl.simps |
|
4761 newconf.simps) |
|
4762 apply(rule_tac x = 0 in exI, rule_tac x = 1 in exI, |
|
4763 rule_tac x = "bl2wc (<lm>)" in exI) |
|
4764 apply(simp) |
|
4765 done |
|
4766 |
|
4767 lemma nonstop_rgt_ex: |
|
4768 "nonstop m (bl2wc (<lm>)) stpa = 0 \<Longrightarrow> \<exists> r. conf m (bl2wc (<lm>)) stpa = trpl 0 0 r" |
|
4769 apply(auto simp: nonstop.simps NSTD.simps split: if_splits) |
|
4770 using conf_trpl_ex[of m lm stpa] |
|
4771 apply(auto) |
|
4772 done |
|
4773 |
|
4774 lemma [elim]: "x > Suc 0 \<Longrightarrow> Max {u. x ^ u dvd x ^ r} = r" |
|
4775 proof(rule_tac Max_eqI) |
|
4776 assume "x > Suc 0" |
|
4777 thus "finite {u. x ^ u dvd x ^ r}" |
|
4778 apply(rule_tac finite_power_dvd, auto) |
|
4779 done |
|
4780 next |
|
4781 fix y |
|
4782 assume "Suc 0 < x" "y \<in> {u. x ^ u dvd x ^ r}" |
|
4783 thus "y \<le> r" |
|
4784 apply(case_tac "y\<le> r", simp) |
|
4785 apply(subgoal_tac "\<exists> d. y = r + d") |
|
4786 apply(auto simp: power_add) |
|
4787 apply(rule_tac x = "y - r" in exI, simp) |
|
4788 done |
|
4789 next |
|
4790 show "r \<in> {u. x ^ u dvd x ^ r}" by simp |
|
4791 qed |
|
4792 |
|
4793 lemma lo_power: "x > Suc 0 \<Longrightarrow> lo (x ^ r) x = r" |
|
4794 apply(auto simp: lo.simps loR.simps mod_dvd_simp) |
|
4795 apply(case_tac "x^r", simp_all) |
|
4796 done |
|
4797 |
|
4798 lemma lo_rgt: "lo (trpl 0 0 r) (Pi 2) = r" |
|
4799 apply(simp add: trpl.simps lo_power) |
|
4800 done |
|
4801 |
|
4802 lemma conf_keep: |
|
4803 "conf m lm stp = trpl 0 0 r \<Longrightarrow> |
|
4804 conf m lm (stp + n) = trpl 0 0 r" |
|
4805 apply(induct n) |
|
4806 apply(auto simp: conf.simps newconf.simps newleft.simps |
|
4807 newrght.simps rght.simps lo_rgt) |
|
4808 done |
|
4809 |
|
4810 lemma halt_state_keep_steps_add: |
|
4811 "\<lbrakk>nonstop m (bl2wc (<lm>)) stpa = 0\<rbrakk> \<Longrightarrow> |
|
4812 conf m (bl2wc (<lm>)) stpa = conf m (bl2wc (<lm>)) (stpa + n)" |
|
4813 apply(drule_tac nonstop_rgt_ex, auto simp: conf_keep) |
|
4814 done |
|
4815 |
|
4816 lemma halt_state_keep: |
|
4817 "\<lbrakk>nonstop m (bl2wc (<lm>)) stpa = 0; nonstop m (bl2wc (<lm>)) stpb = 0\<rbrakk> \<Longrightarrow> |
|
4818 conf m (bl2wc (<lm>)) stpa = conf m (bl2wc (<lm>)) stpb" |
|
4819 apply(case_tac "stpa > stpb") |
|
4820 using halt_state_keep_steps_add[of m lm stpb "stpa - stpb"] |
|
4821 apply simp |
|
4822 using halt_state_keep_steps_add[of m lm stpa "stpb - stpa"] |
|
4823 apply(simp) |
|
4824 done |
|
4825 |
|
4826 text {* |
|
4827 The correntess of @{text "rec_F"} which relates the interpreter function @{text "rec_F"} with the |
|
4828 execution of of TMs. |
|
4829 *} |
|
4830 lemma F_t_halt_eq: |
|
4831 "\<lbrakk>steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>); |
|
4832 turing_basic.t_correct tp; |
|
4833 0<rs\<rbrakk> |
|
4834 \<Longrightarrow> rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] (rs - Suc 0)" |
|
4835 apply(frule_tac halt_least_step, auto) |
|
4836 apply(frule_tac nonstop_t_eq, auto simp: nonstop_lemma) |
|
4837 using rec_t_eq_steps[of tp l lm stp] |
|
4838 apply(simp add: conf_lemma) |
|
4839 proof - |
|
4840 fix stpa |
|
4841 assume h: |
|
4842 "nonstop (code tp) (bl2wc (<lm>)) stpa = 0" |
|
4843 "\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stpa \<le> stp'" |
|
4844 "nonstop (code tp) (bl2wc (<lm>)) stp = 0" |
|
4845 "trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) = conf (code tp) (bl2wc (<lm>)) stp" |
|
4846 "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
|
4847 hence g1: "conf (code tp) (bl2wc (<lm>)) stpa = trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
|
4848 using halt_state_keep[of "code tp" lm stpa stp] |
|
4849 by(simp) |
|
4850 moreover have g2: |
|
4851 "rec_calc_rel rec_halt [code tp, (bl2wc (<lm>))] stpa" |
|
4852 using h |
|
4853 apply(simp add: halt_lemma nonstop_lemma, auto) |
|
4854 done |
|
4855 show |
|
4856 "rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] (rs - Suc 0)" |
|
4857 proof - |
|
4858 have |
|
4859 "rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] |
|
4860 (valu (rght (conf (code tp) (bl2wc (<lm>)) stpa)))" |
|
4861 apply(rule F_lemma) using g2 h by auto |
|
4862 moreover have |
|
4863 "valu (rght (conf (code tp) (bl2wc (<lm>)) stpa)) = rs - Suc 0" |
|
4864 using g1 |
|
4865 apply(simp add: valu.simps trpl_code.simps |
|
4866 bl2wc.simps bl2nat_append lg_power) |
|
4867 done |
|
4868 ultimately show "?thesis" by simp |
|
4869 qed |
|
4870 qed |
|
4871 |
|
4872 |
|
4873 end |