1 theory UF_Rec |
1 theory UF_Rec |
2 imports Recs Turing_Hoare |
2 imports Recs Turing_Hoare |
3 begin |
3 begin |
4 |
4 |
5 section {* Universal Function *} |
5 |
6 |
6 |
7 |
7 |
8 |
8 |
9 text{* coding of the configuration *} |
9 section {* Universal Function in HOL *} |
10 |
10 |
11 text {* |
11 text {* |
12 @{text "Entry sr i"} returns the @{text "i"}-th entry of a list of natural |
12 @{text "Entry sr i"} returns the @{text "i"}-th entry of a list of natural |
13 numbers encoded by number @{text "sr"} using Godel's coding. |
13 numbers encoded by number @{text "sr"} using Godel's coding. |
14 *} |
14 *} |
15 fun Entry where |
15 fun Entry where |
16 "Entry sr i = Lo sr (Pi (Suc i))" |
16 "Entry sr i = Lo sr (Pi (Suc i))" |
17 |
17 |
18 definition |
|
19 "rec_entry = CN rec_lo [Id 2 0, CN rec_pi [CN S [Id 2 1]]]" |
|
20 |
|
21 lemma entry_lemma [simp]: |
|
22 "rec_eval rec_entry [sr, i] = Entry sr i" |
|
23 by(simp add: rec_entry_def) |
|
24 |
|
25 |
|
26 fun Listsum2 :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
18 fun Listsum2 :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
27 where |
19 where |
28 "Listsum2 xs 0 = 0" |
20 "Listsum2 xs 0 = 0" |
29 | "Listsum2 xs (Suc n) = Listsum2 xs n + xs ! n" |
21 | "Listsum2 xs (Suc n) = Listsum2 xs n + xs ! n" |
30 |
22 |
31 fun rec_listsum2 :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
|
32 where |
|
33 "rec_listsum2 vl 0 = CN Z [Id vl 0]" |
|
34 | "rec_listsum2 vl (Suc n) = CN rec_add [rec_listsum2 vl n, Id vl n]" |
|
35 |
|
36 lemma listsum2_lemma [simp]: |
|
37 "length xs = vl \<Longrightarrow> rec_eval (rec_listsum2 vl n) xs = Listsum2 xs n" |
|
38 by (induct n) (simp_all) |
|
39 |
|
40 text {* |
23 text {* |
41 @{text "Strt"} corresponds to the @{text "strt"} function on page 90 of the |
24 @{text "Strt"} corresponds to the @{text "strt"} function on page 90 of the |
42 B book, but this definition generalises the original one to deal with multiple |
25 B book, but this definition generalises the original one in order to deal |
43 input arguments. |
26 with multiple input arguments. *} |
44 *} |
27 |
45 fun Strt' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
28 fun Strt' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
46 where |
29 where |
47 "Strt' xs 0 = 0" |
30 "Strt' xs 0 = 0" |
48 | "Strt' xs (Suc n) = (let dbound = Listsum2 xs n + n |
31 | "Strt' xs (Suc n) = (let dbound = Listsum2 xs n + n |
49 in Strt' xs n + (2 ^ (xs ! n + dbound) - 2 ^ dbound))" |
32 in Strt' xs n + (2 ^ (xs ! n + dbound) - 2 ^ dbound))" |
50 |
33 |
51 fun Strt :: "nat list \<Rightarrow> nat" |
34 fun Strt :: "nat list \<Rightarrow> nat" |
52 where |
35 where |
53 "Strt xs = (let ys = map Suc xs in Strt' ys (length ys))" |
36 "Strt xs = (let ys = map Suc xs in Strt' ys (length ys))" |
|
37 |
|
38 text {* The @{text "Scan"} function on page 90 of B book. *} |
|
39 fun Scan :: "nat \<Rightarrow> nat" |
|
40 where |
|
41 "Scan r = r mod 2" |
|
42 |
|
43 text {* The @{text Newleft} and @{text Newright} functions on page 91 of B book. *} |
|
44 |
|
45 fun Newleft :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
46 where |
|
47 "Newleft p r a = (if a = 0 \<or> a = 1 then p |
|
48 else if a = 2 then p div 2 |
|
49 else if a = 3 then 2 * p + r mod 2 |
|
50 else p)" |
|
51 |
|
52 fun Newright :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
53 where |
|
54 "Newright p r a = (if a = 0 then r - Scan r |
|
55 else if a = 1 then r + 1 - Scan r |
|
56 else if a = 2 then 2 * r + p mod 2 |
|
57 else if a = 3 then r div 2 |
|
58 else r)" |
|
59 |
|
60 text {* |
|
61 The @{text "Actn"} function given on page 92 of B book, which is used to |
|
62 fetch Turing Machine intructions. In @{text "Actn m q r"}, @{text "m"} is |
|
63 the Goedel coding of a Turing Machine, @{text "q"} is the current state of |
|
64 Turing Machine, @{text "r"} is the right number of Turing Machine tape. *} |
|
65 |
|
66 fun Actn :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
67 where |
|
68 "Actn m q r = (if q \<noteq> 0 then Entry m (4 * (q - 1) + 2 * Scan r) else 4)" |
|
69 |
|
70 fun Newstat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
71 where |
|
72 "Newstat m q r = (if q \<noteq> 0 then Entry m (4 * (q - 1) + 2 * Scan r + 1) else 0)" |
|
73 |
|
74 fun Trpl :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
75 where |
|
76 "Trpl p q r = (Pi 0) ^ p * (Pi 1) ^ q * (Pi 2) ^ r" |
|
77 |
|
78 fun Left where |
|
79 "Left c = Lo c (Pi 0)" |
|
80 |
|
81 fun Right where |
|
82 "Right c = Lo c (Pi 2)" |
|
83 |
|
84 fun Stat where |
|
85 "Stat c = Lo c (Pi 1)" |
|
86 |
|
87 lemma mod_dvd_simp: |
|
88 "(x mod y = (0::nat)) = (y dvd x)" |
|
89 by(auto simp add: dvd_def) |
|
90 |
|
91 lemma Trpl_Left [simp]: |
|
92 "Left (Trpl p q r) = p" |
|
93 apply(simp) |
|
94 apply(subst Lo_def) |
|
95 apply(subst dvd_eq_mod_eq_0[symmetric]) |
|
96 apply(simp) |
|
97 apply(auto) |
|
98 thm Lo_def |
|
99 thm mod_dvd_simp |
|
100 apply(simp add: left.simps trpl.simps lo.simps |
|
101 loR.simps mod_dvd_simp, auto simp: conf_decode1) |
|
102 apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", |
|
103 auto) |
|
104 apply(erule_tac x = l in allE, auto) |
|
105 |
|
106 |
|
107 fun Inpt :: "nat \<Rightarrow> nat list \<Rightarrow> nat" |
|
108 where |
|
109 "Inpt m xs = Trpl 0 1 (Strt xs)" |
|
110 |
|
111 fun Newconf :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
112 where |
|
113 "Newconf m c = Trpl (Newleft (Left c) (Right c) (Actn m (Stat c) (Right c))) |
|
114 (Newstat m (Stat c) (Right c)) |
|
115 (Newright (Left c) (Right c) (Actn m (Stat c) (Right c)))" |
|
116 |
|
117 text {* |
|
118 @{text "Conf k m r"} computes the TM configuration after @{text "k"} steps of execution |
|
119 of TM coded as @{text "m"} starting from the initial configuration where the left |
|
120 number equals @{text "0"}, right number equals @{text "r"}. *} |
|
121 |
|
122 fun Conf :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
123 where |
|
124 "Conf 0 m r = Trpl 0 1 r" |
|
125 | "Conf (Suc k) m r = Newconf m (Conf k m r)" |
|
126 |
|
127 text {* |
|
128 @{text "Nstd c"} returns true if the configuration coded |
|
129 by @{text "c"} is not a stardard final configuration. *} |
|
130 |
|
131 fun Nstd :: "nat \<Rightarrow> bool" |
|
132 where |
|
133 "Nstd c = (Stat c \<noteq> 0 \<or> |
|
134 Left c \<noteq> 0 \<or> |
|
135 Right c \<noteq> 2 ^ (Lg (Suc (Right c)) 2) - 1 \<or> |
|
136 Right c = 0)" |
|
137 |
|
138 |
|
139 text{* |
|
140 @{text "Nostop t m r"} means that afer @{text "t"} steps of |
|
141 execution the TM coded by @{text "m"} is not at a stardard |
|
142 final configuration. *} |
|
143 |
|
144 fun Nostop :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" |
|
145 where |
|
146 "Nostop t m r = Nstd (Conf t m r)" |
|
147 |
|
148 fun Value where |
|
149 "Value x = (Lg (Suc x) 2) - 1" |
|
150 |
|
151 text{* |
|
152 @{text "rec_halt"} is the recursive function calculating the steps a TM needs to execute before |
|
153 to reach a stardard final configuration. This recursive function is the only one |
|
154 using @{text "Mn"} combinator. So it is the only non-primitive recursive function |
|
155 needs to be used in the construction of the universal function @{text "rec_uf"}. *} |
|
156 |
|
157 fun Halt :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
158 where |
|
159 "Halt m r = (LEAST t. \<not> Nostop t m r)" |
|
160 |
|
161 fun UF :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
162 where |
|
163 "UF c m = Value (Right (Conf (Halt c m) c m))" |
|
164 |
|
165 |
|
166 section {* Coding of Turing Machines *} |
|
167 |
|
168 text {* |
|
169 The purpose of this section is to construct the coding function of Turing |
|
170 Machine, which is going to be named @{text "code"}. *} |
|
171 |
|
172 fun bl2nat :: "cell list \<Rightarrow> nat \<Rightarrow> nat" |
|
173 where |
|
174 "bl2nat [] n = 0" |
|
175 | "bl2nat (Bk # bl) n = bl2nat bl (Suc n)" |
|
176 | "bl2nat (Oc # bl) n = 2 ^ n + bl2nat bl (Suc n)" |
|
177 |
|
178 fun bl2wc :: "cell list \<Rightarrow> nat" |
|
179 where |
|
180 "bl2wc xs = bl2nat xs 0" |
|
181 |
|
182 lemma bl2nat_double [simp]: |
|
183 "bl2nat xs (Suc n) = 2 * bl2nat xs n" |
|
184 apply(induct xs arbitrary: n) |
|
185 apply(auto) |
|
186 apply(case_tac a) |
|
187 apply(auto) |
|
188 done |
|
189 |
|
190 lemma bl2nat_simps1 [simp]: |
|
191 shows "bl2nat (Bk \<up> y) n = 0" |
|
192 by (induct y) (auto) |
|
193 |
|
194 lemma bl2nat_simps2 [simp]: |
|
195 shows "bl2nat (Oc \<up> y) 0 = 2 ^ y - 1" |
|
196 apply(induct y) |
|
197 apply(auto) |
|
198 apply(case_tac "(2::nat)^ y") |
|
199 apply(auto) |
|
200 done |
|
201 |
|
202 fun Trpl_code :: "config \<Rightarrow> nat" |
|
203 where |
|
204 "Trpl_code (st, l, r) = Trpl (bl2wc l) st (bl2wc r)" |
|
205 |
|
206 fun action_map :: "action \<Rightarrow> nat" |
|
207 where |
|
208 "action_map W0 = 0" |
|
209 | "action_map W1 = 1" |
|
210 | "action_map L = 2" |
|
211 | "action_map R = 3" |
|
212 | "action_map Nop = 4" |
|
213 |
|
214 fun action_map_iff :: "nat \<Rightarrow> action" |
|
215 where |
|
216 "action_map_iff (0::nat) = W0" |
|
217 | "action_map_iff (Suc 0) = W1" |
|
218 | "action_map_iff (Suc (Suc 0)) = L" |
|
219 | "action_map_iff (Suc (Suc (Suc 0))) = R" |
|
220 | "action_map_iff n = Nop" |
|
221 |
|
222 fun block_map :: "cell \<Rightarrow> nat" |
|
223 where |
|
224 "block_map Bk = 0" |
|
225 | "block_map Oc = 1" |
|
226 |
|
227 fun Goedel_code' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
|
228 where |
|
229 "Goedel_code' [] n = 1" |
|
230 | "Goedel_code' (x # xs) n = (Pi n) ^ x * Goedel_code' xs (Suc n) " |
|
231 |
|
232 fun Goedel_code :: "nat list \<Rightarrow> nat" |
|
233 where |
|
234 "Goedel_code xs = 2 ^ (length xs) * (Goedel_code' xs 1)" |
|
235 |
|
236 fun modify_tprog :: "instr list \<Rightarrow> nat list" |
|
237 where |
|
238 "modify_tprog [] = []" |
|
239 | "modify_tprog ((a, s) # nl) = action_map a # s # modify_tprog nl" |
|
240 |
|
241 text {* @{text "Code tp"} gives the Goedel coding of TM program @{text "tp"}. *} |
|
242 fun Code :: "instr list \<Rightarrow> nat" |
|
243 where |
|
244 "Code tp = Goedel_code (modify_tprog tp)" |
|
245 |
|
246 |
|
247 section {* Universal Function as Recursive Functions *} |
|
248 |
|
249 definition |
|
250 "rec_entry = CN rec_lo [Id 2 0, CN rec_pi [CN S [Id 2 1]]]" |
|
251 |
|
252 fun rec_listsum2 :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
|
253 where |
|
254 "rec_listsum2 vl 0 = CN Z [Id vl 0]" |
|
255 | "rec_listsum2 vl (Suc n) = CN rec_add [rec_listsum2 vl n, Id vl n]" |
54 |
256 |
55 fun rec_strt' :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
257 fun rec_strt' :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
56 where |
258 where |
57 "rec_strt' xs 0 = Z" |
259 "rec_strt' xs 0 = Z" |
58 | "rec_strt' xs (Suc n) = |
260 | "rec_strt' xs (Suc n) = |
140 CN rec_if [condn 0, case0, |
296 CN rec_if [condn 0, case0, |
141 CN rec_if [condn 1, case1, |
297 CN rec_if [condn 1, case1, |
142 CN rec_if [condn 2, case2, |
298 CN rec_if [condn 2, case2, |
143 CN rec_if [condn 3, case3, Id 3 1]]]])" |
299 CN rec_if [condn 3, case3, Id 3 1]]]])" |
144 |
300 |
145 lemma newleft_lemma [simp]: |
|
146 "rec_eval rec_newleft [p, r, a] = Newleft p r a" |
|
147 by (simp add: rec_newleft_def Let_def) |
|
148 |
|
149 lemma newright_lemma [simp]: |
|
150 "rec_eval rec_newright [p, r, a] = Newright p r a" |
|
151 by (simp add: rec_newright_def Let_def) |
|
152 |
|
153 text {* |
|
154 The @{text "Actn"} function given on page 92 of B book, which is used to |
|
155 fetch Turing Machine intructions. In @{text "Actn m q r"}, @{text "m"} is |
|
156 the Goedel coding of a Turing Machine, @{text "q"} is the current state of |
|
157 Turing Machine, @{text "r"} is the right number of Turing Machine tape. |
|
158 *} |
|
159 fun Actn :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
160 where |
|
161 "Actn m q r = (if q \<noteq> 0 then Entry m (4 * (q - 1) + 2 * Scan r) else 4)" |
|
162 |
|
163 definition |
301 definition |
164 "rec_actn = (let add1 = CN rec_mult [constn 4, CN rec_pred [Id 3 1]] in |
302 "rec_actn = (let add1 = CN rec_mult [constn 4, CN rec_pred [Id 3 1]] in |
165 let add2 = CN rec_mult [constn 2, CN rec_scan [Id 3 2]] in |
303 let add2 = CN rec_mult [constn 2, CN rec_scan [Id 3 2]] in |
166 let entry = CN rec_entry [Id 3 0, CN rec_add [add1, add2]] |
304 let entry = CN rec_entry [Id 3 0, CN rec_add [add1, add2]] |
167 in CN rec_if [Id 3 1, entry, constn 4])" |
305 in CN rec_if [Id 3 1, entry, constn 4])" |
168 |
306 |
169 lemma actn_lemma [simp]: |
|
170 "rec_eval rec_actn [m, q, r] = Actn m q r" |
|
171 by (simp add: rec_actn_def) |
|
172 |
|
173 fun Newstat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
174 where |
|
175 "Newstat m q r = (if q \<noteq> 0 then Entry m (4 * (q - 1) + 2 * Scan r + 1) else 0)" |
|
176 |
|
177 definition rec_newstat :: "recf" |
307 definition rec_newstat :: "recf" |
178 where |
308 where |
179 "rec_newstat = (let add1 = CN rec_mult [constn 4, CN rec_pred [Id 3 1]] in |
309 "rec_newstat = (let add1 = CN rec_mult [constn 4, CN rec_pred [Id 3 1]] in |
180 let add2 = CN S [CN rec_mult [constn 2, CN rec_scan [Id 3 2]]] in |
310 let add2 = CN S [CN rec_mult [constn 2, CN rec_scan [Id 3 2]]] in |
181 let entry = CN rec_entry [Id 3 0, CN rec_add [add1, add2]] |
311 let entry = CN rec_entry [Id 3 0, CN rec_add [add1, add2]] |
182 in CN rec_if [Id 3 1, entry, Z])" |
312 in CN rec_if [Id 3 1, entry, Z])" |
183 |
313 |
184 lemma newstat_lemma [simp]: |
|
185 "rec_eval rec_newstat [m, q, r] = Newstat m q r" |
|
186 by (simp add: rec_newstat_def) |
|
187 |
|
188 |
|
189 fun Trpl :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
190 where |
|
191 "Trpl p q r = (Pi 0) ^ p * (Pi 1) ^ q * (Pi 2) ^ r" |
|
192 |
|
193 definition |
314 definition |
194 "rec_trpl = CN rec_mult [CN rec_mult |
315 "rec_trpl = CN rec_mult [CN rec_mult |
195 [CN rec_power [constn (Pi 0), Id 3 0], |
316 [CN rec_power [constn (Pi 0), Id 3 0], |
196 CN rec_power [constn (Pi 1), Id 3 1]], |
317 CN rec_power [constn (Pi 1), Id 3 1]], |
197 CN rec_power [constn (Pi 2), Id 3 2]]" |
318 CN rec_power [constn (Pi 2), Id 3 2]]" |
198 |
319 |
199 lemma trpl_lemma [simp]: |
|
200 "rec_eval rec_trpl [p, q, r] = Trpl p q r" |
|
201 by (simp add: rec_trpl_def) |
|
202 |
|
203 |
|
204 |
|
205 fun Left where |
|
206 "Left c = Lo c (Pi 0)" |
|
207 |
|
208 definition |
320 definition |
209 "rec_left = CN rec_lo [Id 1 0, constn (Pi 0)]" |
321 "rec_left = CN rec_lo [Id 1 0, constn (Pi 0)]" |
210 |
322 |
211 lemma left_lemma [simp]: |
|
212 "rec_eval rec_left [c] = Left c" |
|
213 by(simp add: rec_left_def) |
|
214 |
|
215 fun Right where |
|
216 "Right c = Lo c (Pi 2)" |
|
217 |
|
218 definition |
323 definition |
219 "rec_right = CN rec_lo [Id 1 0, constn (Pi 2)]" |
324 "rec_right = CN rec_lo [Id 1 0, constn (Pi 2)]" |
220 |
325 |
221 lemma right_lemma [simp]: |
|
222 "rec_eval rec_right [c] = Right c" |
|
223 by(simp add: rec_right_def) |
|
224 |
|
225 fun Stat where |
|
226 "Stat c = Lo c (Pi 1)" |
|
227 |
|
228 definition |
326 definition |
229 "rec_stat = CN rec_lo [Id 1 0, constn (Pi 1)]" |
327 "rec_stat = CN rec_lo [Id 1 0, constn (Pi 1)]" |
230 |
|
231 lemma stat_lemma [simp]: |
|
232 "rec_eval rec_stat [c] = Stat c" |
|
233 by(simp add: rec_stat_def) |
|
234 |
|
235 fun Inpt :: "nat \<Rightarrow> nat list \<Rightarrow> nat" |
|
236 where |
|
237 "Inpt m xs = Trpl 0 1 (Strt xs)" |
|
238 |
|
239 fun Newconf :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
240 where |
|
241 "Newconf m c = Trpl (Newleft (Left c) (Right c) (Actn m (Stat c) (Right c))) |
|
242 (Newstat m (Stat c) (Right c)) |
|
243 (Newright (Left c) (Right c) (Actn m (Stat c) (Right c)))" |
|
244 |
328 |
245 definition |
329 definition |
246 "rec_newconf = (let act = CN rec_actn [Id 2 0, CN rec_stat [Id 2 1], CN rec_right [Id 2 1]] in |
330 "rec_newconf = (let act = CN rec_actn [Id 2 0, CN rec_stat [Id 2 1], CN rec_right [Id 2 1]] in |
247 let left = CN rec_left [Id 2 1] in |
331 let left = CN rec_left [Id 2 1] in |
248 let right = CN rec_right [Id 2 1] in |
332 let right = CN rec_right [Id 2 1] in |
250 let one = CN rec_newleft [left, right, act] in |
334 let one = CN rec_newleft [left, right, act] in |
251 let two = CN rec_newstat [Id 2 0, stat, right] in |
335 let two = CN rec_newstat [Id 2 0, stat, right] in |
252 let three = CN rec_newright [left, right, act] |
336 let three = CN rec_newright [left, right, act] |
253 in CN rec_trpl [one, two, three])" |
337 in CN rec_trpl [one, two, three])" |
254 |
338 |
255 lemma newconf_lemma [simp]: |
|
256 "rec_eval rec_newconf [m, c] = Newconf m c" |
|
257 by (simp add: rec_newconf_def Let_def) |
|
258 |
|
259 text {* |
|
260 @{text "Conf k m r"} computes the TM configuration after @{text "k"} steps of execution |
|
261 of TM coded as @{text "m"} starting from the initial configuration where the left |
|
262 number equals @{text "0"}, right number equals @{text "r"}. |
|
263 *} |
|
264 fun Conf :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
265 where |
|
266 "Conf 0 m r = Trpl 0 1 r" |
|
267 | "Conf (Suc k) m r = Newconf m (Conf k m r)" |
|
268 |
|
269 definition |
339 definition |
270 "rec_conf = PR (CN rec_trpl [constn 0, constn 1, Id 2 1]) |
340 "rec_conf = PR (CN rec_trpl [constn 0, constn 1, Id 2 1]) |
271 (CN rec_newconf [Id 4 2 , Id 4 1])" |
341 (CN rec_newconf [Id 4 2 , Id 4 1])" |
272 |
|
273 lemma conf_lemma [simp]: |
|
274 "rec_eval rec_conf [k, m, r] = Conf k m r" |
|
275 by(induct k) (simp_all add: rec_conf_def) |
|
276 |
|
277 |
|
278 text {* |
|
279 @{text "Nstd c"} returns true if the configuration coded |
|
280 by @{text "c"} is not a stardard final configuration. |
|
281 *} |
|
282 fun Nstd :: "nat \<Rightarrow> bool" |
|
283 where |
|
284 "Nstd c = (Stat c \<noteq> 0 \<or> |
|
285 Left c \<noteq> 0 \<or> |
|
286 Right c \<noteq> 2 ^ (Lg (Suc (Right c)) 2) - 1 \<or> |
|
287 Right c = 0)" |
|
288 |
342 |
289 definition |
343 definition |
290 "rec_nstd = (let disj1 = CN rec_noteq [rec_stat, constn 0] in |
344 "rec_nstd = (let disj1 = CN rec_noteq [rec_stat, constn 0] in |
291 let disj2 = CN rec_noteq [rec_left, constn 0] in |
345 let disj2 = CN rec_noteq [rec_left, constn 0] in |
292 let rhs = CN rec_pred [CN rec_power [constn 2, CN rec_lg [CN S [rec_right], constn 2]]] in |
346 let rhs = CN rec_pred [CN rec_power [constn 2, CN rec_lg [CN S [rec_right], constn 2]]] in |
293 let disj3 = CN rec_noteq [rec_right, rhs] in |
347 let disj3 = CN rec_noteq [rec_right, rhs] in |
294 let disj4 = CN rec_eq [rec_right, constn 0] in |
348 let disj4 = CN rec_eq [rec_right, constn 0] in |
295 CN rec_disj [CN rec_disj [CN rec_disj [disj1, disj2], disj3], disj4])" |
349 CN rec_disj [CN rec_disj [CN rec_disj [disj1, disj2], disj3], disj4])" |
296 |
350 |
|
351 definition |
|
352 "rec_nostop = CN rec_nstd [rec_conf]" |
|
353 |
|
354 definition |
|
355 "rec_value = CN rec_pred [CN rec_lg [S, constn 2]]" |
|
356 |
|
357 definition |
|
358 "rec_halt = MN rec_nostop" |
|
359 |
|
360 definition |
|
361 "rec_uf = CN rec_value [CN rec_right [CN rec_conf [rec_halt, Id 2 0, Id 2 1]]]" |
|
362 |
|
363 |
|
364 |
|
365 section {* Correctness Proofs for Recursive Functions *} |
|
366 |
|
367 lemma entry_lemma [simp]: |
|
368 "rec_eval rec_entry [sr, i] = Entry sr i" |
|
369 by(simp add: rec_entry_def) |
|
370 |
|
371 lemma listsum2_lemma [simp]: |
|
372 "length xs = vl \<Longrightarrow> rec_eval (rec_listsum2 vl n) xs = Listsum2 xs n" |
|
373 by (induct n) (simp_all) |
|
374 |
|
375 lemma strt'_lemma [simp]: |
|
376 "length xs = vl \<Longrightarrow> rec_eval (rec_strt' vl n) xs = Strt' xs n" |
|
377 by (induct n) (simp_all add: Let_def) |
|
378 |
|
379 lemma map_suc: |
|
380 "map (\<lambda>x. Suc (xs ! x)) [0..<length xs] = map Suc xs" |
|
381 proof - |
|
382 have "Suc \<circ> (\<lambda>x. xs ! x) = (\<lambda>x. Suc (xs ! x))" by (simp add: comp_def) |
|
383 then have "map (\<lambda>x. Suc (xs ! x)) [0..<length xs] = map (Suc \<circ> (\<lambda>x. xs ! x)) [0..<length xs]" by simp |
|
384 also have "... = map Suc (map (\<lambda>x. xs ! x) [0..<length xs])" by simp |
|
385 also have "... = map Suc xs" by (simp add: map_nth) |
|
386 finally show "map (\<lambda>x. Suc (xs ! x)) [0..<length xs] = map Suc xs" . |
|
387 qed |
|
388 |
|
389 lemma strt_lemma [simp]: |
|
390 "length xs = vl \<Longrightarrow> rec_eval (rec_strt vl) xs = Strt xs" |
|
391 by (simp add: comp_def map_suc[symmetric]) |
|
392 |
|
393 lemma scan_lemma [simp]: |
|
394 "rec_eval rec_scan [r] = r mod 2" |
|
395 by(simp add: rec_scan_def) |
|
396 |
|
397 lemma newleft_lemma [simp]: |
|
398 "rec_eval rec_newleft [p, r, a] = Newleft p r a" |
|
399 by (simp add: rec_newleft_def Let_def) |
|
400 |
|
401 lemma newright_lemma [simp]: |
|
402 "rec_eval rec_newright [p, r, a] = Newright p r a" |
|
403 by (simp add: rec_newright_def Let_def) |
|
404 |
|
405 lemma actn_lemma [simp]: |
|
406 "rec_eval rec_actn [m, q, r] = Actn m q r" |
|
407 by (simp add: rec_actn_def) |
|
408 |
|
409 lemma newstat_lemma [simp]: |
|
410 "rec_eval rec_newstat [m, q, r] = Newstat m q r" |
|
411 by (simp add: rec_newstat_def) |
|
412 |
|
413 lemma trpl_lemma [simp]: |
|
414 "rec_eval rec_trpl [p, q, r] = Trpl p q r" |
|
415 by (simp add: rec_trpl_def) |
|
416 |
|
417 lemma left_lemma [simp]: |
|
418 "rec_eval rec_left [c] = Left c" |
|
419 by(simp add: rec_left_def) |
|
420 |
|
421 lemma right_lemma [simp]: |
|
422 "rec_eval rec_right [c] = Right c" |
|
423 by(simp add: rec_right_def) |
|
424 |
|
425 lemma stat_lemma [simp]: |
|
426 "rec_eval rec_stat [c] = Stat c" |
|
427 by(simp add: rec_stat_def) |
|
428 |
|
429 lemma newconf_lemma [simp]: |
|
430 "rec_eval rec_newconf [m, c] = Newconf m c" |
|
431 by (simp add: rec_newconf_def Let_def) |
|
432 |
|
433 lemma conf_lemma [simp]: |
|
434 "rec_eval rec_conf [k, m, r] = Conf k m r" |
|
435 by(induct k) (simp_all add: rec_conf_def) |
|
436 |
297 lemma nstd_lemma [simp]: |
437 lemma nstd_lemma [simp]: |
298 "rec_eval rec_nstd [c] = (if Nstd c then 1 else 0)" |
438 "rec_eval rec_nstd [c] = (if Nstd c then 1 else 0)" |
299 by(simp add: rec_nstd_def) |
439 by(simp add: rec_nstd_def) |
300 |
440 |
301 |
|
302 text{* |
|
303 @{text "Nostop t m r"} means that afer @{text "t"} steps of |
|
304 execution, the TM coded by @{text "m"} is not at a stardard |
|
305 final configuration. |
|
306 *} |
|
307 fun Nostop :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" |
|
308 where |
|
309 "Nostop t m r = Nstd (Conf t m r)" |
|
310 |
|
311 definition |
|
312 "rec_nostop = CN rec_nstd [rec_conf]" |
|
313 |
|
314 lemma nostop_lemma [simp]: |
441 lemma nostop_lemma [simp]: |
315 "rec_eval rec_nostop [t, m, r] = (if Nostop t m r then 1 else 0)" |
442 "rec_eval rec_nostop [t, m, r] = (if Nostop t m r then 1 else 0)" |
316 by (simp add: rec_nostop_def) |
443 by (simp add: rec_nostop_def) |
317 |
444 |
318 |
|
319 fun Value where |
|
320 "Value x = (Lg (Suc x) 2) - 1" |
|
321 |
|
322 definition |
|
323 "rec_value = CN rec_pred [CN rec_lg [S, constn 2]]" |
|
324 |
|
325 lemma value_lemma [simp]: |
445 lemma value_lemma [simp]: |
326 "rec_eval rec_value [x] = Value x" |
446 "rec_eval rec_value [x] = Value x" |
327 by (simp add: rec_value_def) |
447 by (simp add: rec_value_def) |
328 |
448 |
329 text{* |
449 lemma halt_lemma [simp]: |
330 @{text "rec_halt"} is the recursive function calculating the steps a TM needs to execute before |
450 "rec_eval rec_halt [m, r] = Halt m r" |
331 to reach a stardard final configuration. This recursive function is the only one |
451 by (simp add: rec_halt_def) |
332 using @{text "Mn"} combinator. So it is the only non-primitive recursive function |
452 |
333 needs to be used in the construction of the universal function @{text "rec_uf"}. |
453 lemma uf_lemma [simp]: |
334 *} |
454 "rec_eval rec_uf [m, r] = UF m r" |
335 |
455 by (simp add: rec_uf_def) |
336 definition |
456 |
337 "rec_halt = MN rec_nostop" |
|
338 |
|
339 |
|
340 definition |
|
341 "rec_uf = CN rec_value [CN rec_right [CN rec_conf [rec_halt, Id 2 0, Id 2 1]]]" |
|
342 |
|
343 text {* |
|
344 The correctness of @{text "rec_uf"}, nonhalt case. |
|
345 *} |
|
346 |
|
347 subsection {* Coding function of TMs *} |
|
348 |
|
349 text {* |
|
350 The purpose of this section is to get the coding function of Turing Machine, |
|
351 which is going to be named @{text "code"}. |
|
352 *} |
|
353 |
|
354 fun bl2nat :: "cell list \<Rightarrow> nat \<Rightarrow> nat" |
|
355 where |
|
356 "bl2nat [] n = 0" |
|
357 | "bl2nat (Bk # bl) n = bl2nat bl (Suc n)" |
|
358 | "bl2nat (Oc # bl) n = 2 ^ n + bl2nat bl (Suc n)" |
|
359 |
|
360 fun bl2wc :: "cell list \<Rightarrow> nat" |
|
361 where |
|
362 "bl2wc xs = bl2nat xs 0" |
|
363 |
|
364 fun trpl_code :: "config \<Rightarrow> nat" |
|
365 where |
|
366 "trpl_code (st, l, r) = Trpl (bl2wc l) st (bl2wc r)" |
|
367 |
|
368 fun action_map :: "action \<Rightarrow> nat" |
|
369 where |
|
370 "action_map W0 = 0" |
|
371 | "action_map W1 = 1" |
|
372 | "action_map L = 2" |
|
373 | "action_map R = 3" |
|
374 | "action_map Nop = 4" |
|
375 |
|
376 fun action_map_iff :: "nat \<Rightarrow> action" |
|
377 where |
|
378 "action_map_iff (0::nat) = W0" |
|
379 | "action_map_iff (Suc 0) = W1" |
|
380 | "action_map_iff (Suc (Suc 0)) = L" |
|
381 | "action_map_iff (Suc (Suc (Suc 0))) = R" |
|
382 | "action_map_iff n = Nop" |
|
383 |
|
384 fun block_map :: "cell \<Rightarrow> nat" |
|
385 where |
|
386 "block_map Bk = 0" |
|
387 | "block_map Oc = 1" |
|
388 |
|
389 fun Goedel_code' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
|
390 where |
|
391 "Goedel_code' [] n = 1" |
|
392 | "Goedel_code' (x # xs) n = (Pi n) ^ x * Goedel_code' xs (Suc n) " |
|
393 |
|
394 fun Goedel_code :: "nat list \<Rightarrow> nat" |
|
395 where |
|
396 "Goedel_code xs = 2 ^ (length xs) * (Goedel_code' xs 1)" |
|
397 |
|
398 fun modify_tprog :: "instr list \<Rightarrow> nat list" |
|
399 where |
|
400 "modify_tprog [] = []" |
|
401 | "modify_tprog ((a, s) # nl) = action_map a # s # modify_tprog nl" |
|
402 |
|
403 text {* @{text "Code tp"} gives the Goedel coding of TM program @{text "tp"}. *} |
|
404 fun Code :: "instr list \<Rightarrow> nat" |
|
405 where |
|
406 "Code tp = Goedel_code (modify_tprog tp)" |
|
407 |
457 |
408 subsection {* Relating interperter functions to the execution of TMs *} |
458 subsection {* Relating interperter functions to the execution of TMs *} |
409 |
459 |
|
460 lemma rec_step: |
|
461 assumes "(\<lambda> (s, l, r). s \<le> length tp div 2) c" |
|
462 shows "Trpl_code (step0 c tp) = Newconf (Code tp) (Trpl_code c)" |
|
463 apply(cases c) |
|
464 apply(simp only: Trpl_code.simps) |
|
465 apply(simp only: Let_def step.simps) |
|
466 apply(case_tac "fetch tp (a - 0) (read ca)") |
|
467 apply(simp only: prod.cases) |
|
468 apply(case_tac "update aa (b, ca)") |
|
469 apply(simp only: prod.cases) |
|
470 apply(simp only: Trpl_code.simps) |
|
471 apply(simp only: Newconf.simps) |
|
472 apply(simp only: Left.simps) |
|
473 oops |
|
474 |
|
475 lemma rec_steps: |
|
476 assumes "tm_wf0 tp" |
|
477 shows "Trpl_code (steps0 (1, Bk \<up> l, <lm>) tp stp) = Conf stp (Code tp) (bl2wc (<lm>))" |
|
478 apply(induct stp) |
|
479 apply(simp) |
|
480 apply(simp) |
|
481 oops |
|
482 |
410 |
483 |
411 lemma F_correct: |
484 lemma F_correct: |
412 assumes tp: "steps0 (1, Bk \<up> l, <lm>) tp stp = (0, Bk \<up> m, Oc \<up> rs @ Bk \<up> n)" |
485 assumes tm: "steps0 (1, Bk \<up> l, <lm>) tp stp = (0, Bk \<up> m, Oc \<up> rs @ Bk \<up> n)" |
413 and wf: "tm_wf0 tp" "0 < rs" |
486 and wf: "tm_wf0 tp" "0 < rs" |
414 shows "rec_eval rec_uf [Code tp, bl2wc (<lm>)] = (rs - Suc 0)" |
487 shows "rec_eval rec_uf [Code tp, bl2wc (<lm>)] = (rs - Suc 0)" |
|
488 proof - |
|
489 from least_steps[OF tm] |
|
490 obtain stp_least where |
|
491 before: "\<forall>stp' < stp_least. \<not> is_final (steps0 (1, Bk \<up> l, <lm>) tp stp')" and |
|
492 after: "\<forall>stp' \<ge> stp_least. is_final (steps0 (1, Bk \<up> l, <lm>) tp stp')" by blast |
|
493 have "Halt (Code tp) (bl2wc (<lm>)) = stp_least" sorry |
|
494 show ?thesis |
|
495 apply(simp only: uf_lemma) |
|
496 apply(simp only: UF.simps) |
|
497 apply(simp only: Halt.simps) |
|
498 oops |
415 |
499 |
416 |
500 |
417 end |
501 end |
418 |
502 |