author | zhang |
Mon, 15 Oct 2012 13:23:52 +0000 | |
changeset 371 | 48b231495281 |
parent 370 | 1ce04eb1c8ad |
permissions | -rw-r--r-- |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1 |
theory UF |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2 |
imports Main rec_def turing_basic GCD abacus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3 |
begin |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
5 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
6 |
This theory file constructs the Universal Function @{text "rec_F"}, which is the UTM defined |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
7 |
in terms of recursive functions. This @{text "rec_F"} is essentially an |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
8 |
interpreter of Turing Machines. Once the correctness of @{text "rec_F"} is established, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
9 |
UTM can easil be obtained by compling @{text "rec_F"} into the corresponding Turing Machine. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
10 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
11 |
|
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
12 |
section {* Univeral Function *} |
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
13 |
|
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
14 |
subsection {* The construction of component functions *} |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
15 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
16 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
17 |
This section constructs a set of component functions used to construct @{text "rec_F"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
18 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
19 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
20 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
21 |
The recursive function used to do arithmatic addition. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
22 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
23 |
definition rec_add :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
24 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
25 |
"rec_add \<equiv> Pr 1 (id 1 0) (Cn 3 s [id 3 2])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
26 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
27 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
28 |
The recursive function used to do arithmatic multiplication. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
29 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
30 |
definition rec_mult :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
31 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
32 |
"rec_mult = Pr 1 z (Cn 3 rec_add [id 3 0, id 3 2])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
33 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
34 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
35 |
The recursive function used to do arithmatic precede. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
36 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
37 |
definition rec_pred :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
38 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
39 |
"rec_pred = Cn 1 (Pr 1 z (id 3 1)) [id 1 0, id 1 0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
40 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
41 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
42 |
The recursive function used to do arithmatic subtraction. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
43 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
44 |
definition rec_minus :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
45 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
46 |
"rec_minus = Pr 1 (id 1 0) (Cn 3 rec_pred [id 3 2])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
47 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
48 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
49 |
@{text "constn n"} is the recursive function which computes |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
50 |
nature number @{text "n"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
51 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
52 |
fun constn :: "nat \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
53 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
54 |
"constn 0 = z" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
55 |
"constn (Suc n) = Cn 1 s [constn n]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
56 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
57 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
58 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
59 |
Signal function, which returns 1 when the input argument is greater than @{text "0"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
60 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
61 |
definition rec_sg :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
62 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
63 |
"rec_sg = Cn 1 rec_minus [constn 1, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
64 |
Cn 1 rec_minus [constn 1, id 1 0]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
65 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
66 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
67 |
@{text "rec_less"} compares its two arguments, returns @{text "1"} if |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
68 |
the first is less than the second; otherwise returns @{text "0"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
69 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
70 |
definition rec_less :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
71 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
72 |
"rec_less = Cn 2 rec_sg [Cn 2 rec_minus [id 2 1, id 2 0]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
73 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
74 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
75 |
@{text "rec_not"} inverse its argument: returns @{text "1"} when the |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
76 |
argument is @{text "0"}; returns @{text "0"} otherwise. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
77 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
78 |
definition rec_not :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
79 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
80 |
"rec_not = Cn 1 rec_minus [constn 1, id 1 0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
81 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
82 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
83 |
@{text "rec_eq"} compares its two arguments: returns @{text "1"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
84 |
if they are equal; return @{text "0"} otherwise. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
85 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
86 |
definition rec_eq :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
87 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
88 |
"rec_eq = Cn 2 rec_minus [Cn 2 (constn 1) [id 2 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
89 |
Cn 2 rec_add [Cn 2 rec_minus [id 2 0, id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
90 |
Cn 2 rec_minus [id 2 1, id 2 0]]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
91 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
92 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
93 |
@{text "rec_conj"} computes the conjunction of its two arguments, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
94 |
returns @{text "1"} if both of them are non-zero; returns @{text "0"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
95 |
otherwise. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
96 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
97 |
definition rec_conj :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
98 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
99 |
"rec_conj = Cn 2 rec_sg [Cn 2 rec_mult [id 2 0, id 2 1]] " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
100 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
101 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
102 |
@{text "rec_disj"} computes the disjunction of its two arguments, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
103 |
returns @{text "0"} if both of them are zero; returns @{text "0"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
104 |
otherwise. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
105 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
106 |
definition rec_disj :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
107 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
108 |
"rec_disj = Cn 2 rec_sg [Cn 2 rec_add [id 2 0, id 2 1]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
109 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
110 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
111 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
112 |
Computes the arity of recursive function. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
113 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
114 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
115 |
fun arity :: "recf \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
116 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
117 |
"arity z = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
118 |
| "arity s = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
119 |
| "arity (id m n) = m" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
120 |
| "arity (Cn n f gs) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
121 |
| "arity (Pr n f g) = Suc n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
122 |
| "arity (Mn n f) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
123 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
124 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
125 |
@{text "get_fstn_args n (Suc k)"} returns |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
126 |
@{text "[id n 0, id n 1, id n 2, \<dots>, id n k]"}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
127 |
the effect of which is to take out the first @{text "Suc k"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
128 |
arguments out of the @{text "n"} input arguments. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
129 |
*} |
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
130 |
|
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
131 |
fun get_fstn_args :: "nat \<Rightarrow> nat \<Rightarrow> recf list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
132 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
133 |
"get_fstn_args n 0 = []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
134 |
| "get_fstn_args n (Suc y) = get_fstn_args n y @ [id n y]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
135 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
136 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
137 |
@{text "rec_sigma f"} returns the recursive functions which |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
138 |
sums up the results of @{text "f"}: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
139 |
\[ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
140 |
(rec\_sigma f)(x, y) = f(x, 0) + f(x, 1) + \cdots + f(x, y) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
141 |
\] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
142 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
143 |
fun rec_sigma :: "recf \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
144 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
145 |
"rec_sigma rf = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
146 |
(let vl = arity rf in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
147 |
Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
148 |
[Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
149 |
(Cn (Suc vl) rec_add [id (Suc vl) vl, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
150 |
Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
151 |
@ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
152 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
153 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
154 |
@{text "rec_exec"} is the interpreter function for |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
155 |
reursive functions. The function is defined such that |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
156 |
it always returns meaningful results for primitive recursive |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
157 |
functions. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
158 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
159 |
function rec_exec :: "recf \<Rightarrow> nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
160 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
161 |
"rec_exec z xs = 0" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
162 |
"rec_exec s xs = (Suc (xs ! 0))" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
163 |
"rec_exec (id m n) xs = (xs ! n)" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
164 |
"rec_exec (Cn n f gs) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
165 |
(let ys = (map (\<lambda> a. rec_exec a xs) gs) in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
166 |
rec_exec f ys)" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
167 |
"rec_exec (Pr n f g) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
168 |
(if last xs = 0 then |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
169 |
rec_exec f (butlast xs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
170 |
else rec_exec g (butlast xs @ [last xs - 1] @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
171 |
[rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]))" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
172 |
"rec_exec (Mn n f) xs = (LEAST x. rec_exec f (xs @ [x]) = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
173 |
by pat_completeness auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
174 |
termination |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
175 |
proof |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
176 |
show "wf (measures [\<lambda> (r, xs). size r, (\<lambda> (r, xs). last xs)])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
177 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
178 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
179 |
fix n f gs xs x |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
180 |
assume "(x::recf) \<in> set gs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
181 |
thus "((x, xs), Cn n f gs, xs) \<in> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
182 |
measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
183 |
by(induct gs, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
184 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
185 |
fix n f gs xs x |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
186 |
assume "x = map (\<lambda>a. rec_exec a xs) gs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
187 |
"\<And>x. x \<in> set gs \<Longrightarrow> rec_exec_dom (x, xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
188 |
thus "((f, x), Cn n f gs, xs) \<in> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
189 |
measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
190 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
191 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
192 |
fix n f g xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
193 |
show "((f, butlast xs), Pr n f g, xs) \<in> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
194 |
measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
195 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
196 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
197 |
fix n f g xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
198 |
assume "last xs \<noteq> (0::nat)" thus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
199 |
"((Pr n f g, butlast xs @ [last xs - 1]), Pr n f g, xs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
200 |
\<in> measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
201 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
202 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
203 |
fix n f g xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
204 |
show "((g, butlast xs @ [last xs - 1] @ [rec_exec (Pr n f g) (butlast xs @ [last xs - 1])]), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
205 |
Pr n f g, xs) \<in> measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
206 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
207 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
208 |
fix n f xs x |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
209 |
show "((f, xs @ [x]), Mn n f, xs) \<in> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
210 |
measures [\<lambda>(r, xs). size r, \<lambda>(r, xs). last xs]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
211 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
212 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
213 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
214 |
declare rec_exec.simps[simp del] constn.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
215 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
216 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
217 |
Correctness of @{text "rec_add"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
218 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
219 |
lemma add_lemma: "\<And> x y. rec_exec rec_add [x, y] = x + y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
220 |
by(induct_tac y, auto simp: rec_add_def rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
221 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
222 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
223 |
Correctness of @{text "rec_mult"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
224 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
225 |
lemma mult_lemma: "\<And> x y. rec_exec rec_mult [x, y] = x * y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
226 |
by(induct_tac y, auto simp: rec_mult_def rec_exec.simps add_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
227 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
228 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
229 |
Correctness of @{text "rec_pred"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
230 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
231 |
lemma pred_lemma: "\<And> x. rec_exec rec_pred [x] = x - 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
232 |
by(induct_tac x, auto simp: rec_pred_def rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
233 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
234 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
235 |
Correctness of @{text "rec_minus"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
236 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
237 |
lemma minus_lemma: "\<And> x y. rec_exec rec_minus [x, y] = x - y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
238 |
by(induct_tac y, auto simp: rec_exec.simps rec_minus_def pred_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
239 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
240 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
241 |
Correctness of @{text "rec_sg"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
242 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
243 |
lemma sg_lemma: "\<And> x. rec_exec rec_sg [x] = (if x = 0 then 0 else 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
244 |
by(auto simp: rec_sg_def minus_lemma rec_exec.simps constn.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
245 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
246 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
247 |
Correctness of @{text "constn"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
248 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
249 |
lemma constn_lemma: "rec_exec (constn n) [x] = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
250 |
by(induct n, auto simp: rec_exec.simps constn.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
251 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
252 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
253 |
Correctness of @{text "rec_less"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
254 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
255 |
lemma less_lemma: "\<And> x y. rec_exec rec_less [x, y] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
256 |
(if x < y then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
257 |
by(induct_tac y, auto simp: rec_exec.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
258 |
rec_less_def minus_lemma sg_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
259 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
260 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
261 |
Correctness of @{text "rec_not"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
262 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
263 |
lemma not_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
264 |
"\<And> x. rec_exec rec_not [x] = (if x = 0 then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
265 |
by(induct_tac x, auto simp: rec_exec.simps rec_not_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
266 |
constn_lemma minus_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
267 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
268 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
269 |
Correctness of @{text "rec_eq"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
270 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
271 |
lemma eq_lemma: "\<And> x y. rec_exec rec_eq [x, y] = (if x = y then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
272 |
by(induct_tac y, auto simp: rec_exec.simps rec_eq_def constn_lemma add_lemma minus_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
273 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
274 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
275 |
Correctness of @{text "rec_conj"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
276 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
277 |
lemma conj_lemma: "\<And> x y. rec_exec rec_conj [x, y] = (if x = 0 \<or> y = 0 then 0 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
278 |
else 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
279 |
by(induct_tac y, auto simp: rec_exec.simps sg_lemma rec_conj_def mult_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
280 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
281 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
282 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
283 |
Correctness of @{text "rec_disj"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
284 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
285 |
lemma disj_lemma: "\<And> x y. rec_exec rec_disj [x, y] = (if x = 0 \<and> y = 0 then 0 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
286 |
else 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
287 |
by(induct_tac y, auto simp: rec_disj_def sg_lemma add_lemma rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
288 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
289 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
290 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
291 |
@{text "primrec recf n"} is true iff |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
292 |
@{text "recf"} is a primitive recursive function |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
293 |
with arity @{text "n"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
294 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
295 |
inductive primerec :: "recf \<Rightarrow> nat \<Rightarrow> bool" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
296 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
297 |
prime_z[intro]: "primerec z (Suc 0)" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
298 |
prime_s[intro]: "primerec s (Suc 0)" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
299 |
prime_id[intro!]: "\<lbrakk>n < m\<rbrakk> \<Longrightarrow> primerec (id m n) m" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
300 |
prime_cn[intro!]: "\<lbrakk>primerec f k; length gs = k; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
301 |
\<forall> i < length gs. primerec (gs ! i) m; m = n\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
302 |
\<Longrightarrow> primerec (Cn n f gs) m" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
303 |
prime_pr[intro!]: "\<lbrakk>primerec f n; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
304 |
primerec g (Suc (Suc n)); m = Suc n\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
305 |
\<Longrightarrow> primerec (Pr n f g) m" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
306 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
307 |
inductive_cases prime_cn_reverse'[elim]: "primerec (Cn n f gs) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
308 |
inductive_cases prime_mn_reverse: "primerec (Mn n f) m" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
309 |
inductive_cases prime_z_reverse[elim]: "primerec z n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
310 |
inductive_cases prime_s_reverse[elim]: "primerec s n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
311 |
inductive_cases prime_id_reverse[elim]: "primerec (id m n) k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
312 |
inductive_cases prime_cn_reverse[elim]: "primerec (Cn n f gs) m" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
313 |
inductive_cases prime_pr_reverse[elim]: "primerec (Pr n f g) m" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
314 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
315 |
declare mult_lemma[simp] add_lemma[simp] pred_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
316 |
minus_lemma[simp] sg_lemma[simp] constn_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
317 |
less_lemma[simp] not_lemma[simp] eq_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
318 |
conj_lemma[simp] disj_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
319 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
320 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
321 |
@{text "Sigma"} is the logical specification of |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
322 |
the recursive function @{text "rec_sigma"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
323 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
324 |
function Sigma :: "(nat list \<Rightarrow> nat) \<Rightarrow> nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
325 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
326 |
"Sigma g xs = (if last xs = 0 then g xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
327 |
else (Sigma g (butlast xs @ [last xs - 1]) + |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
328 |
g xs)) " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
329 |
by pat_completeness auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
330 |
termination |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
331 |
proof |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
332 |
show "wf (measure (\<lambda> (f, xs). last xs))" by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
333 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
334 |
fix g xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
335 |
assume "last (xs::nat list) \<noteq> 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
336 |
thus "((g, butlast xs @ [last xs - 1]), g, xs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
337 |
\<in> measure (\<lambda>(f, xs). last xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
338 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
339 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
340 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
341 |
declare rec_exec.simps[simp del] get_fstn_args.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
342 |
arity.simps[simp del] Sigma.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
343 |
rec_sigma.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
344 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
345 |
lemma [simp]: "arity z = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
346 |
by(simp add: arity.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
347 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
348 |
lemma rec_pr_0_simp_rewrite: " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
349 |
rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
350 |
by(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
351 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
352 |
lemma rec_pr_0_simp_rewrite_single_param: " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
353 |
rec_exec (Pr n f g) [0] = rec_exec f []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
354 |
by(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
355 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
356 |
lemma rec_pr_Suc_simp_rewrite: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
357 |
"rec_exec (Pr n f g) (xs @ [Suc x]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
358 |
rec_exec g (xs @ [x] @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
359 |
[rec_exec (Pr n f g) (xs @ [x])])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
360 |
by(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
361 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
362 |
lemma rec_pr_Suc_simp_rewrite_single_param: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
363 |
"rec_exec (Pr n f g) ([Suc x]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
364 |
rec_exec g ([x] @ [rec_exec (Pr n f g) ([x])])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
365 |
by(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
366 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
367 |
lemma Sigma_0_simp_rewrite_single_param: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
368 |
"Sigma f [0] = f [0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
369 |
by(simp add: Sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
370 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
371 |
lemma Sigma_0_simp_rewrite: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
372 |
"Sigma f (xs @ [0]) = f (xs @ [0])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
373 |
by(simp add: Sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
374 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
375 |
lemma Sigma_Suc_simp_rewrite: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
376 |
"Sigma f (xs @ [Suc x]) = Sigma f (xs @ [x]) + f (xs @ [Suc x])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
377 |
by(simp add: Sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
378 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
379 |
lemma Sigma_Suc_simp_rewrite_single: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
380 |
"Sigma f ([Suc x]) = Sigma f ([x]) + f ([Suc x])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
381 |
by(simp add: Sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
382 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
383 |
lemma [simp]: "(xs @ ys) ! (Suc (length xs)) = ys ! 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
384 |
by(simp add: nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
385 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
386 |
lemma get_fstn_args_take: "\<lbrakk>length xs = m; n \<le> m\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
387 |
map (\<lambda> f. rec_exec f xs) (get_fstn_args m n)= take n xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
388 |
proof(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
389 |
case 0 thus "?case" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
390 |
by(simp add: get_fstn_args.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
391 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
392 |
case (Suc n) thus "?case" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
393 |
by(simp add: get_fstn_args.simps rec_exec.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
394 |
take_Suc_conv_app_nth) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
395 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
396 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
397 |
lemma [simp]: "primerec f n \<Longrightarrow> arity f = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
398 |
apply(case_tac f) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
399 |
apply(auto simp: arity.simps ) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
400 |
apply(erule_tac prime_mn_reverse) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
401 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
402 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
403 |
lemma rec_sigma_Suc_simp_rewrite: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
404 |
"primerec f (Suc (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
405 |
\<Longrightarrow> rec_exec (rec_sigma f) (xs @ [Suc x]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
406 |
rec_exec (rec_sigma f) (xs @ [x]) + rec_exec f (xs @ [Suc x])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
407 |
apply(induct x) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
408 |
apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
409 |
rec_exec.simps get_fstn_args_take) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
410 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
411 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
412 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
413 |
The correctness of @{text "rec_sigma"} with respect to its specification. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
414 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
415 |
lemma sigma_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
416 |
"primerec rg (Suc (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
417 |
\<Longrightarrow> rec_exec (rec_sigma rg) (xs @ [x]) = Sigma (rec_exec rg) (xs @ [x])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
418 |
apply(induct x) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
419 |
apply(auto simp: rec_exec.simps rec_sigma.simps Let_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
420 |
get_fstn_args_take Sigma_0_simp_rewrite |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
421 |
Sigma_Suc_simp_rewrite) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
422 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
423 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
424 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
425 |
@{text "rec_accum f (x1, x2, \<dots>, xn, k) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
426 |
f(x1, x2, \<dots>, xn, 0) * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
427 |
f(x1, x2, \<dots>, xn, 1) * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
428 |
\<dots> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
429 |
f(x1, x2, \<dots>, xn, k)"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
430 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
431 |
fun rec_accum :: "recf \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
432 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
433 |
"rec_accum rf = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
434 |
(let vl = arity rf in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
435 |
Pr (vl - 1) (Cn (vl - 1) rf (get_fstn_args (vl - 1) (vl - 1) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
436 |
[Cn (vl - 1) (constn 0) [id (vl - 1) 0]])) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
437 |
(Cn (Suc vl) rec_mult [id (Suc vl) (vl), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
438 |
Cn (Suc vl) rf (get_fstn_args (Suc vl) (vl - 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
439 |
@ [Cn (Suc vl) s [id (Suc vl) (vl - 1)]])]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
440 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
441 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
442 |
@{text "Accum"} is the formal specification of @{text "rec_accum"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
443 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
444 |
function Accum :: "(nat list \<Rightarrow> nat) \<Rightarrow> nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
445 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
446 |
"Accum f xs = (if last xs = 0 then f xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
447 |
else (Accum f (butlast xs @ [last xs - 1]) * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
448 |
f xs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
449 |
by pat_completeness auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
450 |
termination |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
451 |
proof |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
452 |
show "wf (measure (\<lambda> (f, xs). last xs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
453 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
454 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
455 |
fix f xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
456 |
assume "last xs \<noteq> (0::nat)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
457 |
thus "((f, butlast xs @ [last xs - 1]), f, xs) \<in> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
458 |
measure (\<lambda>(f, xs). last xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
459 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
460 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
461 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
462 |
lemma rec_accum_Suc_simp_rewrite: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
463 |
"primerec f (Suc (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
464 |
\<Longrightarrow> rec_exec (rec_accum f) (xs @ [Suc x]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
465 |
rec_exec (rec_accum f) (xs @ [x]) * rec_exec f (xs @ [Suc x])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
466 |
apply(induct x) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
467 |
apply(auto simp: rec_sigma.simps Let_def rec_pr_Suc_simp_rewrite |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
468 |
rec_exec.simps get_fstn_args_take) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
469 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
470 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
471 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
472 |
The correctness of @{text "rec_accum"} with respect to its specification. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
473 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
474 |
lemma accum_lemma : |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
475 |
"primerec rg (Suc (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
476 |
\<Longrightarrow> rec_exec (rec_accum rg) (xs @ [x]) = Accum (rec_exec rg) (xs @ [x])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
477 |
apply(induct x) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
478 |
apply(auto simp: rec_exec.simps rec_sigma.simps Let_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
479 |
get_fstn_args_take) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
480 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
481 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
482 |
declare rec_accum.simps [simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
483 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
484 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
485 |
@{text "rec_all t f (x1, x2, \<dots>, xn)"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
486 |
computes the charactrization function of the following FOL formula: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
487 |
@{text "(\<forall> x \<le> t(x1, x2, \<dots>, xn). (f(x1, x2, \<dots>, xn, x) > 0))"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
488 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
489 |
fun rec_all :: "recf \<Rightarrow> recf \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
490 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
491 |
"rec_all rt rf = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
492 |
(let vl = arity rf in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
493 |
Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_accum rf) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
494 |
(get_fstn_args (vl - 1) (vl - 1) @ [rt])])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
495 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
496 |
lemma rec_accum_ex: "primerec rf (Suc (length xs)) \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
497 |
(rec_exec (rec_accum rf) (xs @ [x]) = 0) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
498 |
(\<exists> t \<le> x. rec_exec rf (xs @ [t]) = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
499 |
apply(induct x, simp_all add: rec_accum_Suc_simp_rewrite) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
500 |
apply(simp add: rec_exec.simps rec_accum.simps get_fstn_args_take, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
501 |
auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
502 |
apply(rule_tac x = ta in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
503 |
apply(case_tac "t = Suc x", simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
504 |
apply(rule_tac x = t in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
505 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
506 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
507 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
508 |
The correctness of @{text "rec_all"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
509 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
510 |
lemma all_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
511 |
"\<lbrakk>primerec rf (Suc (length xs)); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
512 |
primerec rt (length xs)\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
514 |
else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
515 |
apply(auto simp: rec_all.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
516 |
apply(simp add: rec_exec.simps map_append get_fstn_args_take split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
517 |
apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
518 |
apply(case_tac "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0", simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
519 |
apply(erule_tac exE, erule_tac x = t in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
520 |
apply(simp add: rec_exec.simps map_append get_fstn_args_take) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
521 |
apply(drule_tac x = "rec_exec rt xs" in rec_accum_ex) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
522 |
apply(case_tac "rec_exec (rec_accum rf) (xs @ [rec_exec rt xs]) = 0", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
523 |
apply(erule_tac x = x in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
524 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
525 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
526 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
527 |
@{text "rec_ex t f (x1, x2, \<dots>, xn)"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
528 |
computes the charactrization function of the following FOL formula: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
529 |
@{text "(\<exists> x \<le> t(x1, x2, \<dots>, xn). (f(x1, x2, \<dots>, xn, x) > 0))"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
530 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
531 |
fun rec_ex :: "recf \<Rightarrow> recf \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
532 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
533 |
"rec_ex rt rf = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
534 |
(let vl = arity rf in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
535 |
Cn (vl - 1) rec_sg [Cn (vl - 1) (rec_sigma rf) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
536 |
(get_fstn_args (vl - 1) (vl - 1) @ [rt])])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
537 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
538 |
lemma rec_sigma_ex: "primerec rf (Suc (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
539 |
\<Longrightarrow> (rec_exec (rec_sigma rf) (xs @ [x]) = 0) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
540 |
(\<forall> t \<le> x. rec_exec rf (xs @ [t]) = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
541 |
apply(induct x, simp_all add: rec_sigma_Suc_simp_rewrite) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
542 |
apply(simp add: rec_exec.simps rec_sigma.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
543 |
get_fstn_args_take, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
544 |
apply(case_tac "t = Suc x", simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
545 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
546 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
547 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
548 |
The correctness of @{text "ex_lemma"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
549 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
550 |
lemma ex_lemma:" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
551 |
\<lbrakk>primerec rf (Suc (length xs)); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
552 |
primerec rt (length xs)\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
553 |
\<Longrightarrow> (rec_exec (rec_ex rt rf) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
554 |
(if (\<exists> x \<le> (rec_exec rt xs). 0 <rec_exec rf (xs @ [x])) then 1 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
555 |
else 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
556 |
apply(auto simp: rec_ex.simps rec_exec.simps map_append get_fstn_args_take |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
557 |
split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
558 |
apply(drule_tac x = "rec_exec rt xs" in rec_sigma_ex, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
559 |
apply(drule_tac x = "rec_exec rt xs" in rec_sigma_ex, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
560 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
561 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
562 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
563 |
Defintiion of @{text "Min[R]"} on page 77 of Boolos's book. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
564 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
565 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
566 |
fun Minr :: "(nat list \<Rightarrow> bool) \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
567 |
where "Minr Rr xs w = (let setx = {y | y. (y \<le> w) \<and> Rr (xs @ [y])} in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
568 |
if (setx = {}) then (Suc w) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
569 |
else (Min setx))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
570 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
571 |
declare Minr.simps[simp del] rec_all.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
572 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
573 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
574 |
The following is a set of auxilliary lemmas about @{text "Minr"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
575 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
576 |
lemma Minr_range: "Minr Rr xs w \<le> w \<or> Minr Rr xs w = Suc w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
577 |
apply(auto simp: Minr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
578 |
apply(subgoal_tac "Min {x. x \<le> w \<and> Rr (xs @ [x])} \<le> x") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
579 |
apply(erule_tac order_trans, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
580 |
apply(rule_tac Min_le, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
581 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
582 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
583 |
lemma [simp]: "{x. x \<le> Suc w \<and> Rr (xs @ [x])} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
584 |
= (if Rr (xs @ [Suc w]) then insert (Suc w) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
585 |
{x. x \<le> w \<and> Rr (xs @ [x])} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
586 |
else {x. x \<le> w \<and> Rr (xs @ [x])})" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
587 |
by(auto, case_tac "x = Suc w", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
588 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
589 |
lemma [simp]: "Minr Rr xs w \<le> w \<Longrightarrow> Minr Rr xs (Suc w) = Minr Rr xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
590 |
apply(simp add: Minr.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
591 |
apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
592 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
593 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
594 |
lemma [simp]: "\<forall>x\<le>w. \<not> Rr (xs @ [x]) \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
595 |
{x. x \<le> w \<and> Rr (xs @ [x])} = {} " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
596 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
597 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
598 |
lemma [simp]: "\<lbrakk>Minr Rr xs w = Suc w; Rr (xs @ [Suc w])\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
599 |
Minr Rr xs (Suc w) = Suc w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
600 |
apply(simp add: Minr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
601 |
apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
602 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
603 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
604 |
lemma [simp]: "\<lbrakk>Minr Rr xs w = Suc w; \<not> Rr (xs @ [Suc w])\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
605 |
Minr Rr xs (Suc w) = Suc (Suc w)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
606 |
apply(simp add: Minr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
607 |
apply(case_tac "\<forall>x\<le>w. \<not> Rr (xs @ [x])", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
608 |
apply(subgoal_tac "Min {x. x \<le> w \<and> Rr (xs @ [x])} \<in> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
609 |
{x. x \<le> w \<and> Rr (xs @ [x])}", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
610 |
apply(rule_tac Min_in, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
611 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
612 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
613 |
lemma Minr_Suc_simp: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
614 |
"Minr Rr xs (Suc w) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
615 |
(if Minr Rr xs w \<le> w then Minr Rr xs w |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
616 |
else if (Rr (xs @ [Suc w])) then (Suc w) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
617 |
else Suc (Suc w))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
618 |
by(insert Minr_range[of Rr xs w], auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
619 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
620 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
621 |
@{text "rec_Minr"} is the recursive function |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
622 |
used to implement @{text "Minr"}: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
623 |
if @{text "Rr"} is implemented by a recursive function @{text "recf"}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
624 |
then @{text "rec_Minr recf"} is the recursive function used to |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
625 |
implement @{text "Minr Rr"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
626 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
627 |
fun rec_Minr :: "recf \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
628 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
629 |
"rec_Minr rf = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
630 |
(let vl = arity rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
631 |
in let rq = rec_all (id vl (vl - 1)) (Cn (Suc vl) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
632 |
rec_not [Cn (Suc vl) rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
633 |
(get_fstn_args (Suc vl) (vl - 1) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
634 |
[id (Suc vl) (vl)])]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
635 |
in rec_sigma rq)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
636 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
637 |
lemma length_getpren_params[simp]: "length (get_fstn_args m n) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
638 |
by(induct n, auto simp: get_fstn_args.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
639 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
640 |
lemma length_app: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
641 |
"(length (get_fstn_args (arity rf - Suc 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
642 |
(arity rf - Suc 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
643 |
@ [Cn (arity rf - Suc 0) (constn 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
644 |
[recf.id (arity rf - Suc 0) 0]])) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
645 |
= (Suc (arity rf - Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
646 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
647 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
648 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
649 |
lemma primerec_accum: "primerec (rec_accum rf) n \<Longrightarrow> primerec rf n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
650 |
apply(auto simp: rec_accum.simps Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
651 |
apply(erule_tac prime_pr_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
652 |
apply(erule_tac prime_cn_reverse, simp only: length_app) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
653 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
654 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
655 |
lemma primerec_all: "primerec (rec_all rt rf) n \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
656 |
primerec rt n \<and> primerec rf (Suc n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
657 |
apply(simp add: rec_all.simps Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
658 |
apply(erule_tac prime_cn_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
659 |
apply(erule_tac prime_cn_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
660 |
apply(erule_tac x = n in allE, simp add: nth_append primerec_accum) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
661 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
662 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
663 |
lemma min_Suc_Suc[simp]: "min (Suc (Suc x)) x = x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
664 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
665 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
666 |
declare numeral_3_eq_3[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
667 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
668 |
lemma [intro]: "primerec rec_pred (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
669 |
apply(simp add: rec_pred_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
670 |
apply(rule_tac prime_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
671 |
apply(case_tac i, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
672 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
673 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
674 |
lemma [intro]: "primerec rec_minus (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
675 |
apply(auto simp: rec_minus_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
676 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
677 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
678 |
lemma [intro]: "primerec (constn n) (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
679 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
680 |
apply(auto simp: constn.simps intro: prime_z prime_cn prime_s) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
681 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
682 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
683 |
lemma [intro]: "primerec rec_sg (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
684 |
apply(simp add: rec_sg_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
685 |
apply(rule_tac k = "Suc (Suc 0)" in prime_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
686 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
687 |
apply(case_tac ia, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
688 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
689 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
690 |
lemma [simp]: "length (get_fstn_args m n) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
691 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
692 |
apply(auto simp: get_fstn_args.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
693 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
694 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
695 |
lemma primerec_getpren[elim]: "\<lbrakk>i < n; n \<le> m\<rbrakk> \<Longrightarrow> primerec (get_fstn_args m n ! i) m" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
696 |
apply(induct n, auto simp: get_fstn_args.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
697 |
apply(case_tac "i = n", auto simp: nth_append intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
698 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
699 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
700 |
lemma [intro]: "primerec rec_add (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
701 |
apply(simp add: rec_add_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
702 |
apply(rule_tac prime_pr, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
703 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
704 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
705 |
lemma [intro]:"primerec rec_mult (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
706 |
apply(simp add: rec_mult_def ) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
707 |
apply(rule_tac prime_pr, auto intro: prime_z) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
708 |
apply(case_tac i, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
709 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
710 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
711 |
lemma [elim]: "\<lbrakk>primerec rf n; n \<ge> Suc (Suc 0)\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
712 |
primerec (rec_accum rf) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
713 |
apply(auto simp: rec_accum.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
714 |
apply(simp add: nth_append, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
715 |
apply(case_tac i, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
716 |
apply(auto simp: nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
717 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
718 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
719 |
lemma primerec_all_iff: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
720 |
"\<lbrakk>primerec rt n; primerec rf (Suc n); n > 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
721 |
primerec (rec_all rt rf) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
722 |
apply(simp add: rec_all.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
723 |
apply(auto, simp add: nth_append, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
724 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
725 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
726 |
lemma [simp]: "Rr (xs @ [0]) \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
727 |
Min {x. x = (0::nat) \<and> Rr (xs @ [x])} = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
728 |
by(rule_tac Min_eqI, simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
729 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
730 |
lemma [intro]: "primerec rec_not (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
731 |
apply(simp add: rec_not_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
732 |
apply(rule prime_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
733 |
apply(case_tac i, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
734 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
735 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
736 |
lemma Min_false1[simp]: "\<lbrakk>\<not> Min {uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])} \<le> w; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
737 |
x \<le> w; 0 < rec_exec rf (xs @ [x])\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
738 |
\<Longrightarrow> False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
739 |
apply(subgoal_tac "finite {uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])}") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
740 |
apply(subgoal_tac "{uu. uu \<le> w \<and> 0 < rec_exec rf (xs @ [uu])} \<noteq> {}") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
741 |
apply(simp add: Min_le_iff, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
742 |
apply(rule_tac x = x in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
743 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
744 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
745 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
746 |
lemma sigma_minr_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
747 |
assumes prrf: "primerec rf (Suc (length xs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
748 |
shows "UF.Sigma (rec_exec (rec_all (recf.id (Suc (length xs)) (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
749 |
(Cn (Suc (Suc (length xs))) rec_not |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
750 |
[Cn (Suc (Suc (length xs))) rf (get_fstn_args (Suc (Suc (length xs))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
751 |
(length xs) @ [recf.id (Suc (Suc (length xs))) (Suc (length xs))])]))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
752 |
(xs @ [w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
753 |
Minr (\<lambda>args. 0 < rec_exec rf args) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
754 |
proof(induct w) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
755 |
let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
756 |
let ?rf = "(Cn (Suc (Suc (length xs))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
757 |
rec_not [Cn (Suc (Suc (length xs))) rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
758 |
(get_fstn_args (Suc (Suc (length xs))) (length xs) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
759 |
[recf.id (Suc (Suc (length xs))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
760 |
(Suc ((length xs)))])])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
761 |
let ?rq = "(rec_all ?rt ?rf)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
762 |
have prrf: "primerec ?rf (Suc (length (xs @ [0]))) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
763 |
primerec ?rt (length (xs @ [0]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
764 |
apply(auto simp: prrf nth_append)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
765 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
766 |
show "Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [0]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
767 |
= Minr (\<lambda>args. 0 < rec_exec rf args) xs 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
768 |
apply(simp add: Sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
769 |
apply(simp only: prrf all_lemma, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
770 |
auto simp: rec_exec.simps get_fstn_args_take Minr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
771 |
apply(rule_tac Min_eqI, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
772 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
773 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
774 |
fix w |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
775 |
let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
776 |
let ?rf = "(Cn (Suc (Suc (length xs))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
777 |
rec_not [Cn (Suc (Suc (length xs))) rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
778 |
(get_fstn_args (Suc (Suc (length xs))) (length xs) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
779 |
[recf.id (Suc (Suc (length xs))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
780 |
(Suc ((length xs)))])])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
781 |
let ?rq = "(rec_all ?rt ?rf)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
782 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
783 |
"Sigma (rec_exec (rec_all ?rt ?rf)) (xs @ [w]) = Minr (\<lambda>args. 0 < rec_exec rf args) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
784 |
have prrf: "primerec ?rf (Suc (length (xs @ [Suc w]))) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
785 |
primerec ?rt (length (xs @ [Suc w]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
786 |
apply(auto simp: prrf nth_append)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
787 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
788 |
show "UF.Sigma (rec_exec (rec_all ?rt ?rf)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
789 |
(xs @ [Suc w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
790 |
Minr (\<lambda>args. 0 < rec_exec rf args) xs (Suc w)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
791 |
apply(auto simp: Sigma_Suc_simp_rewrite ind Minr_Suc_simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
792 |
apply(simp_all only: prrf all_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
793 |
apply(auto simp: rec_exec.simps get_fstn_args_take Let_def Minr.simps split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
794 |
apply(drule_tac Min_false1, simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
795 |
apply(case_tac "x = Suc w", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
796 |
apply(drule_tac Min_false1, simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
797 |
apply(drule_tac Min_false1, simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
798 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
799 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
800 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
801 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
802 |
The correctness of @{text "rec_Minr"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
803 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
804 |
lemma Minr_lemma: " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
805 |
\<lbrakk>primerec rf (Suc (length xs))\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
806 |
\<Longrightarrow> rec_exec (rec_Minr rf) (xs @ [w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
807 |
Minr (\<lambda> args. (0 < rec_exec rf args)) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
808 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
809 |
let ?rt = "(recf.id (Suc (length xs)) ((length xs)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
810 |
let ?rf = "(Cn (Suc (Suc (length xs))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
811 |
rec_not [Cn (Suc (Suc (length xs))) rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
812 |
(get_fstn_args (Suc (Suc (length xs))) (length xs) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
813 |
[recf.id (Suc (Suc (length xs))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
814 |
(Suc ((length xs)))])])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
815 |
let ?rq = "(rec_all ?rt ?rf)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
816 |
assume h: "primerec rf (Suc (length xs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
817 |
have h1: "primerec ?rq (Suc (length xs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
818 |
apply(rule_tac primerec_all_iff) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
819 |
apply(auto simp: h nth_append)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
820 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
821 |
moreover have "arity rf = Suc (length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
822 |
using h by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
823 |
ultimately show "rec_exec (rec_Minr rf) (xs @ [w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
824 |
Minr (\<lambda> args. (0 < rec_exec rf args)) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
825 |
apply(simp add: rec_exec.simps rec_Minr.simps arity.simps Let_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
826 |
sigma_lemma all_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
827 |
apply(rule_tac sigma_minr_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
828 |
apply(simp add: h) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
829 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
830 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
831 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
832 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
833 |
@{text "rec_le"} is the comparasion function |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
834 |
which compares its two arguments, testing whether the |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
835 |
first is less or equal to the second. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
836 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
837 |
definition rec_le :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
838 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
839 |
"rec_le = Cn (Suc (Suc 0)) rec_disj [rec_less, rec_eq]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
840 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
841 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
842 |
The correctness of @{text "rec_le"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
843 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
844 |
lemma le_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
845 |
"\<And>x y. rec_exec rec_le [x, y] = (if (x \<le> y) then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
846 |
by(auto simp: rec_le_def rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
847 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
848 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
849 |
Defintiion of @{text "Max[Rr]"} on page 77 of Boolos's book. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
850 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
851 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
852 |
fun Maxr :: "(nat list \<Rightarrow> bool) \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
853 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
854 |
"Maxr Rr xs w = (let setx = {y. y \<le> w \<and> Rr (xs @[y])} in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
855 |
if setx = {} then 0 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
856 |
else Max setx)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
857 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
858 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
859 |
@{text "rec_maxr"} is the recursive function |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
860 |
used to implementation @{text "Maxr"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
861 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
862 |
fun rec_maxr :: "recf \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
863 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
864 |
"rec_maxr rr = (let vl = arity rr in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
865 |
let rt = id (Suc vl) (vl - 1) in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
866 |
let rf1 = Cn (Suc (Suc vl)) rec_le |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
867 |
[id (Suc (Suc vl)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
868 |
((Suc vl)), id (Suc (Suc vl)) (vl)] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
869 |
let rf2 = Cn (Suc (Suc vl)) rec_not |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
870 |
[Cn (Suc (Suc vl)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
871 |
rr (get_fstn_args (Suc (Suc vl)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
872 |
(vl - 1) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
873 |
[id (Suc (Suc vl)) ((Suc vl))])] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
874 |
let rf = Cn (Suc (Suc vl)) rec_disj [rf1, rf2] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
875 |
let rq = rec_all rt rf in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
876 |
let Qf = Cn (Suc vl) rec_not [rec_all rt rf] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
877 |
in Cn vl (rec_sigma Qf) (get_fstn_args vl vl @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
878 |
[id vl (vl - 1)]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
879 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
880 |
declare rec_maxr.simps[simp del] Maxr.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
881 |
declare le_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
882 |
lemma [simp]: "(min (Suc (Suc (Suc (x)))) (x)) = x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
883 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
884 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
885 |
declare numeral_2_eq_2[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
886 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
887 |
lemma [intro]: "primerec rec_disj (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
888 |
apply(simp add: rec_disj_def, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
889 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
890 |
apply(case_tac ia, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
891 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
892 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
893 |
lemma [intro]: "primerec rec_less (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
894 |
apply(simp add: rec_less_def, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
895 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
896 |
apply(case_tac ia , auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
897 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
898 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
899 |
lemma [intro]: "primerec rec_eq (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
900 |
apply(simp add: rec_eq_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
901 |
apply(rule_tac prime_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
902 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
903 |
apply(case_tac ia, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
904 |
apply(case_tac [!] i, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
905 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
906 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
907 |
lemma [intro]: "primerec rec_le (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
908 |
apply(simp add: rec_le_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
909 |
apply(rule_tac prime_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
910 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
911 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
912 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
913 |
lemma [simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
914 |
"length ys = Suc n \<Longrightarrow> (take n ys @ [ys ! n, ys ! n]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
915 |
ys @ [ys ! n]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
916 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
917 |
apply(subgoal_tac "\<exists> xs y. ys = xs @ [y]", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
918 |
apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
919 |
apply(case_tac "ys = []", simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
920 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
921 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
922 |
lemma Maxr_Suc_simp: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
923 |
"Maxr Rr xs (Suc w) =(if Rr (xs @ [Suc w]) then Suc w |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
924 |
else Maxr Rr xs w)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
925 |
apply(auto simp: Maxr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
926 |
apply(rule_tac max_absorb1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
927 |
apply(subgoal_tac "(Max {y. y \<le> w \<and> Rr (xs @ [y])} \<le> (Suc w)) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
928 |
(\<forall>a\<in>{y. y \<le> w \<and> Rr (xs @ [y])}. a \<le> (Suc w))", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
929 |
apply(rule_tac Max_le_iff, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
930 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
931 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
932 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
933 |
lemma [simp]: "min (Suc n) n = n" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
934 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
935 |
lemma Sigma_0: "\<forall> i \<le> n. (f (xs @ [i]) = 0) \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
936 |
Sigma f (xs @ [n]) = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
937 |
apply(induct n, simp add: Sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
938 |
apply(simp add: Sigma_Suc_simp_rewrite) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
939 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
940 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
941 |
lemma [elim]: "\<forall>k<Suc w. f (xs @ [k]) = Suc 0 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
942 |
\<Longrightarrow> Sigma f (xs @ [w]) = Suc w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
943 |
apply(induct w) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
944 |
apply(simp add: Sigma.simps, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
945 |
apply(simp add: Sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
946 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
947 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
948 |
lemma Sigma_max_point: "\<lbrakk>\<forall> k < ma. f (xs @ [k]) = 1; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
949 |
\<forall> k \<ge> ma. f (xs @ [k]) = 0; ma \<le> w\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
950 |
\<Longrightarrow> Sigma f (xs @ [w]) = ma" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
951 |
apply(induct w, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
952 |
apply(rule_tac Sigma_0, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
953 |
apply(simp add: Sigma_Suc_simp_rewrite) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
954 |
apply(case_tac "ma = Suc w", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
955 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
956 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
957 |
lemma Sigma_Max_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
958 |
assumes prrf: "primerec rf (Suc (length xs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
959 |
shows "UF.Sigma (rec_exec (Cn (Suc (Suc (length xs))) rec_not |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
960 |
[rec_all (recf.id (Suc (Suc (length xs))) (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
961 |
(Cn (Suc (Suc (Suc (length xs)))) rec_disj |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
962 |
[Cn (Suc (Suc (Suc (length xs)))) rec_le |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
963 |
[recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs))), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
964 |
recf.id (Suc (Suc (Suc (length xs)))) (Suc (length xs))], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
965 |
Cn (Suc (Suc (Suc (length xs)))) rec_not |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
966 |
[Cn (Suc (Suc (Suc (length xs)))) rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
967 |
(get_fstn_args (Suc (Suc (Suc (length xs)))) (length xs) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
968 |
[recf.id (Suc (Suc (Suc (length xs)))) (Suc (Suc (length xs)))])]])])) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
969 |
((xs @ [w]) @ [w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
970 |
Maxr (\<lambda>args. 0 < rec_exec rf args) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
971 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
972 |
let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
973 |
let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
974 |
rec_le [recf.id (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
975 |
((Suc (Suc (length xs)))), recf.id |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
976 |
(Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
977 |
let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
978 |
(get_fstn_args (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
979 |
(length xs) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
980 |
[recf.id (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
981 |
((Suc (Suc (length xs))))])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
982 |
let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
983 |
let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
984 |
let ?rq = "rec_all ?rt ?rf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
985 |
let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
986 |
show "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
987 |
proof(auto simp: Maxr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
988 |
assume h: "\<forall>x\<le>w. rec_exec rf (xs @ [x]) = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
989 |
have "primerec ?rf (Suc (length (xs @ [w, i]))) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
990 |
primerec ?rt (length (xs @ [w, i]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
991 |
using prrf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
992 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
993 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
994 |
apply(case_tac ia, auto simp: h nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
995 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
996 |
hence "Sigma (rec_exec ?notrq) ((xs@[w])@[w]) = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
997 |
apply(rule_tac Sigma_0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
998 |
apply(auto simp: rec_exec.simps all_lemma |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
999 |
get_fstn_args_take nth_append h) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1000 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1001 |
thus "UF.Sigma (rec_exec ?notrq) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1002 |
(xs @ [w, w]) = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1003 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1004 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1005 |
fix x |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1006 |
assume h: "x \<le> w" "0 < rec_exec rf (xs @ [x])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1007 |
hence "\<exists> ma. Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} = ma" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1008 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1009 |
from this obtain ma where k1: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1010 |
"Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])} = ma" .. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1011 |
hence k2: "ma \<le> w \<and> 0 < rec_exec rf (xs @ [ma])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1012 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1013 |
apply(subgoal_tac |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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])}") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1015 |
apply(erule_tac CollectE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1016 |
apply(rule_tac Max_in, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1017 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1018 |
hence k3: "\<forall> k < ma. (rec_exec ?notrq (xs @ [w, k]) = 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1019 |
apply(auto simp: nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1020 |
apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1021 |
primerec ?rt (length (xs @ [w, k]))") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1022 |
apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1023 |
using prrf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1024 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1025 |
apply(case_tac ia, auto simp: h nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1026 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1027 |
have k4: "\<forall> k \<ge> ma. (rec_exec ?notrq (xs @ [w, k]) = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1028 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1029 |
apply(subgoal_tac "primerec ?rf (Suc (length (xs @ [w, k]))) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1030 |
primerec ?rt (length (xs @ [w, k]))") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1031 |
apply(auto simp: rec_exec.simps all_lemma get_fstn_args_take nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1032 |
apply(subgoal_tac "x \<le> Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])}", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1033 |
simp add: k1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1034 |
apply(rule_tac Max_ge, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1035 |
using prrf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1036 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1037 |
apply(case_tac ia, auto simp: h nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1038 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1039 |
from k3 k4 k1 have "Sigma (rec_exec ?notrq) ((xs @ [w]) @ [w]) = ma" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1040 |
apply(rule_tac Sigma_max_point, simp, simp, simp add: k2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1041 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1042 |
from k1 and this show "Sigma (rec_exec ?notrq) (xs @ [w, w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1043 |
Max {y. y \<le> w \<and> 0 < rec_exec rf (xs @ [y])}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1044 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1045 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1046 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1047 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1048 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1049 |
The correctness of @{text "rec_maxr"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1050 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1051 |
lemma Maxr_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1052 |
assumes h: "primerec rf (Suc (length xs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1053 |
shows "rec_exec (rec_maxr rf) (xs @ [w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1054 |
Maxr (\<lambda> args. 0 < rec_exec rf args) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1055 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1056 |
from h have "arity rf = Suc (length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1057 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1058 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1059 |
proof(simp add: rec_exec.simps rec_maxr.simps nth_append get_fstn_args_take) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1060 |
let ?rt = "(recf.id (Suc (Suc (length xs))) ((length xs)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1061 |
let ?rf1 = "Cn (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1062 |
rec_le [recf.id (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1063 |
((Suc (Suc (length xs)))), recf.id |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1064 |
(Suc (Suc (Suc (length xs)))) ((Suc (length xs)))]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1065 |
let ?rf2 = "Cn (Suc (Suc (Suc (length xs)))) rf |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1066 |
(get_fstn_args (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1067 |
(length xs) @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1068 |
[recf.id (Suc (Suc (Suc (length xs)))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1069 |
((Suc (Suc (length xs))))])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1070 |
let ?rf3 = "Cn (Suc (Suc (Suc (length xs)))) rec_not [?rf2]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1071 |
let ?rf = "Cn (Suc (Suc (Suc (length xs)))) rec_disj [?rf1, ?rf3]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1072 |
let ?rq = "rec_all ?rt ?rf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1073 |
let ?notrq = "Cn (Suc (Suc (length xs))) rec_not [?rq]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1074 |
have prt: "primerec ?rt (Suc (Suc (length xs)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1075 |
by(auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1076 |
have prrf: "primerec ?rf (Suc (Suc (Suc (length xs))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1077 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1078 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1079 |
apply(case_tac ia, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1080 |
apply(simp add: h) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1081 |
apply(simp add: nth_append, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1082 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1083 |
from prt and prrf have prrq: "primerec ?rq |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1084 |
(Suc (Suc (length xs)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1085 |
by(erule_tac primerec_all_iff, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1086 |
hence prnotrp: "primerec ?notrq (Suc (length ((xs @ [w]))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1087 |
by(rule_tac prime_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1088 |
have g1: "rec_exec (rec_sigma ?notrq) ((xs @ [w]) @ [w]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1089 |
= Maxr (\<lambda>args. 0 < rec_exec rf args) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1090 |
using prnotrp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1091 |
using sigma_lemma |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1092 |
apply(simp only: sigma_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1093 |
apply(rule_tac Sigma_Max_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1094 |
apply(simp add: h) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1095 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1096 |
thus "rec_exec (rec_sigma ?notrq) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1097 |
(xs @ [w, w]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1098 |
Maxr (\<lambda>args. 0 < rec_exec rf args) xs w" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1099 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1100 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1101 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1102 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1103 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1104 |
text {* |
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
1105 |
@{text "quo"} is the formal specification of division. |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1106 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1107 |
fun quo :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1108 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1109 |
"quo [x, y] = (let Rr = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1110 |
(\<lambda> zs. ((zs ! (Suc 0) * zs ! (Suc (Suc 0)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1111 |
\<le> zs ! 0) \<and> zs ! Suc 0 \<noteq> (0::nat))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1112 |
in Maxr Rr [x, y] x)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1113 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1114 |
declare quo.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1115 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1116 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1117 |
The following lemmas shows more directly the menaing of @{text "quo"}: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1118 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1119 |
lemma [elim!]: "y > 0 \<Longrightarrow> quo [x, y] = x div y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1120 |
proof(simp add: quo.simps Maxr.simps, auto, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1121 |
rule_tac Max_eqI, simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1122 |
fix xa ya |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1123 |
assume h: "y * ya \<le> x" "y > 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1124 |
hence "(y * ya) div y \<le> x div y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1125 |
by(insert div_le_mono[of "y * ya" x y], simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1126 |
from this and h show "ya \<le> x div y" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1127 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1128 |
fix xa |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1129 |
show "y * (x div y) \<le> x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1130 |
apply(subgoal_tac "y * (x div y) + x mod y = x") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1131 |
apply(rule_tac k = "x mod y" in add_leD1, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1132 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1133 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1134 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1135 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1136 |
lemma [intro]: "quo [x, 0] = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1137 |
by(simp add: quo.simps Maxr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1138 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1139 |
lemma quo_div: "quo [x, y] = x div y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1140 |
by(case_tac "y=0", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1141 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1142 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1143 |
@{text "rec_noteq"} is the recursive function testing whether its |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1144 |
two arguments are not equal. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1145 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1146 |
definition rec_noteq:: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1147 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1148 |
"rec_noteq = Cn (Suc (Suc 0)) rec_not [Cn (Suc (Suc 0)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1149 |
rec_eq [id (Suc (Suc 0)) (0), id (Suc (Suc 0)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1150 |
((Suc 0))]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1151 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1152 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1153 |
The correctness of @{text "rec_noteq"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1154 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1155 |
lemma noteq_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1156 |
"\<And> x y. rec_exec rec_noteq [x, y] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1157 |
(if x \<noteq> y then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1158 |
by(simp add: rec_exec.simps rec_noteq_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1159 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1160 |
declare noteq_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1161 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1162 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1163 |
@{text "rec_quo"} is the recursive function used to implement @{text "quo"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1164 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1165 |
definition rec_quo :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1166 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1167 |
"rec_quo = (let rR = Cn (Suc (Suc (Suc 0))) rec_conj |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1168 |
[Cn (Suc (Suc (Suc 0))) rec_le |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1169 |
[Cn (Suc (Suc (Suc 0))) rec_mult |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1170 |
[id (Suc (Suc (Suc 0))) (Suc 0), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1171 |
id (Suc (Suc (Suc 0))) ((Suc (Suc 0)))], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1172 |
id (Suc (Suc (Suc 0))) (0)], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1173 |
Cn (Suc (Suc (Suc 0))) rec_noteq |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1174 |
[id (Suc (Suc (Suc 0))) (Suc (0)), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1175 |
Cn (Suc (Suc (Suc 0))) (constn 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1176 |
[id (Suc (Suc (Suc 0))) (0)]]] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1177 |
in Cn (Suc (Suc 0)) (rec_maxr rR)) [id (Suc (Suc 0)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1178 |
(0),id (Suc (Suc 0)) (Suc (0)), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1179 |
id (Suc (Suc 0)) (0)]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1180 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1181 |
lemma [intro]: "primerec rec_conj (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1182 |
apply(simp add: rec_conj_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1183 |
apply(rule_tac prime_cn, auto)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1184 |
apply(case_tac i, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1185 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1186 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1187 |
lemma [intro]: "primerec rec_noteq (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1188 |
apply(simp add: rec_noteq_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1189 |
apply(rule_tac prime_cn, auto)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1190 |
apply(case_tac i, auto intro: prime_id) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1191 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1192 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1193 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1194 |
lemma quo_lemma1: "rec_exec rec_quo [x, y] = quo [x, y]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1195 |
proof(simp add: rec_exec.simps rec_quo_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1196 |
let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_conj |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1197 |
[Cn (Suc (Suc (Suc 0))) rec_le |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1198 |
[Cn (Suc (Suc (Suc 0))) rec_mult |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1199 |
[recf.id (Suc (Suc (Suc 0))) (Suc (0)), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1200 |
recf.id (Suc (Suc (Suc 0))) (Suc (Suc (0)))], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1201 |
recf.id (Suc (Suc (Suc 0))) (0)], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1202 |
Cn (Suc (Suc (Suc 0))) rec_noteq |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1203 |
[recf.id (Suc (Suc (Suc 0))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1204 |
(Suc (0)), Cn (Suc (Suc (Suc 0))) (constn 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1205 |
[recf.id (Suc (Suc (Suc 0))) (0)]]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1206 |
have "rec_exec (rec_maxr ?rR) ([x, y]@ [ x]) = Maxr (\<lambda> args. 0 < rec_exec ?rR args) [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1207 |
proof(rule_tac Maxr_lemma, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1208 |
show "primerec ?rR (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1209 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1210 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1211 |
apply(case_tac [!] ia, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1212 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1213 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1214 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1215 |
hence g1: "rec_exec (rec_maxr ?rR) ([x, y, x]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1216 |
Maxr (\<lambda> args. if rec_exec ?rR args = 0 then False |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1217 |
else True) [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1218 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1219 |
have g2: "Maxr (\<lambda> args. if rec_exec ?rR args = 0 then False |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1220 |
else True) [x, y] x = quo [x, y]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1221 |
apply(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1222 |
apply(simp add: Maxr.simps quo.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1223 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1224 |
from g1 and g2 show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1225 |
"rec_exec (rec_maxr ?rR) ([x, y, x]) = quo [x, y]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1226 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1227 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1228 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1229 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1230 |
The correctness of @{text "quo"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1231 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1232 |
lemma quo_lemma2: "rec_exec rec_quo [x, y] = x div y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1233 |
using quo_lemma1[of x y] quo_div[of x y] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1234 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1235 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1236 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1237 |
@{text "rec_mod"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1238 |
the reminder function. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1239 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1240 |
definition rec_mod :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1241 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1242 |
"rec_mod = Cn (Suc (Suc 0)) rec_minus [id (Suc (Suc 0)) (0), |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1243 |
Cn (Suc (Suc 0)) rec_mult [rec_quo, id (Suc (Suc 0)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1244 |
(Suc (0))]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1245 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1246 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1247 |
The correctness of @{text "rec_mod"}: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1248 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1249 |
lemma mod_lemma: "\<And> x y. rec_exec rec_mod [x, y] = (x mod y)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1250 |
proof(simp add: rec_exec.simps rec_mod_def quo_lemma2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1251 |
fix x y |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1252 |
show "x - x div y * y = x mod (y::nat)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1253 |
using mod_div_equality2[of y x] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1254 |
apply(subgoal_tac "y * (x div y) = (x div y ) * y", arith, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1255 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1256 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1257 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1258 |
text{* lemmas for embranch function*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1259 |
type_synonym ftype = "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1260 |
type_synonym rtype = "nat list \<Rightarrow> bool" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1261 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1262 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1263 |
The specifation of the mutli-way branching statement on |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1264 |
page 79 of Boolos's book. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1265 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1266 |
fun Embranch :: "(ftype * rtype) list \<Rightarrow> nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1267 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1268 |
"Embranch [] xs = 0" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1269 |
"Embranch (gc # gcs) xs = ( |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1270 |
let (g, c) = gc in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1271 |
if c xs then g xs else Embranch gcs xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1272 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1273 |
fun rec_embranch' :: "(recf * recf) list \<Rightarrow> nat \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1274 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1275 |
"rec_embranch' [] vl = Cn vl z [id vl (vl - 1)]" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1276 |
"rec_embranch' ((rg, rc) # rgcs) vl = Cn vl rec_add |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1277 |
[Cn vl rec_mult [rg, rc], rec_embranch' rgcs vl]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1278 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1279 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1280 |
@{text "rec_embrach"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1281 |
@{text "Embranch"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1282 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1283 |
fun rec_embranch :: "(recf * recf) list \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1284 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1285 |
"rec_embranch ((rg, rc) # rgcs) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1286 |
(let vl = arity rg in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1287 |
rec_embranch' ((rg, rc) # rgcs) vl)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1288 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1289 |
declare Embranch.simps[simp del] rec_embranch.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1290 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1291 |
lemma embranch_all0: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1292 |
"\<lbrakk>\<forall> j < length rcs. rec_exec (rcs ! j) xs = 0; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1293 |
length rgs = length rcs; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1294 |
rcs \<noteq> []; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1295 |
list_all (\<lambda> rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1296 |
rec_exec (rec_embranch (zip rgs rcs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1297 |
proof(induct rcs arbitrary: rgs, simp, case_tac rgs, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1298 |
fix a rcs rgs aa list |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1299 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1300 |
"\<And>rgs. \<lbrakk>\<forall>j<length rcs. rec_exec (rcs ! j) xs = 0; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1301 |
length rgs = length rcs; rcs \<noteq> []; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1302 |
list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1303 |
rec_exec (rec_embranch (zip rgs rcs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1304 |
and h: "\<forall>j<length (a # rcs). rec_exec ((a # rcs) ! j) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1305 |
"length rgs = length (a # rcs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1306 |
"a # rcs \<noteq> []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1307 |
"list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ a # rcs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1308 |
"rgs = aa # list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1309 |
have g: "rcs \<noteq> [] \<Longrightarrow> rec_exec (rec_embranch (zip list rcs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1310 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1311 |
by(rule_tac ind, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1312 |
show "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1313 |
proof(case_tac "rcs = []", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1314 |
show "rec_exec (rec_embranch (zip rgs [a])) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1315 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1316 |
apply(simp add: rec_embranch.simps rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1317 |
apply(erule_tac x = 0 in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1318 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1319 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1320 |
assume "rcs \<noteq> []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1321 |
hence "rec_exec (rec_embranch (zip list rcs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1322 |
using g by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1323 |
thus "rec_exec (rec_embranch (zip rgs (a # rcs))) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1324 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1325 |
apply(simp add: rec_embranch.simps rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1326 |
apply(case_tac rcs, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1327 |
auto simp: rec_exec.simps rec_embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1328 |
apply(case_tac list, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1329 |
auto simp: rec_exec.simps rec_embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1330 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1331 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1332 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1333 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1334 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1335 |
lemma embranch_exec_0: "\<lbrakk>rec_exec aa xs = 0; zip rgs list \<noteq> []; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1336 |
list_all (\<lambda> rf. primerec rf (length xs)) ([a, aa] @ rgs @ list)\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1337 |
\<Longrightarrow> rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1338 |
= rec_exec (rec_embranch (zip rgs list)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1339 |
apply(simp add: rec_exec.simps rec_embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1340 |
apply(case_tac "zip rgs list", simp, case_tac ab, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1341 |
simp add: rec_embranch.simps rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1342 |
apply(subgoal_tac "arity a = length xs", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1343 |
apply(subgoal_tac "arity aaa = length xs", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1344 |
apply(case_tac rgs, simp, case_tac list, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1345 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1346 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1347 |
lemma zip_null_iff: "\<lbrakk>length xs = k; length ys = k; zip xs ys = []\<rbrakk> \<Longrightarrow> xs = [] \<and> ys = []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1348 |
apply(case_tac xs, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1349 |
apply(case_tac ys, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1350 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1351 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1352 |
lemma zip_null_gr: "\<lbrakk>length xs = k; length ys = k; zip xs ys \<noteq> []\<rbrakk> \<Longrightarrow> 0 < k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1353 |
apply(case_tac xs, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1354 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1355 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1356 |
lemma Embranch_0: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1357 |
"\<lbrakk>length rgs = k; length rcs = k; k > 0; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1358 |
\<forall> j < k. rec_exec (rcs ! j) xs = 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1359 |
Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1360 |
proof(induct rgs arbitrary: rcs k, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1361 |
fix a rgs rcs k |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1362 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1363 |
"\<And>rcs k. \<lbrakk>length rgs = k; length rcs = k; 0 < k; \<forall>j<k. rec_exec (rcs ! j) xs = 0\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1364 |
\<Longrightarrow> Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1365 |
and h: "Suc (length rgs) = k" "length rcs = k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1366 |
"\<forall>j<k. rec_exec (rcs ! j) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1367 |
from h show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1368 |
"Embranch (zip (rec_exec a # map rec_exec rgs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1369 |
(map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1370 |
apply(case_tac rcs, simp, case_tac "rgs = []", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1371 |
apply(simp add: Embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1372 |
apply(erule_tac x = 0 in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1373 |
apply(simp add: Embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1374 |
apply(erule_tac x = 0 in all_dupE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1375 |
apply(rule_tac ind, simp, simp, simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1376 |
apply(erule_tac x = "Suc j" in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1377 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1378 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1379 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1380 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1381 |
The correctness of @{text "rec_embranch"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1382 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1383 |
lemma embranch_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1384 |
assumes branch_num: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1385 |
"length rgs = n" "length rcs = n" "n > 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1386 |
and partition: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1387 |
"(\<exists> i < n. (rec_exec (rcs ! i) xs = 1 \<and> (\<forall> j < n. j \<noteq> i \<longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1388 |
rec_exec (rcs ! j) xs = 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1389 |
and prime_all: "list_all (\<lambda> rf. primerec rf (length xs)) (rgs @ rcs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1390 |
shows "rec_exec (rec_embranch (zip rgs rcs)) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1391 |
Embranch (zip (map rec_exec rgs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1392 |
(map (\<lambda> r args. 0 < rec_exec r args) rcs)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1393 |
using branch_num partition prime_all |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1394 |
proof(induct rgs arbitrary: rcs n, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1395 |
fix a rgs rcs n |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1396 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1397 |
"\<And>rcs n. \<lbrakk>length rgs = n; length rcs = n; 0 < n; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1398 |
\<exists>i<n. rec_exec (rcs ! i) xs = 1 \<and> (\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec (rcs ! j) xs = 0); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1399 |
list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ rcs)\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1400 |
\<Longrightarrow> rec_exec (rec_embranch (zip rgs rcs)) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1401 |
Embranch (zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) rcs)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1402 |
and h: "length (a # rgs) = n" "length (rcs::recf list) = n" "0 < n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1403 |
" \<exists>i<n. rec_exec (rcs ! i) xs = 1 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1404 |
(\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec (rcs ! j) xs = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1405 |
"list_all (\<lambda>rf. primerec rf (length xs)) ((a # rgs) @ rcs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1406 |
from h show "rec_exec (rec_embranch (zip (a # rgs) rcs)) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1407 |
Embranch (zip (map rec_exec (a # rgs)) (map (\<lambda>r args. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1408 |
0 < rec_exec r args) rcs)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1409 |
apply(case_tac rcs, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1410 |
apply(case_tac "rec_exec aa xs = 0") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1411 |
apply(case_tac [!] "zip rgs list = []", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1412 |
apply(subgoal_tac "rgs = [] \<and> list = []", simp add: Embranch.simps rec_exec.simps rec_embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1413 |
apply(rule_tac zip_null_iff, simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1414 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1415 |
fix aa list |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1416 |
assume g: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1417 |
"Suc (length rgs) = n" "Suc (length list) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1418 |
"\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1419 |
(\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1420 |
"primerec a (length xs) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1421 |
list_all (\<lambda>rf. primerec rf (length xs)) rgs \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1422 |
primerec aa (length xs) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1423 |
list_all (\<lambda>rf. primerec rf (length xs)) list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1424 |
"rec_exec aa xs = 0" "rcs = aa # list" "zip rgs list \<noteq> []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1425 |
have "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1426 |
= rec_exec (rec_embranch (zip rgs list)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1427 |
apply(rule embranch_exec_0, simp_all add: g) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1428 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1429 |
from g and this show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1430 |
Embranch ((rec_exec a, \<lambda>args. 0 < rec_exec aa args) # |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1431 |
zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) list)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1432 |
apply(simp add: Embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1433 |
apply(rule_tac n = "n - Suc 0" in ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1434 |
apply(case_tac n, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1435 |
apply(case_tac n, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1436 |
apply(case_tac n, simp, simp add: zip_null_gr ) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1437 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1438 |
apply(case_tac i, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1439 |
apply(rule_tac x = nat in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1440 |
apply(rule_tac allI, erule_tac x = "Suc j" in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1441 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1442 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1443 |
fix aa list |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1444 |
assume g: "Suc (length rgs) = n" "Suc (length list) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1445 |
"\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1446 |
(\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1447 |
"primerec a (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) rgs \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1448 |
primerec aa (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1449 |
"rcs = aa # list" "rec_exec aa xs \<noteq> 0" "zip rgs list = []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1450 |
thus "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1451 |
Embranch ((rec_exec a, \<lambda>args. 0 < rec_exec aa args) # |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1452 |
zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) list)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1453 |
apply(subgoal_tac "rgs = [] \<and> list = []", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1454 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1455 |
apply(rule_tac zip_null_iff, simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1456 |
apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1457 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1458 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1459 |
fix aa list |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1460 |
assume g: "Suc (length rgs) = n" "Suc (length list) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1461 |
"\<exists>i<n. rec_exec ((aa # list) ! i) xs = Suc 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1462 |
(\<forall>j<n. j \<noteq> i \<longrightarrow> rec_exec ((aa # list) ! j) xs = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1463 |
"primerec a (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) rgs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1464 |
\<and> primerec aa (length xs) \<and> list_all (\<lambda>rf. primerec rf (length xs)) list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1465 |
"rcs = aa # list" "rec_exec aa xs \<noteq> 0" "zip rgs list \<noteq> []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1466 |
have "rec_exec aa xs = Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1467 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1468 |
apply(case_tac "rec_exec aa xs", simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1469 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1470 |
moreover have "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1471 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1472 |
have "rec_embranch' (zip rgs list) (length xs) = rec_embranch (zip rgs list)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1473 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1474 |
apply(case_tac "zip rgs list", simp, case_tac ab) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1475 |
apply(simp add: rec_embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1476 |
apply(subgoal_tac "arity aaa = length xs", simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1477 |
apply(case_tac rgs, simp, simp, case_tac list, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1478 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1479 |
moreover have "rec_exec (rec_embranch (zip rgs list)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1480 |
proof(rule embranch_all0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1481 |
show " \<forall>j<length list. rec_exec (list ! j) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1482 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1483 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1484 |
apply(case_tac i, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1485 |
apply(erule_tac x = "Suc j" in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1486 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1487 |
apply(erule_tac x = 0 in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1488 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1489 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1490 |
show "length rgs = length list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1491 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1492 |
apply(case_tac n, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1493 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1494 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1495 |
show "list \<noteq> []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1496 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1497 |
apply(case_tac list, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1498 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1499 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1500 |
show "list_all (\<lambda>rf. primerec rf (length xs)) (rgs @ list)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1501 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1502 |
apply auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1503 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1504 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1505 |
ultimately show "rec_exec (rec_embranch' (zip rgs list) (length xs)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1506 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1507 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1508 |
moreover have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1509 |
"Embranch (zip (map rec_exec rgs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1510 |
(map (\<lambda>r args. 0 < rec_exec r args) list)) xs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1511 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1512 |
apply(rule_tac k = "length rgs" in Embranch_0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1513 |
apply(simp, case_tac n, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1514 |
apply(case_tac rgs, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1515 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1516 |
apply(case_tac i, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1517 |
apply(erule_tac x = "Suc j" in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1518 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1519 |
apply(rule_tac x = 0 in allE, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1520 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1521 |
moreover have "arity a = length xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1522 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1523 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1524 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1525 |
ultimately show "rec_exec (rec_embranch ((a, aa) # zip rgs list)) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1526 |
Embranch ((rec_exec a, \<lambda>args. 0 < rec_exec aa args) # |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1527 |
zip (map rec_exec rgs) (map (\<lambda>r args. 0 < rec_exec r args) list)) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1528 |
apply(simp add: rec_exec.simps rec_embranch.simps Embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1529 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1530 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1531 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1532 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1533 |
text{* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1534 |
@{text "prime n"} means @{text "n"} is a prime number. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1535 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1536 |
fun Prime :: "nat \<Rightarrow> bool" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1537 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1538 |
"Prime x = (1 < x \<and> (\<forall> u < x. (\<forall> v < x. u * v \<noteq> x)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1539 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1540 |
declare Prime.simps [simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1541 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1542 |
lemma primerec_all1: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1543 |
"primerec (rec_all rt rf) n \<Longrightarrow> primerec rt n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1544 |
by (simp add: primerec_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1545 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1546 |
lemma primerec_all2: "primerec (rec_all rt rf) n \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1547 |
primerec rf (Suc n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1548 |
by(insert primerec_all[of rt rf n], simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1549 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1550 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1551 |
@{text "rec_prime"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1552 |
@{text "Prime"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1553 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1554 |
definition rec_prime :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1555 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1556 |
"rec_prime = Cn (Suc 0) rec_conj |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1557 |
[Cn (Suc 0) rec_less [constn 1, id (Suc 0) (0)], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1558 |
rec_all (Cn 1 rec_minus [id 1 0, constn 1]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1559 |
(rec_all (Cn 2 rec_minus [id 2 0, Cn 2 (constn 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1560 |
[id 2 0]]) (Cn 3 rec_noteq |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1561 |
[Cn 3 rec_mult [id 3 1, id 3 2], id 3 0]))]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1562 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1563 |
declare numeral_2_eq_2[simp del] numeral_3_eq_3[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1564 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1565 |
lemma exec_tmp: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1566 |
"rec_exec (rec_all (Cn 2 rec_minus [recf.id 2 0, Cn 2 (constn (Suc 0)) [recf.id 2 0]]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1567 |
(Cn 3 rec_noteq [Cn 3 rec_mult [recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])) [x, k] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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]). |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1570 |
([x, k] @ [w])) then 1 else 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1571 |
apply(rule_tac all_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1572 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1573 |
apply(case_tac [!] i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1574 |
apply(case_tac ia, auto simp: numeral_3_eq_3 numeral_2_eq_2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1575 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1576 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1577 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1578 |
The correctness of @{text "Prime"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1579 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1580 |
lemma prime_lemma: "rec_exec rec_prime [x] = (if Prime x then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1581 |
proof(simp add: rec_exec.simps rec_prime_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1582 |
let ?rt1 = "(Cn 2 rec_minus [recf.id 2 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1583 |
Cn 2 (constn (Suc 0)) [recf.id 2 0]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1584 |
let ?rf1 = "(Cn 3 rec_noteq [Cn 3 rec_mult |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1585 |
[recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 (0)])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1586 |
let ?rt2 = "(Cn (Suc 0) rec_minus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1587 |
[recf.id (Suc 0) 0, constn (Suc 0)])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1588 |
let ?rf2 = "rec_all ?rt1 ?rf1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1589 |
have h1: "rec_exec (rec_all ?rt2 ?rf2) ([x]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1590 |
(if (\<forall>k\<le>rec_exec ?rt2 ([x]). 0 < rec_exec ?rf2 ([x] @ [k])) then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1591 |
proof(rule_tac all_lemma, simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1592 |
show "primerec ?rf2 (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1593 |
apply(rule_tac primerec_all_iff) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1594 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1595 |
apply(case_tac [!] i, auto simp: numeral_2_eq_2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1596 |
apply(case_tac ia, auto simp: numeral_3_eq_3) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1597 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1598 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1599 |
show "primerec (Cn (Suc 0) rec_minus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1600 |
[recf.id (Suc 0) 0, constn (Suc 0)]) (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1601 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1602 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1603 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1604 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1605 |
from h1 show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1606 |
"(Suc 0 < x \<longrightarrow> (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 \<longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1607 |
\<not> Prime x) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1608 |
(0 < rec_exec (rec_all ?rt2 ?rf2) [x] \<longrightarrow> Prime x)) \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1609 |
(\<not> Suc 0 < x \<longrightarrow> \<not> Prime x \<and> (rec_exec (rec_all ?rt2 ?rf2) [x] = 0 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1610 |
\<longrightarrow> \<not> Prime x))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1611 |
apply(auto simp:rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1612 |
apply(simp add: exec_tmp rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1613 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1614 |
assume "\<forall>k\<le>x - Suc 0. (0::nat) < (if \<forall>w\<le>x - Suc 0. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1615 |
0 < (if k * w \<noteq> x then 1 else (0 :: nat)) then 1 else 0)" "Suc 0 < x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1616 |
thus "Prime x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1617 |
apply(simp add: rec_exec.simps split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1618 |
apply(simp add: Prime.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1619 |
apply(erule_tac x = u in allE, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1620 |
apply(case_tac u, simp, case_tac nat, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1621 |
apply(case_tac v, simp, case_tac nat, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1622 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1623 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1624 |
assume "\<not> Suc 0 < x" "Prime x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1625 |
thus "False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1626 |
apply(simp add: Prime.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1627 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1628 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1629 |
fix k |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1630 |
assume "rec_exec (rec_all ?rt1 ?rf1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1631 |
[x, k] = 0" "k \<le> x - Suc 0" "Prime x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1632 |
thus "False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1633 |
apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1634 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1635 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1636 |
fix k |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1637 |
assume "rec_exec (rec_all ?rt1 ?rf1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1638 |
[x, k] = 0" "k \<le> x - Suc 0" "Prime x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1639 |
thus "False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1640 |
apply(simp add: exec_tmp rec_exec.simps Prime.simps split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1641 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1642 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1643 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1644 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1645 |
definition rec_dummyfac :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1646 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1647 |
"rec_dummyfac = Pr 1 (constn 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1648 |
(Cn 3 rec_mult [id 3 2, Cn 3 s [id 3 1]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1649 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1650 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1651 |
The recursive function used to implment factorization. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1652 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1653 |
definition rec_fac :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1654 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1655 |
"rec_fac = Cn 1 rec_dummyfac [id 1 0, id 1 0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1656 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1657 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1658 |
Formal specification of factorization. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1659 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1660 |
fun fac :: "nat \<Rightarrow> nat" ("_!" [100] 99) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1661 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1662 |
"fac 0 = 1" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1663 |
"fac (Suc x) = (Suc x) * fac x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1664 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1665 |
lemma [simp]: "rec_exec rec_dummyfac [0, 0] = Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1666 |
by(simp add: rec_dummyfac_def rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1667 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1668 |
lemma rec_cn_simp: "rec_exec (Cn n f gs) xs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1669 |
(let rgs = map (\<lambda> g. rec_exec g xs) gs in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1670 |
rec_exec f rgs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1671 |
by(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1672 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1673 |
lemma rec_id_simp: "rec_exec (id m n) xs = xs ! n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1674 |
by(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1675 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1676 |
lemma fac_dummy: "rec_exec rec_dummyfac [x, y] = y !" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1677 |
apply(induct y) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1678 |
apply(auto simp: rec_dummyfac_def rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1679 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1680 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1681 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1682 |
The correctness of @{text "rec_fac"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1683 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1684 |
lemma fac_lemma: "rec_exec rec_fac [x] = x!" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1685 |
apply(simp add: rec_fac_def rec_exec.simps fac_dummy) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1686 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1687 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1688 |
declare fac.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1689 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1690 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1691 |
@{text "Np x"} returns the first prime number after @{text "x"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1692 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1693 |
fun Np ::"nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1694 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1695 |
"Np x = Min {y. y \<le> Suc (x!) \<and> x < y \<and> Prime y}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1696 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1697 |
declare Np.simps[simp del] rec_Minr.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1698 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1699 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1700 |
@{text "rec_np"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1701 |
@{text "Np"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1702 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1703 |
definition rec_np :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1704 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1705 |
"rec_np = (let Rr = Cn 2 rec_conj [Cn 2 rec_less [id 2 0, id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1706 |
Cn 2 rec_prime [id 2 1]] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1707 |
in Cn 1 (rec_Minr Rr) [id 1 0, Cn 1 s [rec_fac]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1708 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1709 |
lemma [simp]: "n < Suc (n!)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1710 |
apply(induct n, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1711 |
apply(simp add: fac.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1712 |
apply(case_tac n, auto simp: fac.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1713 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1714 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1715 |
lemma divsor_ex: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1716 |
"\<lbrakk>\<not> Prime x; x > Suc 0\<rbrakk> \<Longrightarrow> (\<exists> u > Suc 0. (\<exists> v > Suc 0. u * v = x))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1717 |
by(auto simp: Prime.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1718 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1719 |
lemma divsor_prime_ex: "\<lbrakk>\<not> Prime x; x > Suc 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1720 |
\<exists> p. Prime p \<and> p dvd x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1721 |
apply(induct x rule: wf_induct[where r = "measure (\<lambda> y. y)"], simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1722 |
apply(drule_tac divsor_ex, simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1723 |
apply(erule_tac x = u in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1724 |
apply(case_tac "Prime u", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1725 |
apply(rule_tac x = u in exI, simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1726 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1727 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1728 |
lemma [intro]: "0 < n!" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1729 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1730 |
apply(auto simp: fac.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1731 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1732 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1733 |
lemma fac_Suc: "Suc n! = (Suc n) * (n!)" by(simp add: fac.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1734 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1735 |
lemma fac_dvd: "\<lbrakk>0 < q; q \<le> n\<rbrakk> \<Longrightarrow> q dvd n!" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1736 |
apply(induct n, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1737 |
apply(case_tac "q \<le> n", simp add: fac_Suc) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1738 |
apply(subgoal_tac "q = Suc n", simp only: fac_Suc) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1739 |
apply(rule_tac dvd_mult2, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1740 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1741 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1742 |
lemma fac_dvd2: "\<lbrakk>Suc 0 < q; q dvd n!; q \<le> n\<rbrakk> \<Longrightarrow> \<not> q dvd Suc (n!)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1743 |
proof(auto simp: dvd_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1744 |
fix k ka |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1745 |
assume h1: "Suc 0 < q" "q \<le> n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1746 |
and h2: "Suc (q * k) = q * ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1747 |
have "k < ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1748 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1749 |
have "q * k < q * ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1750 |
using h2 by arith |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1751 |
thus "k < ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1752 |
using h1 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1753 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1754 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1755 |
hence "\<exists>d. d > 0 \<and> ka = d + k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1756 |
by(rule_tac x = "ka - k" in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1757 |
from this obtain d where "d > 0 \<and> ka = d + k" .. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1758 |
from h2 and this and h1 show "False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1759 |
by(simp add: add_mult_distrib2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1760 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1761 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1762 |
lemma prime_ex: "\<exists> p. n < p \<and> p \<le> Suc (n!) \<and> Prime p" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1763 |
proof(cases "Prime (n! + 1)") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1764 |
case True thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1765 |
by(rule_tac x = "Suc (n!)" in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1766 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1767 |
assume h: "\<not> Prime (n! + 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1768 |
hence "\<exists> p. Prime p \<and> p dvd (n! + 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1769 |
by(erule_tac divsor_prime_ex, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1770 |
from this obtain q where k: "Prime q \<and> q dvd (n! + 1)" .. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1771 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1772 |
proof(cases "q > n") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1773 |
case True thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1774 |
using k |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1775 |
apply(rule_tac x = q in exI, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1776 |
apply(rule_tac dvd_imp_le, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1777 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1778 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1779 |
case False thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1780 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1781 |
assume g: "\<not> n < q" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1782 |
have j: "q > Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1783 |
using k by(case_tac q, auto simp: Prime.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1784 |
hence "q dvd n!" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1785 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1786 |
apply(rule_tac fac_dvd, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1787 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1788 |
hence "\<not> q dvd Suc (n!)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1789 |
using g j |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1790 |
by(rule_tac fac_dvd2, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1791 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1792 |
using k by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1793 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1794 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1795 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1796 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1797 |
lemma Suc_Suc_induct[elim!]: "\<lbrakk>i < Suc (Suc 0); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1798 |
primerec (ys ! 0) n; primerec (ys ! 1) n\<rbrakk> \<Longrightarrow> primerec (ys ! i) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1799 |
by(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1800 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1801 |
lemma [intro]: "primerec rec_prime (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1802 |
apply(auto simp: rec_prime_def, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1803 |
apply(rule_tac primerec_all_iff, auto, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1804 |
apply(rule_tac primerec_all_iff, auto, auto simp: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1805 |
numeral_2_eq_2 numeral_3_eq_3) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1806 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1807 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1808 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1809 |
The correctness of @{text "rec_np"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1810 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1811 |
lemma np_lemma: "rec_exec rec_np [x] = Np x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1812 |
proof(auto simp: rec_np_def rec_exec.simps Let_def fac_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1813 |
let ?rr = "(Cn 2 rec_conj [Cn 2 rec_less [recf.id 2 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1814 |
recf.id 2 (Suc 0)], Cn 2 rec_prime [recf.id 2 (Suc 0)]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1815 |
let ?R = "\<lambda> zs. zs ! 0 < zs ! 1 \<and> Prime (zs ! 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1816 |
have g1: "rec_exec (rec_Minr ?rr) ([x] @ [Suc (x!)]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1817 |
Minr (\<lambda> args. 0 < rec_exec ?rr args) [x] (Suc (x!))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1818 |
by(rule_tac Minr_lemma, auto simp: rec_exec.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1819 |
prime_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1820 |
have g2: "Minr (\<lambda> args. 0 < rec_exec ?rr args) [x] (Suc (x!)) = Np x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1821 |
using prime_ex[of x] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1822 |
apply(auto simp: Minr.simps Np.simps rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1823 |
apply(erule_tac x = p in allE, simp add: prime_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1824 |
apply(simp add: prime_lemma split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1825 |
apply(subgoal_tac |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1826 |
"{uu. (Prime uu \<longrightarrow> (x < uu \<longrightarrow> uu \<le> Suc (x!)) \<and> x < uu) \<and> Prime uu} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1827 |
= {y. y \<le> Suc (x!) \<and> x < y \<and> Prime y}", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1828 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1829 |
from g1 and g2 show "rec_exec (rec_Minr ?rr) ([x, Suc (x!)]) = Np x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1830 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1831 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1832 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1833 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1834 |
@{text "rec_power"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1835 |
power function. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1836 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1837 |
definition rec_power :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1838 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1839 |
"rec_power = Pr 1 (constn 1) (Cn 3 rec_mult [id 3 0, id 3 2])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1840 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1841 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1842 |
The correctness of @{text "rec_power"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1843 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1844 |
lemma power_lemma: "rec_exec rec_power [x, y] = x^y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1845 |
by(induct y, auto simp: rec_exec.simps rec_power_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1846 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1847 |
text{* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1848 |
@{text "Pi k"} returns the @{text "k"}-th prime number. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1849 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1850 |
fun Pi :: "nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1851 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1852 |
"Pi 0 = 2" | |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1853 |
"Pi (Suc x) = Np (Pi x)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1854 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1855 |
definition rec_dummy_pi :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1856 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1857 |
"rec_dummy_pi = Pr 1 (constn 2) (Cn 3 rec_np [id 3 2])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1858 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1859 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1860 |
@{text "rec_pi"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1861 |
@{text "Pi"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1862 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1863 |
definition rec_pi :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1864 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1865 |
"rec_pi = Cn 1 rec_dummy_pi [id 1 0, id 1 0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1866 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1867 |
lemma pi_dummy_lemma: "rec_exec rec_dummy_pi [x, y] = Pi y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1868 |
apply(induct y) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1869 |
by(auto simp: rec_exec.simps rec_dummy_pi_def Pi.simps np_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1870 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1871 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1872 |
The correctness of @{text "rec_pi"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1873 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1874 |
lemma pi_lemma: "rec_exec rec_pi [x] = Pi x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1875 |
apply(simp add: rec_pi_def rec_exec.simps pi_dummy_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1876 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1877 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1878 |
fun loR :: "nat list \<Rightarrow> bool" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1879 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1880 |
"loR [x, y, u] = (x mod (y^u) = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1881 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1882 |
declare loR.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1883 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1884 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1885 |
@{text "Lo"} specifies the @{text "lo"} function given on page 79 of |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1886 |
Boolos's book. It is one of the two notions of integeral logarithmatic |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1887 |
operation on that page. The other is @{text "lg"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1888 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1889 |
fun lo :: " nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1890 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1891 |
"lo x y = (if x > 1 \<and> y > 1 \<and> {u. loR [x, y, u]} \<noteq> {} then Max {u. loR [x, y, u]} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1892 |
else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1893 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1894 |
declare lo.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1895 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1896 |
lemma [elim]: "primerec rf n \<Longrightarrow> n > 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1897 |
apply(induct rule: primerec.induct, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1898 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1899 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1900 |
lemma primerec_sigma[intro!]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1901 |
"\<lbrakk>n > Suc 0; primerec rf n\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1902 |
primerec (rec_sigma rf) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1903 |
apply(simp add: rec_sigma.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1904 |
apply(auto, auto simp: nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1905 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1906 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1907 |
lemma [intro!]: "\<lbrakk>primerec rf n; n > 0\<rbrakk> \<Longrightarrow> primerec (rec_maxr rf) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1908 |
apply(simp add: rec_maxr.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1909 |
apply(rule_tac prime_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1910 |
apply(rule_tac primerec_all_iff, auto, auto simp: nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1911 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1912 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1913 |
lemma Suc_Suc_Suc_induct[elim!]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1914 |
"\<lbrakk>i < Suc (Suc (Suc (0::nat))); primerec (ys ! 0) n; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1915 |
primerec (ys ! 1) n; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1916 |
primerec (ys ! 2) n\<rbrakk> \<Longrightarrow> primerec (ys ! i) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1917 |
apply(case_tac i, auto, case_tac nat, simp, simp add: numeral_2_eq_2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1918 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1919 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1920 |
lemma [intro]: "primerec rec_quo (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1921 |
apply(simp add: rec_quo_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1922 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1923 |
@{thm prime_id}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1924 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1925 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1926 |
lemma [intro]: "primerec rec_mod (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1927 |
apply(simp add: rec_mod_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1928 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1929 |
@{thm prime_id}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1930 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1931 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1932 |
lemma [intro]: "primerec rec_power (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1933 |
apply(simp add: rec_power_def numeral_2_eq_2 numeral_3_eq_3) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1934 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1935 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1936 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1937 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1938 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1939 |
@{text "rec_lo"} is the recursive function used to implement @{text "Lo"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1940 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1941 |
definition rec_lo :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1942 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1943 |
"rec_lo = (let rR = Cn 3 rec_eq [Cn 3 rec_mod [id 3 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1944 |
Cn 3 rec_power [id 3 1, id 3 2]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1945 |
Cn 3 (constn 0) [id 3 1]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1946 |
let rb = Cn 2 (rec_maxr rR) [id 2 0, id 2 1, id 2 0] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1947 |
let rcond = Cn 2 rec_conj [Cn 2 rec_less [Cn 2 (constn 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1948 |
[id 2 0], id 2 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1949 |
Cn 2 rec_less [Cn 2 (constn 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1950 |
[id 2 0], id 2 1]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1951 |
let rcond2 = Cn 2 rec_minus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1952 |
[Cn 2 (constn 1) [id 2 0], rcond] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1953 |
in Cn 2 rec_add [Cn 2 rec_mult [rb, rcond], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1954 |
Cn 2 rec_mult [Cn 2 (constn 0) [id 2 0], rcond2]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1955 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1956 |
lemma rec_lo_Maxr_lor: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1957 |
"\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1958 |
rec_exec rec_lo [x, y] = Maxr loR [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1959 |
proof(auto simp: rec_exec.simps rec_lo_def Let_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1960 |
numeral_2_eq_2 numeral_3_eq_3) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1961 |
let ?rR = "(Cn (Suc (Suc (Suc 0))) rec_eq |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1962 |
[Cn (Suc (Suc (Suc 0))) rec_mod [recf.id (Suc (Suc (Suc 0))) 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1963 |
Cn (Suc (Suc (Suc 0))) rec_power [recf.id (Suc (Suc (Suc 0))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1964 |
(Suc 0), recf.id (Suc (Suc (Suc 0))) (Suc (Suc 0))]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1965 |
Cn (Suc (Suc (Suc 0))) (constn 0) [recf.id (Suc (Suc (Suc 0))) (Suc 0)]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1966 |
have h: "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1967 |
Maxr (\<lambda> args. 0 < rec_exec ?rR args) [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1968 |
by(rule_tac Maxr_lemma, auto simp: rec_exec.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1969 |
mod_lemma power_lemma, auto simp: numeral_2_eq_2 numeral_3_eq_3) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1970 |
have "Maxr loR [x, y] x = Maxr (\<lambda> args. 0 < rec_exec ?rR args) [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1971 |
apply(simp add: rec_exec.simps mod_lemma power_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1972 |
apply(simp add: Maxr.simps loR.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1973 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1974 |
from h and this show "rec_exec (rec_maxr ?rR) [x, y, x] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1975 |
Maxr loR [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1976 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1977 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1978 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1979 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1980 |
lemma [simp]: "Max {ya. ya = 0 \<and> loR [0, y, ya]} = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1981 |
apply(rule_tac Max_eqI, auto simp: loR.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1982 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1983 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1984 |
lemma [simp]: "Suc 0 < y \<Longrightarrow> Suc (Suc 0) < y * y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1985 |
apply(induct y, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1986 |
apply(case_tac y, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1987 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1988 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1989 |
lemma less_mult: "\<lbrakk>x > 0; y > Suc 0\<rbrakk> \<Longrightarrow> x < y * x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1990 |
apply(case_tac y, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1991 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1992 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1993 |
lemma x_less_exp: "\<lbrakk>y > Suc 0\<rbrakk> \<Longrightarrow> x < y^x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1994 |
apply(induct x, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1995 |
apply(case_tac x, simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1996 |
apply(rule_tac y = "y* y^nat" in le_less_trans, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1997 |
apply(rule_tac less_mult, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1998 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
1999 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2000 |
lemma le_mult: "y \<noteq> (0::nat) \<Longrightarrow> x \<le> x * y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2001 |
by(induct y, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2002 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2003 |
lemma uplimit_loR: "\<lbrakk>Suc 0 < x; Suc 0 < y; loR [x, y, xa]\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2004 |
xa \<le> x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2005 |
apply(simp add: loR.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2006 |
apply(rule_tac classical, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2007 |
apply(subgoal_tac "xa < y^xa") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2008 |
apply(subgoal_tac "y^xa \<le> y^xa * q", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2009 |
apply(rule_tac le_mult, case_tac q, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2010 |
apply(rule_tac x_less_exp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2011 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2012 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2013 |
lemma [simp]: "\<lbrakk>xa \<le> x; loR [x, y, xa]; Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2014 |
{u. loR [x, y, u]} = {ya. ya \<le> x \<and> loR [x, y, ya]}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2015 |
apply(rule_tac Collect_cong, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2016 |
apply(erule_tac uplimit_loR, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2017 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2018 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2019 |
lemma Maxr_lo: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2020 |
Maxr loR [x, y] x = lo x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2021 |
apply(simp add: Maxr.simps lo.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2022 |
apply(erule_tac x = xa in allE, simp, simp add: uplimit_loR) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2023 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2024 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2025 |
lemma lo_lemma': "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2026 |
rec_exec rec_lo [x, y] = lo x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2027 |
by(simp add: Maxr_lo rec_lo_Maxr_lor) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2028 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2029 |
lemma lo_lemma'': "\<lbrakk>\<not> Suc 0 < x\<rbrakk> \<Longrightarrow> rec_exec rec_lo [x, y] = lo x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2030 |
apply(case_tac x, auto simp: rec_exec.simps rec_lo_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2031 |
Let_def lo.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2032 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2033 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2034 |
lemma lo_lemma''': "\<lbrakk>\<not> Suc 0 < y\<rbrakk> \<Longrightarrow> rec_exec rec_lo [x, y] = lo x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2035 |
apply(case_tac y, auto simp: rec_exec.simps rec_lo_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2036 |
Let_def lo.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2037 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2038 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2039 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2040 |
The correctness of @{text "rec_lo"}: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2041 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2042 |
lemma lo_lemma: "rec_exec rec_lo [x, y] = lo x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2043 |
apply(case_tac "Suc 0 < x \<and> Suc 0 < y") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2044 |
apply(auto simp: lo_lemma' lo_lemma'' lo_lemma''') |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2045 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2046 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2047 |
fun lgR :: "nat list \<Rightarrow> bool" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2048 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2049 |
"lgR [x, y, u] = (y^u \<le> x)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2050 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2051 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2052 |
@{text "lg"} specifies the @{text "lg"} function given on page 79 of |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2053 |
Boolos's book. It is one of the two notions of integeral logarithmatic |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2054 |
operation on that page. The other is @{text "lo"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2055 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2056 |
fun lg :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2057 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2058 |
"lg x y = (if x > 1 \<and> y > 1 \<and> {u. lgR [x, y, u]} \<noteq> {} then |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2059 |
Max {u. lgR [x, y, u]} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2060 |
else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2061 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2062 |
declare lg.simps[simp del] lgR.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2063 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2064 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2065 |
@{text "rec_lg"} is the recursive function used to implement @{text "lg"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2066 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2067 |
definition rec_lg :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2068 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2069 |
"rec_lg = (let rec_lgR = Cn 3 rec_le |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2070 |
[Cn 3 rec_power [id 3 1, id 3 2], id 3 0] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2071 |
let conR1 = Cn 2 rec_conj [Cn 2 rec_less |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2072 |
[Cn 2 (constn 1) [id 2 0], id 2 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2073 |
Cn 2 rec_less [Cn 2 (constn 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2074 |
[id 2 0], id 2 1]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2075 |
let conR2 = Cn 2 rec_not [conR1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2076 |
Cn 2 rec_add [Cn 2 rec_mult |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2077 |
[conR1, Cn 2 (rec_maxr rec_lgR) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2078 |
[id 2 0, id 2 1, id 2 0]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2079 |
Cn 2 rec_mult [conR2, Cn 2 (constn 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2080 |
[id 2 0]]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2081 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2082 |
lemma lg_maxr: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2083 |
rec_exec rec_lg [x, y] = Maxr lgR [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2084 |
proof(simp add: rec_exec.simps rec_lg_def Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2085 |
assume h: "Suc 0 < x" "Suc 0 < y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2086 |
let ?rR = "(Cn 3 rec_le [Cn 3 rec_power |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2087 |
[recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2088 |
have "rec_exec (rec_maxr ?rR) ([x, y] @ [x]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2089 |
= Maxr ((\<lambda> args. 0 < rec_exec ?rR args)) [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2090 |
proof(rule Maxr_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2091 |
show "primerec (Cn 3 rec_le [Cn 3 rec_power |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2092 |
[recf.id 3 (Suc 0), recf.id 3 2], recf.id 3 0]) (Suc (length [x, y]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2093 |
apply(auto simp: numeral_3_eq_3)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2094 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2095 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2096 |
moreover have "Maxr lgR [x, y] x = Maxr ((\<lambda> args. 0 < rec_exec ?rR args)) [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2097 |
apply(simp add: rec_exec.simps power_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2098 |
apply(simp add: Maxr.simps lgR.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2099 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2100 |
ultimately show "rec_exec (rec_maxr ?rR) [x, y, x] = Maxr lgR [x, y] x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2101 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2102 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2103 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2104 |
lemma [simp]: "\<lbrakk>Suc 0 < y; lgR [x, y, xa]\<rbrakk> \<Longrightarrow> xa \<le> x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2105 |
apply(simp add: lgR.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2106 |
apply(subgoal_tac "y^xa > xa", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2107 |
apply(erule x_less_exp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2108 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2109 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2110 |
lemma [simp]: "\<lbrakk>Suc 0 < x; Suc 0 < y; lgR [x, y, xa]\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2111 |
{u. lgR [x, y, u]} = {ya. ya \<le> x \<and> lgR [x, y, ya]}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2112 |
apply(rule_tac Collect_cong, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2113 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2114 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2115 |
lemma maxr_lg: "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> Maxr lgR [x, y] x = lg x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2116 |
apply(simp add: lg.simps Maxr.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2117 |
apply(erule_tac x = xa in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2118 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2119 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2120 |
lemma lg_lemma': "\<lbrakk>Suc 0 < x; Suc 0 < y\<rbrakk> \<Longrightarrow> rec_exec rec_lg [x, y] = lg x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2121 |
apply(simp add: maxr_lg lg_maxr) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2122 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2123 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2124 |
lemma lg_lemma'': "\<not> Suc 0 < x \<Longrightarrow> rec_exec rec_lg [x, y] = lg x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2125 |
apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2126 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2127 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2128 |
lemma lg_lemma''': "\<not> Suc 0 < y \<Longrightarrow> rec_exec rec_lg [x, y] = lg x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2129 |
apply(simp add: rec_exec.simps rec_lg_def Let_def lg.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2130 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2131 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2132 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2133 |
The correctness of @{text "rec_lg"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2134 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2135 |
lemma lg_lemma: "rec_exec rec_lg [x, y] = lg x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2136 |
apply(case_tac "Suc 0 < x \<and> Suc 0 < y", auto simp: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2137 |
lg_lemma' lg_lemma'' lg_lemma''') |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2138 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2139 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2140 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2141 |
@{text "Entry sr i"} returns the @{text "i"}-th entry of a list of natural |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2142 |
numbers encoded by number @{text "sr"} using Godel's coding. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2143 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2144 |
fun Entry :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2145 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2146 |
"Entry sr i = lo sr (Pi (Suc i))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2147 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2148 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2149 |
@{text "rec_entry"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2150 |
@{text "Entry"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2151 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2152 |
definition rec_entry:: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2153 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2154 |
"rec_entry = Cn 2 rec_lo [id 2 0, Cn 2 rec_pi [Cn 2 s [id 2 1]]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2155 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2156 |
declare Pi.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2157 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2158 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2159 |
The correctness of @{text "rec_entry"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2160 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2161 |
lemma entry_lemma: "rec_exec rec_entry [str, i] = Entry str i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2162 |
by(simp add: rec_entry_def rec_exec.simps lo_lemma pi_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2163 |
|
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
2164 |
|
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
2165 |
subsection {* The construction of F *} |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2166 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2167 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2168 |
Using the auxilliary functions obtained in last section, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2169 |
we are going to contruct the function @{text "F"}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2170 |
which is an interpreter of Turing Machines. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2171 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2172 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2173 |
fun listsum2 :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2174 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2175 |
"listsum2 xs 0 = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2176 |
| "listsum2 xs (Suc n) = listsum2 xs n + xs ! n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2177 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2178 |
fun rec_listsum2 :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2179 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2180 |
"rec_listsum2 vl 0 = Cn vl z [id vl 0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2181 |
| "rec_listsum2 vl (Suc n) = Cn vl rec_add |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2182 |
[rec_listsum2 vl n, id vl (n)]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2183 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2184 |
declare listsum2.simps[simp del] rec_listsum2.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2185 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2186 |
lemma listsum2_lemma: "\<lbrakk>length xs = vl; n \<le> vl\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2187 |
rec_exec (rec_listsum2 vl n) xs = listsum2 xs n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2188 |
apply(induct n, simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2189 |
apply(simp_all add: rec_exec.simps rec_listsum2.simps listsum2.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2190 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2191 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2192 |
fun strt' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2193 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2194 |
"strt' xs 0 = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2195 |
| "strt' xs (Suc n) = (let dbound = listsum2 xs n + n in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2196 |
strt' xs n + (2^(xs ! n + dbound) - 2^dbound))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2197 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2198 |
fun rec_strt' :: "nat \<Rightarrow> nat \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2199 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2200 |
"rec_strt' vl 0 = Cn vl z [id vl 0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2201 |
| "rec_strt' vl (Suc n) = (let rec_dbound = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2202 |
Cn vl rec_add [rec_listsum2 vl n, Cn vl (constn n) [id vl 0]] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2203 |
in Cn vl rec_add [rec_strt' vl n, Cn vl rec_minus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2204 |
[Cn vl rec_power [Cn vl (constn 2) [id vl 0], Cn vl rec_add |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2205 |
[id vl (n), rec_dbound]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2206 |
Cn vl rec_power [Cn vl (constn 2) [id vl 0], rec_dbound]]])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2207 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2208 |
declare strt'.simps[simp del] rec_strt'.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2209 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2210 |
lemma strt'_lemma: "\<lbrakk>length xs = vl; n \<le> vl\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2211 |
rec_exec (rec_strt' vl n) xs = strt' xs n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2212 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2213 |
apply(simp_all add: rec_exec.simps rec_strt'.simps strt'.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2214 |
Let_def power_lemma listsum2_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2215 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2216 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2217 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2218 |
@{text "strt"} corresponds to the @{text "strt"} function on page 90 of B book, but |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2219 |
this definition generalises the original one to deal with multiple input arguments. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2220 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2221 |
fun strt :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2222 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2223 |
"strt xs = (let ys = map Suc xs in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2224 |
strt' ys (length ys))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2225 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2226 |
fun rec_map :: "recf \<Rightarrow> nat \<Rightarrow> recf list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2227 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2228 |
"rec_map rf vl = map (\<lambda> i. Cn vl rf [id vl (i)]) [0..<vl]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2229 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2230 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2231 |
@{text "rec_strt"} is the recursive function used to implement @{text "strt"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2232 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2233 |
fun rec_strt :: "nat \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2234 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2235 |
"rec_strt vl = Cn vl (rec_strt' vl vl) (rec_map s vl)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2236 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2237 |
lemma map_s_lemma: "length xs = vl \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2238 |
map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn vl s [recf.id vl i])) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2239 |
[0..<vl] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2240 |
= map Suc xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2241 |
apply(induct vl arbitrary: xs, simp, auto simp: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2242 |
apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2243 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2244 |
fix ys y |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2245 |
assume ind: "\<And>xs. length xs = length (ys::nat list) \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2246 |
map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn (length ys) s |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2247 |
[recf.id (length ys) (i)])) [0..<length ys] = map Suc xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2248 |
show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2249 |
"map ((\<lambda>a. rec_exec a (ys @ [y])) \<circ> (\<lambda>i. Cn (Suc (length ys)) s |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2250 |
[recf.id (Suc (length ys)) (i)])) [0..<length ys] = map Suc ys" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2251 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2252 |
have "map ((\<lambda>a. rec_exec a ys) \<circ> (\<lambda>i. Cn (length ys) s |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2253 |
[recf.id (length ys) (i)])) [0..<length ys] = map Suc ys" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2254 |
apply(rule_tac ind, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2255 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2256 |
moreover have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2257 |
"map ((\<lambda>a. rec_exec a (ys @ [y])) \<circ> (\<lambda>i. Cn (Suc (length ys)) s |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2258 |
[recf.id (Suc (length ys)) (i)])) [0..<length ys] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2259 |
= map ((\<lambda>a. rec_exec a ys) \<circ> (\<lambda>i. Cn (length ys) s |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2260 |
[recf.id (length ys) (i)])) [0..<length ys]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2261 |
apply(rule_tac map_ext, auto simp: rec_exec.simps nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2262 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2263 |
ultimately show "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2264 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2265 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2266 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2267 |
fix vl xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2268 |
assume "length xs = Suc vl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2269 |
thus "\<exists>ys y. xs = ys @ [y]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2270 |
apply(rule_tac x = "butlast xs" in exI, rule_tac x = "last xs" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2271 |
apply(subgoal_tac "xs \<noteq> []", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2272 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2273 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2274 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2275 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2276 |
The correctness of @{text "rec_strt"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2277 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2278 |
lemma strt_lemma: "length xs = vl \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2279 |
rec_exec (rec_strt vl) xs = strt xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2280 |
apply(simp add: strt.simps rec_exec.simps strt'_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2281 |
apply(subgoal_tac "(map ((\<lambda>a. rec_exec a xs) \<circ> (\<lambda>i. Cn vl s [recf.id vl (i)])) [0..<vl]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2282 |
= map Suc xs", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2283 |
apply(rule map_s_lemma, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2284 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2285 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2286 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2287 |
The @{text "scan"} function on page 90 of B book. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2288 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2289 |
fun scan :: "nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2290 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2291 |
"scan r = r mod 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2292 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2293 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2294 |
@{text "rec_scan"} is the implemention of @{text "scan"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2295 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2296 |
definition rec_scan :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2297 |
where "rec_scan = Cn 1 rec_mod [id 1 0, constn 2]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2298 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2299 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2300 |
The correctness of @{text "scan"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2301 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2302 |
lemma scan_lemma: "rec_exec rec_scan [r] = r mod 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2303 |
by(simp add: rec_exec.simps rec_scan_def mod_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2304 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2305 |
fun newleft0 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2306 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2307 |
"newleft0 [p, r] = p" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2308 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2309 |
definition rec_newleft0 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2310 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2311 |
"rec_newleft0 = id 2 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2312 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2313 |
fun newrgt0 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2314 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2315 |
"newrgt0 [p, r] = r - scan r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2316 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2317 |
definition rec_newrgt0 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2318 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2319 |
"rec_newrgt0 = Cn 2 rec_minus [id 2 1, Cn 2 rec_scan [id 2 1]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2320 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2321 |
(*newleft1, newrgt1: left rgt number after execute on step*) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2322 |
fun newleft1 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2323 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2324 |
"newleft1 [p, r] = p" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2325 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2326 |
definition rec_newleft1 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2327 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2328 |
"rec_newleft1 = id 2 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2329 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2330 |
fun newrgt1 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2331 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2332 |
"newrgt1 [p, r] = r + 1 - scan r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2333 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2334 |
definition rec_newrgt1 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2335 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2336 |
"rec_newrgt1 = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2337 |
Cn 2 rec_minus [Cn 2 rec_add [id 2 1, Cn 2 (constn 1) [id 2 0]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2338 |
Cn 2 rec_scan [id 2 1]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2339 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2340 |
fun newleft2 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2341 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2342 |
"newleft2 [p, r] = p div 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2343 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2344 |
definition rec_newleft2 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2345 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2346 |
"rec_newleft2 = Cn 2 rec_quo [id 2 0, Cn 2 (constn 2) [id 2 0]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2347 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2348 |
fun newrgt2 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2349 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2350 |
"newrgt2 [p, r] = 2 * r + p mod 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2351 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2352 |
definition rec_newrgt2 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2353 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2354 |
"rec_newrgt2 = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2355 |
Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2356 |
Cn 2 rec_mod [id 2 0, Cn 2 (constn 2) [id 2 0]]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2357 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2358 |
fun newleft3 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2359 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2360 |
"newleft3 [p, r] = 2 * p + r mod 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2361 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2362 |
definition rec_newleft3 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2363 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2364 |
"rec_newleft3 = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2365 |
Cn 2 rec_add [Cn 2 rec_mult [Cn 2 (constn 2) [id 2 0], id 2 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2366 |
Cn 2 rec_mod [id 2 1, Cn 2 (constn 2) [id 2 0]]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2367 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2368 |
fun newrgt3 :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2369 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2370 |
"newrgt3 [p, r] = r div 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2371 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2372 |
definition rec_newrgt3 :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2373 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2374 |
"rec_newrgt3 = Cn 2 rec_quo [id 2 1, Cn 2 (constn 2) [id 2 0]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2375 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2376 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2377 |
The @{text "new_left"} function on page 91 of B book. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2378 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2379 |
fun newleft :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2380 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2381 |
"newleft p r a = (if a = 0 \<or> a = 1 then newleft0 [p, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2382 |
else if a = 2 then newleft2 [p, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2383 |
else if a = 3 then newleft3 [p, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2384 |
else p)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2385 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2386 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2387 |
@{text "rec_newleft"} is the recursive function used to |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2388 |
implement @{text "newleft"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2389 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2390 |
definition rec_newleft :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2391 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2392 |
"rec_newleft = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2393 |
(let g0 = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2394 |
Cn 3 rec_newleft0 [id 3 0, id 3 1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2395 |
let g1 = Cn 3 rec_newleft2 [id 3 0, id 3 1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2396 |
let g2 = Cn 3 rec_newleft3 [id 3 0, id 3 1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2397 |
let g3 = id 3 0 in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2398 |
let r0 = Cn 3 rec_disj |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2399 |
[Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2400 |
Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2401 |
let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2402 |
let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2403 |
let r3 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2404 |
let gs = [g0, g1, g2, g3] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2405 |
let rs = [r0, r1, r2, r3] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2406 |
rec_embranch (zip gs rs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2407 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2408 |
declare newleft.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2409 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2410 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2411 |
lemma Suc_Suc_Suc_Suc_induct: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2412 |
"\<lbrakk>i < Suc (Suc (Suc (Suc 0))); i = 0 \<Longrightarrow> P i; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2413 |
i = 1 \<Longrightarrow> P i; i =2 \<Longrightarrow> P i; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2414 |
i =3 \<Longrightarrow> P i\<rbrakk> \<Longrightarrow> P i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2415 |
apply(case_tac i, simp, case_tac nat, simp, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2416 |
case_tac nata, simp, case_tac natb, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2417 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2418 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2419 |
declare quo_lemma2[simp] mod_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2420 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2421 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2422 |
The correctness of @{text "rec_newleft"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2423 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2424 |
lemma newleft_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2425 |
"rec_exec rec_newleft [p, r, a] = newleft p r a" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2426 |
proof(simp only: rec_newleft_def Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2427 |
let ?rgs = "[Cn 3 rec_newleft0 [recf.id 3 0, recf.id 3 1], Cn 3 rec_newleft2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2429 |
let ?rrs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2430 |
"[Cn 3 rec_disj [Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2431 |
[recf.id 3 0]], Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 1) [recf.id 3 0]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2432 |
Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 2) [recf.id 3 0]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2433 |
Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2434 |
Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2435 |
thm embranch_lemma |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2436 |
have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2437 |
= Embranch (zip (map rec_exec ?rgs) (map (\<lambda>r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2438 |
apply(rule_tac embranch_lemma ) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2439 |
apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newleft0_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2440 |
rec_newleft1_def rec_newleft2_def rec_newleft3_def)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2441 |
apply(case_tac "a = 0 \<or> a = 1", rule_tac x = 0 in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2442 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2443 |
apply(case_tac "a = 2", rule_tac x = "Suc 0" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2444 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2445 |
apply(case_tac "a = 3", rule_tac x = "2" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2446 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2447 |
apply(case_tac "a > 3", rule_tac x = "3" in exI, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2448 |
apply(auto simp: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2449 |
apply(erule_tac [!] Suc_Suc_Suc_Suc_induct, auto simp: rec_exec.simps) |
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
2450 |
done |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2452 |
apply(simp add: Embranch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2453 |
apply(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2454 |
apply(auto simp: newleft.simps rec_newleft0_def rec_exec.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2455 |
rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2456 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2457 |
from k1 and k2 show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2458 |
"rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = newleft p r a" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2459 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2460 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2461 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2462 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2463 |
The @{text "newrght"} function is one similar to @{text "newleft"}, but used to |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2464 |
compute the right number. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2465 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2466 |
fun newrght :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2467 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2468 |
"newrght p r a = (if a = 0 then newrgt0 [p, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2469 |
else if a = 1 then newrgt1 [p, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2470 |
else if a = 2 then newrgt2 [p, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2471 |
else if a = 3 then newrgt3 [p, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2472 |
else r)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2473 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2474 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2475 |
@{text "rec_newrght"} is the recursive function used to implement |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2476 |
@{text "newrgth"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2477 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2478 |
definition rec_newrght :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2479 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2480 |
"rec_newrght = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2481 |
(let g0 = Cn 3 rec_newrgt0 [id 3 0, id 3 1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2482 |
let g1 = Cn 3 rec_newrgt1 [id 3 0, id 3 1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2483 |
let g2 = Cn 3 rec_newrgt2 [id 3 0, id 3 1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2484 |
let g3 = Cn 3 rec_newrgt3 [id 3 0, id 3 1] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2485 |
let g4 = id 3 1 in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2486 |
let r0 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 0) [id 3 0]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2487 |
let r1 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 1) [id 3 0]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2488 |
let r2 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 2) [id 3 0]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2489 |
let r3 = Cn 3 rec_eq [id 3 2, Cn 3 (constn 3) [id 3 0]] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2490 |
let r4 = Cn 3 rec_less [Cn 3 (constn 3) [id 3 0], id 3 2] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2491 |
let gs = [g0, g1, g2, g3, g4] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2492 |
let rs = [r0, r1, r2, r3, r4] in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2493 |
rec_embranch (zip gs rs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2494 |
declare newrght.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2495 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2496 |
lemma numeral_4_eq_4: "4 = Suc 3" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2497 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2498 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2499 |
lemma Suc_5_induct: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2500 |
"\<lbrakk>i < Suc (Suc (Suc (Suc (Suc 0)))); i = 0 \<Longrightarrow> P 0; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2502 |
apply(case_tac i, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2503 |
apply(case_tac nat, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2504 |
apply(case_tac nata, auto simp: numeral_2_eq_2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2505 |
apply(case_tac nat, auto simp: numeral_3_eq_3 numeral_4_eq_4) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2506 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2507 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2508 |
lemma [intro]: "primerec rec_scan (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2509 |
apply(auto simp: rec_scan_def, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2510 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2511 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2512 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2513 |
The correctness of @{text "rec_newrght"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2514 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2515 |
lemma newrght_lemma: "rec_exec rec_newrght [p, r, a] = newrght p r a" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2516 |
proof(simp only: rec_newrght_def Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2517 |
let ?gs' = "[newrgt0, newrgt1, newrgt2, newrgt3, \<lambda> zs. zs ! 1]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2518 |
let ?r0 = "\<lambda> zs. zs ! 2 = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2519 |
let ?r1 = "\<lambda> zs. zs ! 2 = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2520 |
let ?r2 = "\<lambda> zs. zs ! 2 = 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2521 |
let ?r3 = "\<lambda> zs. zs ! 2 = 3" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2522 |
let ?r4 = "\<lambda> zs. zs ! 2 > 3" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2523 |
let ?gs = "map (\<lambda> g. (\<lambda> zs. g [zs ! 0, zs ! 1])) ?gs'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2524 |
let ?rs = "[?r0, ?r1, ?r2, ?r3, ?r4]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2525 |
let ?rgs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2526 |
"[Cn 3 rec_newrgt0 [recf.id 3 0, recf.id 3 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2527 |
Cn 3 rec_newrgt1 [recf.id 3 0, recf.id 3 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2528 |
Cn 3 rec_newrgt2 [recf.id 3 0, recf.id 3 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2529 |
Cn 3 rec_newrgt3 [recf.id 3 0, recf.id 3 1], recf.id 3 1]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2530 |
let ?rrs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2533 |
Cn 3 rec_eq [recf.id 3 2, Cn 3 (constn 3) [recf.id 3 0]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2534 |
Cn 3 rec_less [Cn 3 (constn 3) [recf.id 3 0], recf.id 3 2]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2535 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2536 |
have k1: "rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2537 |
= Embranch (zip (map rec_exec ?rgs) (map (\<lambda>r args. 0 < rec_exec r args) ?rrs)) [p, r, a]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2538 |
apply(rule_tac embranch_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2539 |
apply(auto simp: numeral_3_eq_3 numeral_2_eq_2 rec_newrgt0_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2540 |
rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2541 |
apply(case_tac "a = 0", rule_tac x = 0 in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2542 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2543 |
apply(case_tac "a = 1", rule_tac x = "Suc 0" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2544 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2545 |
apply(case_tac "a = 2", rule_tac x = "2" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2546 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2547 |
apply(case_tac "a = 3", rule_tac x = "3" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2548 |
prefer 2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2549 |
apply(case_tac "a > 3", rule_tac x = "4" in exI, auto simp: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2550 |
apply(erule_tac [!] Suc_5_induct, auto simp: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2551 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2552 |
have k2: "Embranch (zip (map rec_exec ?rgs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2553 |
(map (\<lambda>r args. 0 < rec_exec r args) ?rrs)) [p, r, a] = newrght p r a" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2554 |
apply(auto simp:Embranch.simps rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2555 |
apply(auto simp: newrght.simps rec_newrgt3_def rec_newrgt2_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2556 |
rec_newrgt1_def rec_newrgt0_def rec_exec.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2557 |
scan_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2558 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2559 |
from k1 and k2 show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2560 |
"rec_exec (rec_embranch (zip ?rgs ?rrs)) [p, r, a] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2561 |
newrght p r a" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2562 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2563 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2564 |
declare Entry.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2565 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2566 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2567 |
The @{text "actn"} function given on page 92 of B book, which is used to |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2568 |
fetch Turing Machine intructions. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2569 |
In @{text "actn m q r"}, @{text "m"} is the Godel coding of a Turing Machine, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2570 |
@{text "q"} is the current state of Turing Machine, @{text "r"} is the |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2571 |
right number of Turing Machine tape. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2572 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2573 |
fun actn :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2574 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2575 |
"actn m q r = (if q \<noteq> 0 then Entry m (4*(q - 1) + 2 * scan r) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2576 |
else 4)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2577 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2578 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2579 |
@{text "rec_actn"} is the recursive function used to implement @{text "actn"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2580 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2581 |
definition rec_actn :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2582 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2583 |
"rec_actn = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2584 |
Cn 3 rec_add [Cn 3 rec_mult |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2585 |
[Cn 3 rec_entry [id 3 0, Cn 3 rec_add [Cn 3 rec_mult |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2586 |
[Cn 3 (constn 4) [id 3 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2587 |
Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2588 |
Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2589 |
Cn 3 rec_scan [id 3 2]]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2590 |
Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2591 |
Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2592 |
Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2593 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2594 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2595 |
The correctness of @{text "actn"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2596 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2597 |
lemma actn_lemma: "rec_exec rec_actn [m, q, r] = actn m q r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2598 |
by(auto simp: rec_actn_def rec_exec.simps entry_lemma scan_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2599 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2600 |
fun newstat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2601 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2602 |
"newstat m q r = (if q \<noteq> 0 then Entry m (4*(q - 1) + 2*scan r + 1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2603 |
else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2604 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2605 |
definition rec_newstat :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2606 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2607 |
"rec_newstat = Cn 3 rec_add |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2608 |
[Cn 3 rec_mult [Cn 3 rec_entry [id 3 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2609 |
Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 4) [id 3 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2610 |
Cn 3 rec_minus [id 3 1, Cn 3 (constn 1) [id 3 0]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2611 |
Cn 3 rec_add [Cn 3 rec_mult [Cn 3 (constn 2) [id 3 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2612 |
Cn 3 rec_scan [id 3 2]], Cn 3 (constn 1) [id 3 0]]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2613 |
Cn 3 rec_noteq [id 3 1, Cn 3 (constn 0) [id 3 0]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2614 |
Cn 3 rec_mult [Cn 3 (constn 0) [id 3 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2615 |
Cn 3 rec_eq [id 3 1, Cn 3 (constn 0) [id 3 0]]]] " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2616 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2617 |
lemma newstat_lemma: "rec_exec rec_newstat [m, q, r] = newstat m q r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2618 |
by(auto simp: rec_exec.simps entry_lemma scan_lemma rec_newstat_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2619 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2620 |
declare newstat.simps[simp del] actn.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2621 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2622 |
text{*code the configuration*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2623 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2624 |
fun trpl :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2625 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2626 |
"trpl p q r = (Pi 0)^p * (Pi 1)^q * (Pi 2)^r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2627 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2628 |
definition rec_trpl :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2629 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2630 |
"rec_trpl = Cn 3 rec_mult [Cn 3 rec_mult |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2631 |
[Cn 3 rec_power [Cn 3 (constn (Pi 0)) [id 3 0], id 3 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2632 |
Cn 3 rec_power [Cn 3 (constn (Pi 1)) [id 3 0], id 3 1]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2633 |
Cn 3 rec_power [Cn 3 (constn (Pi 2)) [id 3 0], id 3 2]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2634 |
declare trpl.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2635 |
lemma trpl_lemma: "rec_exec rec_trpl [p, q, r] = trpl p q r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2636 |
by(auto simp: rec_trpl_def rec_exec.simps power_lemma trpl.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2637 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2638 |
text{*left, stat, rght: decode func*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2639 |
fun left :: "nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2640 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2641 |
"left c = lo c (Pi 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2642 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2643 |
fun stat :: "nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2644 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2645 |
"stat c = lo c (Pi 1)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2646 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2647 |
fun rght :: "nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2648 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2649 |
"rght c = lo c (Pi 2)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2650 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2651 |
thm Prime.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2652 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2653 |
fun inpt :: "nat \<Rightarrow> nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2654 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2655 |
"inpt m xs = trpl 0 1 (strt xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2656 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2657 |
fun newconf :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2658 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2659 |
"newconf m c = trpl (newleft (left c) (rght c) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2660 |
(actn m (stat c) (rght c))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2661 |
(newstat m (stat c) (rght c)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2662 |
(newrght (left c) (rght c) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2663 |
(actn m (stat c) (rght c)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2664 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2665 |
declare left.simps[simp del] stat.simps[simp del] rght.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2666 |
inpt.simps[simp del] newconf.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2667 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2668 |
definition rec_left :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2669 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2670 |
"rec_left = Cn 1 rec_lo [id 1 0, constn (Pi 0)]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2671 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2672 |
definition rec_right :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2673 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2674 |
"rec_right = Cn 1 rec_lo [id 1 0, constn (Pi 2)]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2675 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2676 |
definition rec_stat :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2677 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2678 |
"rec_stat = Cn 1 rec_lo [id 1 0, constn (Pi 1)]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2679 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2680 |
definition rec_inpt :: "nat \<Rightarrow> recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2681 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2682 |
"rec_inpt vl = Cn vl rec_trpl |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2683 |
[Cn vl (constn 0) [id vl 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2684 |
Cn vl (constn 1) [id vl 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2685 |
Cn vl (rec_strt (vl - 1)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2686 |
(map (\<lambda> i. id vl (i)) [1..<vl])]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2687 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2688 |
lemma left_lemma: "rec_exec rec_left [c] = left c" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2689 |
by(simp add: rec_exec.simps rec_left_def left.simps lo_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2690 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2691 |
lemma right_lemma: "rec_exec rec_right [c] = rght c" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2692 |
by(simp add: rec_exec.simps rec_right_def rght.simps lo_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2693 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2694 |
lemma stat_lemma: "rec_exec rec_stat [c] = stat c" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2695 |
by(simp add: rec_exec.simps rec_stat_def stat.simps lo_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2696 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2697 |
declare rec_strt.simps[simp del] strt.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2698 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2699 |
lemma map_cons_eq: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2700 |
"(map ((\<lambda>a. rec_exec a (m # xs)) \<circ> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2701 |
(\<lambda>i. recf.id (Suc (length xs)) (i))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2702 |
[Suc 0..<Suc (length xs)]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2703 |
= map (\<lambda> i. xs ! (i - 1)) [Suc 0..<Suc (length xs)]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2704 |
apply(rule map_ext, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2705 |
apply(auto simp: rec_exec.simps nth_append nth_Cons split: nat.split) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2706 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2707 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2708 |
lemma list_map_eq: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2709 |
"vl = length (xs::nat list) \<Longrightarrow> map (\<lambda> i. xs ! (i - 1)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2710 |
[Suc 0..<Suc vl] = xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2711 |
apply(induct vl arbitrary: xs, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2712 |
apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2713 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2714 |
fix ys y |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2715 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2716 |
"\<And>xs. length (ys::nat list) = length (xs::nat list) \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2717 |
map (\<lambda>i. xs ! (i - Suc 0)) [Suc 0..<length xs] @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2718 |
[xs ! (length xs - Suc 0)] = xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2719 |
and h: "Suc 0 \<le> length (ys::nat list)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2720 |
have "map (\<lambda>i. ys ! (i - Suc 0)) [Suc 0..<length ys] @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2721 |
[ys ! (length ys - Suc 0)] = ys" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2722 |
apply(rule_tac ind, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2723 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2724 |
moreover have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2725 |
"map (\<lambda>i. (ys @ [y]) ! (i - Suc 0)) [Suc 0..<length ys] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2726 |
= map (\<lambda>i. ys ! (i - Suc 0)) [Suc 0..<length ys]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2727 |
apply(rule map_ext) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2728 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2729 |
apply(auto simp: nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2730 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2731 |
ultimately show "map (\<lambda>i. (ys @ [y]) ! (i - Suc 0)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2732 |
[Suc 0..<length ys] @ [(ys @ [y]) ! (length ys - Suc 0)] = ys" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2733 |
apply(simp del: map_eq_conv add: nth_append, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2734 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2735 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2736 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2737 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2738 |
fix vl xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2739 |
assume "Suc vl = length (xs::nat list)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2740 |
thus "\<exists>ys y. xs = ys @ [y]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2741 |
apply(rule_tac x = "butlast xs" in exI, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2742 |
rule_tac x = "last xs" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2743 |
apply(case_tac "xs \<noteq> []", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2744 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2745 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2746 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2747 |
lemma [elim]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2748 |
"Suc 0 \<le> length xs \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2749 |
(map ((\<lambda>a. rec_exec a (m # xs)) \<circ> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2750 |
(\<lambda>i. recf.id (Suc (length xs)) (i))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2751 |
[Suc 0..<length xs] @ [(m # xs) ! length xs]) = xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2752 |
using map_cons_eq[of m xs] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2753 |
apply(simp del: map_eq_conv add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2754 |
using list_map_eq[of "length xs" xs] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2755 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2756 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2757 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2758 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2759 |
lemma inpt_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2760 |
"\<lbrakk>Suc (length xs) = vl\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2761 |
rec_exec (rec_inpt vl) (m # xs) = inpt m xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2762 |
apply(auto simp: rec_exec.simps rec_inpt_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2763 |
trpl_lemma inpt.simps strt_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2764 |
apply(subgoal_tac |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2765 |
"(map ((\<lambda>a. rec_exec a (m # xs)) \<circ> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2766 |
(\<lambda>i. recf.id (Suc (length xs)) (i))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2767 |
[Suc 0..<length xs] @ [(m # xs) ! length xs]) = xs", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2768 |
apply(auto, case_tac xs, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2769 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2770 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2771 |
definition rec_newconf:: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2772 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2773 |
"rec_newconf = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2774 |
Cn 2 rec_trpl |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2775 |
[Cn 2 rec_newleft [Cn 2 rec_left [id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2776 |
Cn 2 rec_right [id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2777 |
Cn 2 rec_actn [id 2 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2778 |
Cn 2 rec_stat [id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2779 |
Cn 2 rec_right [id 2 1]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2780 |
Cn 2 rec_newstat [id 2 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2781 |
Cn 2 rec_stat [id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2782 |
Cn 2 rec_right [id 2 1]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2783 |
Cn 2 rec_newrght [Cn 2 rec_left [id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2784 |
Cn 2 rec_right [id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2785 |
Cn 2 rec_actn [id 2 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2786 |
Cn 2 rec_stat [id 2 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2787 |
Cn 2 rec_right [id 2 1]]]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2788 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2789 |
lemma newconf_lemma: "rec_exec rec_newconf [m ,c] = newconf m c" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2790 |
by(auto simp: rec_newconf_def rec_exec.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2791 |
trpl_lemma newleft_lemma left_lemma |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2792 |
right_lemma stat_lemma newrght_lemma actn_lemma |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2793 |
newstat_lemma stat_lemma newconf.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2794 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2795 |
declare newconf_lemma[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2796 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2797 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2798 |
@{text "conf m r k"} computes the TM configuration after @{text "k"} steps of execution |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2799 |
of TM coded as @{text "m"} starting from the initial configuration where the left number equals @{text "0"}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2800 |
right number equals @{text "r"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2801 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2802 |
fun conf :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2803 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2804 |
"conf m r 0 = trpl 0 (Suc 0) r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2805 |
| "conf m r (Suc t) = newconf m (conf m r t)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2806 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2807 |
declare conf.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2808 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2809 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2810 |
@{text "conf"} is implemented by the following recursive function @{text "rec_conf"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2811 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2812 |
definition rec_conf :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2813 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2815 |
(Cn 4 rec_newconf [id 4 0, id 4 3])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2816 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2817 |
lemma conf_step: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2818 |
"rec_exec rec_conf [m, r, Suc t] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2819 |
rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2820 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2821 |
have "rec_exec rec_conf ([m, r] @ [Suc t]) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2822 |
rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2823 |
by(simp only: rec_conf_def rec_pr_Suc_simp_rewrite, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2824 |
simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2825 |
thus "rec_exec rec_conf [m, r, Suc t] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2826 |
rec_exec rec_newconf [m, rec_exec rec_conf [m, r, t]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2827 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2828 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2829 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2830 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2831 |
The correctness of @{text "rec_conf"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2832 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2833 |
lemma conf_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2834 |
"rec_exec rec_conf [m, r, t] = conf m r t" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2835 |
apply(induct t) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2836 |
apply(simp add: rec_conf_def rec_exec.simps conf.simps inpt_lemma trpl_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2837 |
apply(simp add: conf_step conf.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2838 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2839 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2840 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2841 |
@{text "NSTD c"} returns true if the configureation coded by @{text "c"} is no a stardard |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2842 |
final configuration. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2843 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2844 |
fun NSTD :: "nat \<Rightarrow> bool" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2845 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2846 |
"NSTD c = (stat c \<noteq> 0 \<or> left c \<noteq> 0 \<or> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2847 |
rght c \<noteq> 2^(lg (rght c + 1) 2) - 1 \<or> rght c = 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2848 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2849 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2850 |
@{text "rec_NSTD"} is the recursive function implementing @{text "NSTD"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2851 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2852 |
definition rec_NSTD :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2853 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2854 |
"rec_NSTD = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2855 |
Cn 1 rec_disj [ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2856 |
Cn 1 rec_disj [ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2857 |
Cn 1 rec_disj |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2858 |
[Cn 1 rec_noteq [rec_stat, constn 0], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2859 |
Cn 1 rec_noteq [rec_left, constn 0]] , |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2860 |
Cn 1 rec_noteq [rec_right, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2861 |
Cn 1 rec_minus [Cn 1 rec_power |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2862 |
[constn 2, Cn 1 rec_lg |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2863 |
[Cn 1 rec_add |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2864 |
[rec_right, constn 1], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2865 |
constn 2]], constn 1]]], |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2866 |
Cn 1 rec_eq [rec_right, constn 0]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2867 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2868 |
lemma NSTD_lemma1: "rec_exec rec_NSTD [c] = Suc 0 \<or> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2869 |
rec_exec rec_NSTD [c] = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2870 |
by(simp add: rec_exec.simps rec_NSTD_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2871 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2872 |
declare NSTD.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2873 |
lemma NSTD_lemma2': "(rec_exec rec_NSTD [c] = Suc 0) \<Longrightarrow> NSTD c" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2874 |
apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma left_lemma |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2875 |
lg_lemma right_lemma power_lemma NSTD.simps eq_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2876 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2877 |
apply(case_tac "0 < left c", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2878 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2879 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2880 |
lemma NSTD_lemma2'': |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2881 |
"NSTD c \<Longrightarrow> (rec_exec rec_NSTD [c] = Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2882 |
apply(simp add: rec_exec.simps rec_NSTD_def stat_lemma |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2883 |
left_lemma lg_lemma right_lemma power_lemma NSTD.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2884 |
apply(auto split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2885 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2886 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2887 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2888 |
The correctness of @{text "NSTD"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2889 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2890 |
lemma NSTD_lemma2: "(rec_exec rec_NSTD [c] = Suc 0) = NSTD c" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2891 |
using NSTD_lemma1 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2892 |
apply(auto intro: NSTD_lemma2' NSTD_lemma2'') |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2893 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2894 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2895 |
fun nstd :: "nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2896 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2897 |
"nstd c = (if NSTD c then 1 else 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2898 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2899 |
lemma nstd_lemma: "rec_exec rec_NSTD [c] = nstd c" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2900 |
using NSTD_lemma1 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2901 |
apply(simp add: NSTD_lemma2, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2902 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2903 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2904 |
text{* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2905 |
@{text "nonstep m r t"} means afer @{text "t"} steps of execution, the TM coded by @{text "m"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2906 |
is not at a stardard final configuration. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2907 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2908 |
fun nonstop :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2909 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2910 |
"nonstop m r t = nstd (conf m r t)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2911 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2912 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2913 |
@{text "rec_nonstop"} is the recursive function implementing @{text "nonstop"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2914 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2915 |
definition rec_nonstop :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2916 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2917 |
"rec_nonstop = Cn 3 rec_NSTD [rec_conf]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2918 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2919 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2920 |
The correctness of @{text "rec_nonstop"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2921 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2922 |
lemma nonstop_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2923 |
"rec_exec rec_nonstop [m, r, t] = nonstop m r t" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2924 |
apply(simp add: rec_exec.simps rec_nonstop_def nstd_lemma conf_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2925 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2926 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2927 |
text{* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2928 |
@{text "rec_halt"} is the recursive function calculating the steps a TM needs to execute before |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2929 |
to reach a stardard final configuration. This recursive function is the only one |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2930 |
using @{text "Mn"} combinator. So it is the only non-primitive recursive function |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2931 |
needs to be used in the construction of the universal function @{text "F"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2932 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2933 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2934 |
definition rec_halt :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2935 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2936 |
"rec_halt = Mn (Suc (Suc 0)) (rec_nonstop)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2937 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2938 |
declare nonstop.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2939 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2940 |
lemma primerec_not0: "primerec f n \<Longrightarrow> n > 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2941 |
by(induct f n rule: primerec.induct, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2942 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2943 |
lemma [elim]: "primerec f 0 \<Longrightarrow> RR" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2944 |
apply(drule_tac primerec_not0, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2945 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2946 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2947 |
lemma [simp]: "length xs = Suc n \<Longrightarrow> length (butlast xs) = n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2948 |
apply(subgoal_tac "\<exists> y ys. xs = ys @ [y]", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2949 |
apply(rule_tac x = "last xs" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2950 |
apply(rule_tac x = "butlast xs" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2951 |
apply(case_tac "xs = []", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2952 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2953 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2954 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2955 |
The lemma relates the interpreter of primitive fucntions with |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2956 |
the calculation relation of general recursive functions. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2957 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2958 |
lemma prime_rel_exec_eq: "primerec r (length xs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2959 |
\<Longrightarrow> rec_calc_rel r xs rs = (rec_exec r xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2960 |
proof(induct r xs arbitrary: rs rule: rec_exec.induct, simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2961 |
fix xs rs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2962 |
assume "primerec z (length (xs::nat list))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2963 |
hence "length xs = Suc 0" by(erule_tac prime_z_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2964 |
thus "rec_calc_rel z xs rs = (rec_exec z xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2965 |
apply(case_tac xs, simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2966 |
apply(erule_tac calc_z_reverse, simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2967 |
apply(simp add: rec_exec.simps, rule_tac calc_z) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2968 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2969 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2970 |
fix xs rs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2971 |
assume "primerec s (length (xs::nat list))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2972 |
hence "length xs = Suc 0" .. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2973 |
thus "rec_calc_rel s xs rs = (rec_exec s xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2974 |
by(case_tac xs, auto simp: rec_exec.simps intro: calc_s |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2975 |
elim: calc_s_reverse) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2976 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2977 |
fix m n xs rs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2978 |
assume "primerec (recf.id m n) (length (xs::nat list))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2979 |
thus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2980 |
"rec_calc_rel (recf.id m n) xs rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2981 |
(rec_exec (recf.id m n) xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2982 |
apply(erule_tac prime_id_reverse) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2983 |
apply(simp add: rec_exec.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2984 |
apply(erule_tac calc_id_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2985 |
apply(rule_tac calc_id, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2986 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2987 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2988 |
fix n f gs xs rs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2989 |
assume ind1: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2990 |
"\<And>x rs. \<lbrakk>x \<in> set gs; primerec x (length xs)\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2991 |
rec_calc_rel x xs rs = (rec_exec x xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2992 |
and ind2: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2993 |
"\<And>x rs. \<lbrakk>x = map (\<lambda>a. rec_exec a xs) gs; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2994 |
primerec f (length gs)\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2995 |
rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2996 |
(rec_exec f (map (\<lambda>a. rec_exec a xs) gs) = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2997 |
and h: "primerec (Cn n f gs) (length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2998 |
show "rec_calc_rel (Cn n f gs) xs rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
2999 |
(rec_exec (Cn n f gs) xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3000 |
proof(auto simp: rec_exec.simps, erule_tac calc_cn_reverse, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3001 |
fix ys |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3002 |
assume g1:"\<forall>k<length gs. rec_calc_rel (gs ! k) xs (ys ! k)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3003 |
and g2: "length ys = length gs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3004 |
and g3: "rec_calc_rel f ys rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3005 |
have "rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3006 |
(rec_exec f (map (\<lambda>a. rec_exec a xs) gs) = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3007 |
apply(rule_tac ind2, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3008 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3009 |
apply(erule_tac prime_cn_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3010 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3011 |
moreover have "ys = (map (\<lambda>a. rec_exec a xs) gs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3012 |
proof(rule_tac nth_equalityI, auto simp: g2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3013 |
fix i |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3014 |
assume "i < length gs" thus "ys ! i = rec_exec (gs!i) xs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3015 |
using ind1[of "gs ! i" "ys ! i"] g1 h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3016 |
apply(erule_tac prime_cn_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3017 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3018 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3019 |
ultimately show "rec_exec f (map (\<lambda>a. rec_exec a xs) gs) = rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3020 |
using g3 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3021 |
by(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3022 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3023 |
from h show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3024 |
"rec_calc_rel (Cn n f gs) xs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3025 |
(rec_exec f (map (\<lambda>a. rec_exec a xs) gs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3026 |
apply(rule_tac rs = "(map (\<lambda>a. rec_exec a xs) gs)" in calc_cn, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3027 |
auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3028 |
apply(erule_tac [!] prime_cn_reverse, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3029 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3030 |
fix k |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3031 |
assume "k < length gs" "primerec f (length gs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3032 |
"\<forall>i<length gs. primerec (gs ! i) (length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3033 |
thus "rec_calc_rel (gs ! k) xs (rec_exec (gs ! k) xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3034 |
using ind1[of "gs!k" "(rec_exec (gs ! k) xs)"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3035 |
by(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3036 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3037 |
assume "primerec f (length gs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3038 |
"\<forall>i<length gs. primerec (gs ! i) (length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3039 |
thus "rec_calc_rel f (map (\<lambda>a. rec_exec a xs) gs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3040 |
(rec_exec f (map (\<lambda>a. rec_exec a xs) gs))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3041 |
using ind2[of "(map (\<lambda>a. rec_exec a xs) gs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3042 |
"(rec_exec f (map (\<lambda>a. rec_exec a xs) gs))"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3043 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3044 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3045 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3046 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3047 |
fix n f g xs rs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3048 |
assume ind1: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3049 |
"\<And>rs. \<lbrakk>last xs = 0; primerec f (length xs - Suc 0)\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3050 |
\<Longrightarrow> rec_calc_rel f (butlast xs) rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3051 |
(rec_exec f (butlast xs) = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3052 |
and ind2 : |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3053 |
"\<And>rs. \<lbrakk>0 < last xs; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3054 |
primerec (Pr n f g) (Suc (length xs - Suc 0))\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3055 |
rec_calc_rel (Pr n f g) (butlast xs @ [last xs - Suc 0]) rs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3056 |
= (rec_exec (Pr n f g) (butlast xs @ [last xs - Suc 0]) = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3057 |
and ind3: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3058 |
"\<And>rs. \<lbrakk>0 < last xs; primerec g (Suc (Suc (length xs - Suc 0)))\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3059 |
\<Longrightarrow> rec_calc_rel g (butlast xs @ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3060 |
[last xs - Suc 0, rec_exec (Pr n f g) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3061 |
(butlast xs @ [last xs - Suc 0])]) rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3062 |
(rec_exec g (butlast xs @ [last xs - Suc 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3063 |
rec_exec (Pr n f g) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3064 |
(butlast xs @ [last xs - Suc 0])]) = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3065 |
and h: "primerec (Pr n f g) (length (xs::nat list))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3066 |
show "rec_calc_rel (Pr n f g) xs rs = (rec_exec (Pr n f g) xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3067 |
proof(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3068 |
assume "rec_calc_rel (Pr n f g) xs rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3069 |
thus "rec_exec (Pr n f g) xs = rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3070 |
proof(erule_tac calc_pr_reverse) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3071 |
fix l |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3072 |
assume g: "xs = l @ [0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3073 |
"rec_calc_rel f l rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3074 |
"n = length l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3075 |
thus "rec_exec (Pr n f g) xs = rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3076 |
using ind1[of rs] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3077 |
apply(simp add: rec_exec.simps, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3078 |
erule_tac prime_pr_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3079 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3080 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3081 |
fix l y ry |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3082 |
assume d:"xs = l @ [Suc y]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3083 |
"rec_calc_rel (Pr (length l) f g) (l @ [y]) ry" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3084 |
"n = length l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3085 |
"rec_calc_rel g (l @ [y, ry]) rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3086 |
moreover hence "primerec g (Suc (Suc n))" using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3087 |
proof(erule_tac prime_pr_reverse) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3088 |
assume "primerec g (Suc (Suc n))" "length xs = Suc n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3089 |
thus "?thesis" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3090 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3091 |
ultimately show "rec_exec (Pr n f g) xs = rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3092 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3093 |
using ind3[of rs] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3094 |
apply(simp add: rec_pr_Suc_simp_rewrite) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3095 |
using ind2[of ry] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3096 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3097 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3098 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3099 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3100 |
show "rec_calc_rel (Pr n f g) xs (rec_exec (Pr n f g) xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3101 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3102 |
have "rec_calc_rel (Pr n f g) (butlast xs @ [last xs]) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3103 |
(rec_exec (Pr n f g) (butlast xs @ [last xs]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3104 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3105 |
apply(erule_tac prime_pr_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3106 |
apply(case_tac "last xs", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3107 |
apply(rule_tac calc_pr_zero, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3108 |
using ind1[of "rec_exec (Pr n f g) (butlast xs @ [0])"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3109 |
apply(simp add: rec_exec.simps, simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3110 |
thm calc_pr_ind |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3111 |
apply(rule_tac rk = "rec_exec (Pr n f g) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3112 |
(butlast xs@[last xs - Suc 0])" in calc_pr_ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3113 |
using ind2[of "rec_exec (Pr n f g) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3114 |
(butlast xs @ [last xs - Suc 0])"] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3115 |
apply(simp, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3116 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3117 |
fix nat |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3118 |
assume "length xs = Suc n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3119 |
"primerec g (Suc (Suc n))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3120 |
"last xs = Suc nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3121 |
thus |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3122 |
"rec_calc_rel g (butlast xs @ [nat, rec_exec (Pr n f g) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3123 |
(butlast xs @ [nat])]) (rec_exec (Pr n f g) (butlast xs @ [Suc nat]))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3124 |
using ind3[of "rec_exec (Pr n f g) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3125 |
(butlast xs @ [Suc nat])"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3126 |
apply(simp add: rec_exec.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3127 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3128 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3129 |
thus "rec_calc_rel (Pr n f g) xs (rec_exec (Pr n f g) xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3130 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3131 |
apply(erule_tac prime_pr_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3132 |
apply(subgoal_tac "butlast xs @ [last xs] = xs", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3133 |
apply(case_tac xs, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3134 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3135 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3136 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3137 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3138 |
fix n f xs rs |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3139 |
assume "primerec (Mn n f) (length (xs::nat list))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3140 |
thus "rec_calc_rel (Mn n f) xs rs = (rec_exec (Mn n f) xs = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3141 |
by(erule_tac prime_mn_reverse) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3142 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3143 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3144 |
declare numeral_2_eq_2[simp] numeral_3_eq_3[simp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3145 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3146 |
lemma [intro]: "primerec rec_right (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3147 |
apply(simp add: rec_right_def rec_lo_def Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3148 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3149 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3150 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3151 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3152 |
lemma [simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3153 |
"rec_calc_rel rec_right [r] rs = (rec_exec rec_right [r] = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3154 |
apply(rule_tac prime_rel_exec_eq, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3155 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3156 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3157 |
lemma [intro]: "primerec rec_pi (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3158 |
apply(simp add: rec_pi_def rec_dummy_pi_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3159 |
rec_np_def rec_fac_def rec_prime_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3160 |
rec_Minr.simps Let_def get_fstn_args.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3161 |
arity.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3162 |
rec_all.simps rec_sigma.simps rec_accum.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3163 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3164 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3165 |
apply(simp add: rec_dummyfac_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3166 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3167 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3168 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3169 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3170 |
lemma [intro]: "primerec rec_trpl (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3171 |
apply(simp add: rec_trpl_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3172 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3173 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3174 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3175 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3176 |
lemma [intro!]: "\<lbrakk>0 < vl; n \<le> vl\<rbrakk> \<Longrightarrow> primerec (rec_listsum2 vl n) vl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3177 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3178 |
apply(simp_all add: rec_strt'.simps Let_def rec_listsum2.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3179 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3180 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3181 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3182 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3183 |
lemma [elim]: "\<lbrakk>0 < vl; n \<le> vl\<rbrakk> \<Longrightarrow> primerec (rec_strt' vl n) vl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3184 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3185 |
apply(simp_all add: rec_strt'.simps Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3186 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3187 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3188 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3189 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3190 |
lemma [elim]: "vl > 0 \<Longrightarrow> primerec (rec_strt vl) vl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3191 |
apply(simp add: rec_strt.simps rec_strt'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3192 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3193 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3194 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3195 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3196 |
lemma [elim]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3197 |
"i < vl \<Longrightarrow> primerec ((map (\<lambda>i. recf.id (Suc vl) (i)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3198 |
[Suc 0..<vl] @ [recf.id (Suc vl) (vl)]) ! i) (Suc vl)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3199 |
apply(induct i, auto simp: nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3200 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3201 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3202 |
lemma [intro]: "primerec rec_newleft0 ((Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3203 |
apply(simp add: rec_newleft_def rec_embranch.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3204 |
Let_def arity.simps rec_newleft0_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3205 |
rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3206 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3207 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3208 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3209 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3210 |
lemma [intro]: "primerec rec_newleft1 ((Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3211 |
apply(simp add: rec_newleft_def rec_embranch.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3212 |
Let_def arity.simps rec_newleft0_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3213 |
rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3214 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3215 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3216 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3217 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3218 |
lemma [intro]: "primerec rec_newleft2 ((Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3219 |
apply(simp add: rec_newleft_def rec_embranch.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3220 |
Let_def arity.simps rec_newleft0_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3221 |
rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3222 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3223 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3224 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3225 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3226 |
lemma [intro]: "primerec rec_newleft3 ((Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3227 |
apply(simp add: rec_newleft_def rec_embranch.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3228 |
Let_def arity.simps rec_newleft0_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3229 |
rec_newleft1_def rec_newleft2_def rec_newleft3_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3230 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3231 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3232 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3233 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3234 |
lemma [intro]: "primerec rec_newleft (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3235 |
apply(simp add: rec_newleft_def rec_embranch.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3236 |
Let_def arity.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3237 |
apply(rule_tac prime_cn, auto+) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3238 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3239 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3240 |
lemma [intro]: "primerec rec_left (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3241 |
apply(simp add: rec_left_def rec_lo_def rec_entry_def Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3242 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3243 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3244 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3245 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3246 |
lemma [intro]: "primerec rec_actn (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3247 |
apply(simp add: rec_left_def rec_lo_def rec_entry_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3248 |
Let_def rec_actn_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3249 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3250 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3251 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3252 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3253 |
lemma [intro]: "primerec rec_stat (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3254 |
apply(simp add: rec_left_def rec_lo_def rec_entry_def Let_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3255 |
rec_actn_def rec_stat_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3256 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3257 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3258 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3259 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3260 |
lemma [intro]: "primerec rec_newstat (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3261 |
apply(simp add: rec_left_def rec_lo_def rec_entry_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3262 |
Let_def rec_actn_def rec_stat_def rec_newstat_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3263 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3264 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3265 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3266 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3267 |
lemma [intro]: "primerec rec_newrght (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3268 |
apply(simp add: rec_newrght_def rec_embranch.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3269 |
Let_def arity.simps rec_newrgt0_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3270 |
rec_newrgt1_def rec_newrgt2_def rec_newrgt3_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3271 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3272 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3273 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3274 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3275 |
lemma [intro]: "primerec rec_newconf (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3276 |
apply(simp add: rec_newconf_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3277 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3278 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3279 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3280 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3281 |
lemma [intro]: "0 < vl \<Longrightarrow> primerec (rec_inpt (Suc vl)) (Suc vl)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3282 |
apply(simp add: rec_inpt_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3283 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3284 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3285 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3286 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3287 |
lemma [intro]: "primerec rec_conf (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3288 |
apply(simp add: rec_conf_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3289 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3290 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3291 |
apply(auto simp: numeral_4_eq_4) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3292 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3293 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3294 |
lemma [simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3295 |
"rec_calc_rel rec_conf [m, r, t] rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3296 |
(rec_exec rec_conf [m, r, t] = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3297 |
apply(rule_tac prime_rel_exec_eq, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3298 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3299 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3300 |
lemma [intro]: "primerec rec_lg (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3301 |
apply(simp add: rec_lg_def Let_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3302 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3303 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3304 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3305 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3306 |
lemma [intro]: "primerec rec_nonstop (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3307 |
apply(simp add: rec_nonstop_def rec_NSTD_def rec_stat_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3308 |
rec_lo_def Let_def rec_left_def rec_right_def rec_newconf_def |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3309 |
rec_newstat_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3310 |
apply(tactic {* resolve_tac [@{thm prime_cn}, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3311 |
@{thm prime_id}, @{thm prime_pr}] 1*}, auto+)+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3312 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3313 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3314 |
lemma nonstop_eq[simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3315 |
"rec_calc_rel rec_nonstop [m, r, t] rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3316 |
(rec_exec rec_nonstop [m, r, t] = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3317 |
apply(rule prime_rel_exec_eq, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3318 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3319 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3320 |
lemma halt_lemma': |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3321 |
"rec_calc_rel rec_halt [m, r] t = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3322 |
(rec_calc_rel rec_nonstop [m, r, t] 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3323 |
(\<forall> t'< t. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3324 |
(\<exists> y. rec_calc_rel rec_nonstop [m, r, t'] y \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3325 |
y \<noteq> 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3326 |
apply(auto simp: rec_halt_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3327 |
apply(erule calc_mn_reverse, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3328 |
apply(erule_tac calc_mn_reverse) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3329 |
apply(erule_tac x = t' in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3330 |
apply(rule_tac calc_mn, simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3331 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3332 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3333 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3334 |
The following lemma gives the correctness of @{text "rec_halt"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3335 |
It says: if @{text "rec_halt"} calculates that the TM coded by @{text "m"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3336 |
will reach a standard final configuration after @{text "t"} steps of execution, then it is indeed so. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3337 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3338 |
lemma halt_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3339 |
"rec_calc_rel (rec_halt) [m, r] t = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3340 |
(rec_exec rec_nonstop [m, r, t] = 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3341 |
(\<forall> t'< t. (\<exists> y. rec_exec rec_nonstop [m, r, t'] = y |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3342 |
\<and> y \<noteq> 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3343 |
using halt_lemma'[of m r t] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3344 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3345 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3346 |
text {*F: universal machine*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3347 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3348 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3349 |
@{text "valu r"} extracts computing result out of the right number @{text "r"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3350 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3351 |
fun valu :: "nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3352 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3353 |
"valu r = (lg (r + 1) 2) - 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3354 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3355 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3356 |
@{text "rec_valu"} is the recursive function implementing @{text "valu"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3357 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3358 |
definition rec_valu :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3359 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3360 |
"rec_valu = Cn 1 rec_minus [Cn 1 rec_lg [s, constn 2], constn 1]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3361 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3362 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3363 |
The correctness of @{text "rec_valu"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3364 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3365 |
lemma value_lemma: "rec_exec rec_valu [r] = valu r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3366 |
apply(simp add: rec_exec.simps rec_valu_def lg_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3367 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3368 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3369 |
lemma [intro]: "primerec rec_valu (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3370 |
apply(simp add: rec_valu_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3371 |
apply(rule_tac k = "Suc (Suc 0)" in prime_cn) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3372 |
apply(auto simp: prime_s) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3373 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3374 |
show "primerec rec_lg (Suc (Suc 0))" by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3375 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3376 |
show "Suc (Suc 0) = Suc (Suc 0)" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3377 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3378 |
show "primerec (constn (Suc (Suc 0))) (Suc 0)" by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3379 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3380 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3381 |
lemma [simp]: "rec_calc_rel rec_valu [r] rs = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3382 |
(rec_exec rec_valu [r] = rs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3383 |
apply(rule_tac prime_rel_exec_eq, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3384 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3385 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3386 |
declare valu.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3387 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3388 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3389 |
The definition of the universal function @{text "rec_F"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3390 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3391 |
definition rec_F :: "recf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3392 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3393 |
"rec_F = Cn (Suc (Suc 0)) rec_valu [Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3394 |
rec_conf ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3395 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3396 |
lemma get_fstn_args_nth: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3397 |
"k < n \<Longrightarrow> (get_fstn_args m n ! k) = id m (k)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3398 |
apply(induct n, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3399 |
apply(case_tac "k = n", simp_all add: get_fstn_args.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3400 |
nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3401 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3402 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3403 |
lemma [simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3404 |
"\<lbrakk>ys \<noteq> []; k < length ys\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3405 |
(get_fstn_args (length ys) (length ys) ! k) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3406 |
id (length ys) (k)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3407 |
by(erule_tac get_fstn_args_nth) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3408 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3409 |
lemma calc_rel_get_pren: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3410 |
"\<lbrakk>ys \<noteq> []; k < length ys\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3411 |
rec_calc_rel (get_fstn_args (length ys) (length ys) ! k) ys |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3412 |
(ys ! k)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3413 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3414 |
apply(rule_tac calc_id, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3415 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3416 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3417 |
lemma [elim]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3418 |
"\<lbrakk>xs \<noteq> []; k < Suc (length xs)\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3419 |
rec_calc_rel (get_fstn_args (Suc (length xs)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3420 |
(Suc (length xs)) ! k) (m # xs) ((m # xs) ! k)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3421 |
using calc_rel_get_pren[of "m#xs" k] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3422 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3423 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3424 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3425 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3426 |
The correctness of @{text "rec_F"}, halt case. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3427 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3428 |
lemma F_lemma: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3429 |
"rec_calc_rel rec_halt [m, r] t \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3430 |
rec_calc_rel rec_F [m, r] (valu (rght (conf m r t)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3431 |
apply(simp add: rec_F_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3432 |
apply(rule_tac rs = "[rght (conf m r t)]" in calc_cn, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3433 |
auto simp: value_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3434 |
apply(rule_tac rs = "[conf m r t]" in calc_cn, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3435 |
auto simp: right_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3436 |
apply(rule_tac rs = "[m, r, t]" in calc_cn, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3437 |
apply(subgoal_tac " k = 0 \<or> k = Suc 0 \<or> k = Suc (Suc 0)", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3438 |
auto simp:nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3439 |
apply(rule_tac [1-2] calc_id, simp_all add: conf_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3440 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3441 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3442 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3443 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3444 |
The correctness of @{text "rec_F"}, nonhalt case. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3445 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3446 |
lemma F_lemma2: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3447 |
"\<forall> t. \<not> rec_calc_rel rec_halt [m, r] t \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3448 |
\<forall> rs. \<not> rec_calc_rel rec_F [m, r] rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3449 |
apply(auto simp: rec_F_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3450 |
apply(erule_tac calc_cn_reverse, simp (no_asm_use))+ |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3451 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3452 |
fix rs rsa rsb rsc |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3453 |
assume h: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3454 |
"\<forall>t. \<not> rec_calc_rel rec_halt [m, r] t" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3455 |
"length rsa = Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3456 |
"rec_calc_rel rec_valu rsa rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3457 |
"length rsb = Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3458 |
"rec_calc_rel rec_right rsb (rsa ! 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3459 |
"length rsc = (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3460 |
"rec_calc_rel rec_conf rsc (rsb ! 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3461 |
and g: "\<forall>k<Suc (Suc (Suc 0)). rec_calc_rel ([recf.id (Suc (Suc 0)) 0, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3462 |
recf.id (Suc (Suc 0)) (Suc 0), rec_halt] ! k) [m, r] (rsc ! k)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3463 |
have "rec_calc_rel (rec_halt ) [m, r] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3464 |
(rsc ! (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3465 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3466 |
apply(erule_tac x = "(Suc (Suc 0))" in allE) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3467 |
apply(simp add:nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3468 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3469 |
thus "False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3470 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3471 |
apply(erule_tac x = "ysb ! (Suc (Suc 0))" in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3472 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3473 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3474 |
|
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
3475 |
|
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
3476 |
subsection {* Coding function of TMs *} |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3477 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3478 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3479 |
The purpose of this section is to get the coding function of Turing Machine, which is |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3480 |
going to be named @{text "code"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3481 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3482 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3483 |
fun bl2nat :: "block list \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3484 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3485 |
"bl2nat [] n = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3486 |
| "bl2nat (Bk#bl) n = bl2nat bl (Suc n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3487 |
| "bl2nat (Oc#bl) n = 2^n + bl2nat bl (Suc n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3488 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3489 |
fun bl2wc :: "block list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3490 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3491 |
"bl2wc xs = bl2nat xs 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3492 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3493 |
fun trpl_code :: "t_conf \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3494 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3495 |
"trpl_code (st, l, r) = trpl (bl2wc l) st (bl2wc r)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3496 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3497 |
declare bl2nat.simps[simp del] bl2wc.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3498 |
trpl_code.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3499 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3500 |
fun action_map :: "taction \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3501 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3502 |
"action_map W0 = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3503 |
| "action_map W1 = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3504 |
| "action_map L = 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3505 |
| "action_map R = 3" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3506 |
| "action_map Nop = 4" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3507 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3508 |
fun action_map_iff :: "nat \<Rightarrow> taction" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3509 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3510 |
"action_map_iff (0::nat) = W0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3511 |
| "action_map_iff (Suc 0) = W1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3512 |
| "action_map_iff (Suc (Suc 0)) = L" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3513 |
| "action_map_iff (Suc (Suc (Suc 0))) = R" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3514 |
| "action_map_iff n = Nop" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3515 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3516 |
fun block_map :: "block \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3517 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3518 |
"block_map Bk = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3519 |
| "block_map Oc = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3520 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3521 |
fun godel_code' :: "nat list \<Rightarrow> nat \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3522 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3523 |
"godel_code' [] n = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3524 |
| "godel_code' (x#xs) n = (Pi n)^x * godel_code' xs (Suc n) " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3525 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3526 |
fun godel_code :: "nat list \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3527 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3528 |
"godel_code xs = (let lh = length xs in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3529 |
2^lh * (godel_code' xs (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3530 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3531 |
fun modify_tprog :: "tprog \<Rightarrow> nat list" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3532 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3533 |
"modify_tprog [] = []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3534 |
| "modify_tprog ((ac, ns)#nl) = action_map ac # ns # modify_tprog nl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3535 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3536 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3537 |
@{text "code tp"} gives the Godel coding of TM program @{text "tp"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3538 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3539 |
fun code :: "tprog \<Rightarrow> nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3540 |
where |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3541 |
"code tp = (let nl = modify_tprog tp in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3542 |
godel_code nl)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3543 |
|
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
3544 |
subsection {* Relating interperter functions to the execution of TMs *} |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3545 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3546 |
lemma [simp]: "bl2wc [] = 0" by(simp add: bl2wc.simps bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3547 |
term trpl |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3548 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3549 |
lemma [simp]: "\<lbrakk>fetch tp 0 b = (nact, ns)\<rbrakk> \<Longrightarrow> action_map nact = 4" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3550 |
apply(simp add: fetch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3551 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3552 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3553 |
lemma Pi_gr_1[simp]: "Pi n > Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3554 |
proof(induct n, auto simp: Pi.simps Np.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3555 |
fix n |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3556 |
let ?setx = "{y. y \<le> Suc (Pi n!) \<and> Pi n < y \<and> Prime y}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3557 |
have "finite ?setx" by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3558 |
moreover have "?setx \<noteq> {}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3559 |
using prime_ex[of "Pi n"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3560 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3561 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3562 |
ultimately show "Suc 0 < Min ?setx" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3563 |
apply(simp add: Min_gr_iff) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3564 |
apply(auto simp: Prime.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3565 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3566 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3567 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3568 |
lemma Pi_not_0[simp]: "Pi n > 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3569 |
using Pi_gr_1[of n] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3570 |
by arith |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3571 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3572 |
declare godel_code.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3573 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3574 |
lemma [simp]: "0 < godel_code' nl n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3575 |
apply(induct nl arbitrary: n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3576 |
apply(auto simp: godel_code'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3577 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3578 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3579 |
lemma godel_code_great: "godel_code nl > 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3580 |
apply(simp add: godel_code.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3581 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3582 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3583 |
lemma godel_code_eq_1: "(godel_code nl = 1) = (nl = [])" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3584 |
apply(auto simp: godel_code.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3585 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3586 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3587 |
lemma [elim]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3588 |
"\<lbrakk>i < length nl; \<not> Suc 0 < godel_code nl\<rbrakk> \<Longrightarrow> nl ! i = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3589 |
using godel_code_great[of nl] godel_code_eq_1[of nl] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3590 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3591 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3592 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3593 |
term set_of |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3594 |
lemma prime_coprime: "\<lbrakk>Prime x; Prime y; x\<noteq>y\<rbrakk> \<Longrightarrow> coprime x y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3595 |
proof(simp only: Prime.simps coprime_nat, auto simp: dvd_def, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3596 |
rule_tac classical, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3597 |
fix d k ka |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3598 |
assume case_ka: "\<forall>u<d * ka. \<forall>v<d * ka. u * v \<noteq> d * ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3599 |
and case_k: "\<forall>u<d * k. \<forall>v<d * k. u * v \<noteq> d * k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3600 |
and h: "(0::nat) < d" "d \<noteq> Suc 0" "Suc 0 < d * ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3601 |
"ka \<noteq> k" "Suc 0 < d * k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3602 |
from h have "k > Suc 0 \<or> ka >Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3603 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3604 |
apply(case_tac ka, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3605 |
apply(case_tac k, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3606 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3607 |
from this show "False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3608 |
proof(erule_tac disjE) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3609 |
assume "(Suc 0::nat) < k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3610 |
hence "k < d*k \<and> d < d*k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3611 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3612 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3613 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3614 |
using case_k |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3615 |
apply(erule_tac x = d in allE) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3616 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3617 |
apply(erule_tac x = k in allE) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3618 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3619 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3620 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3621 |
assume "(Suc 0::nat) < ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3622 |
hence "ka < d * ka \<and> d < d*ka" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3623 |
using h by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3624 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3625 |
using case_ka |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3626 |
apply(erule_tac x = d in allE) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3627 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3628 |
apply(erule_tac x = ka in allE) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3629 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3630 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3631 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3632 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3633 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3634 |
lemma Pi_inc: "Pi (Suc i) > Pi i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3635 |
proof(simp add: Pi.simps Np.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3636 |
let ?setx = "{y. y \<le> Suc (Pi i!) \<and> Pi i < y \<and> Prime y}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3637 |
have "finite ?setx" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3638 |
moreover have "?setx \<noteq> {}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3639 |
using prime_ex[of "Pi i"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3640 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3641 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3642 |
ultimately show "Pi i < Min ?setx" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3643 |
apply(simp add: Min_gr_iff) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3644 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3645 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3646 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3647 |
lemma Pi_inc_gr: "i < j \<Longrightarrow> Pi i < Pi j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3648 |
proof(induct j, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3649 |
fix j |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3650 |
assume ind: "i < j \<Longrightarrow> Pi i < Pi j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3651 |
and h: "i < Suc j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3652 |
from h show "Pi i < Pi (Suc j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3653 |
proof(cases "i < j") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3654 |
case True thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3655 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3656 |
assume "i < j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3657 |
hence "Pi i < Pi j" by(erule_tac ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3658 |
moreover have "Pi j < Pi (Suc j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3659 |
apply(simp add: Pi_inc) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3660 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3661 |
ultimately show "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3662 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3663 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3664 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3665 |
assume "i < Suc j" "\<not> i < j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3666 |
hence "i = j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3667 |
by arith |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3668 |
thus "Pi i < Pi (Suc j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3669 |
apply(simp add: Pi_inc) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3670 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3671 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3672 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3673 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3674 |
lemma Pi_notEq: "i \<noteq> j \<Longrightarrow> Pi i \<noteq> Pi j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3675 |
apply(case_tac "i < j") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3676 |
using Pi_inc_gr[of i j] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3677 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3678 |
using Pi_inc_gr[of j i] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3679 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3680 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3681 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3682 |
lemma [intro]: "Prime (Suc (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3683 |
apply(auto simp: Prime.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3684 |
apply(case_tac u, simp, case_tac nat, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3685 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3686 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3687 |
lemma Prime_Pi[intro]: "Prime (Pi n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3688 |
proof(induct n, auto simp: Pi.simps Np.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3689 |
fix n |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3690 |
let ?setx = "{y. y \<le> Suc (Pi n!) \<and> Pi n < y \<and> Prime y}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3691 |
show "Prime (Min ?setx)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3692 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3693 |
have "finite ?setx" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3694 |
moreover have "?setx \<noteq> {}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3695 |
using prime_ex[of "Pi n"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3696 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3697 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3698 |
ultimately show "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3699 |
apply(drule_tac Min_in, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3700 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3701 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3702 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3703 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3704 |
lemma Pi_coprime: "i \<noteq> j \<Longrightarrow> coprime (Pi i) (Pi j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3705 |
using Prime_Pi[of i] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3706 |
using Prime_Pi[of j] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3707 |
apply(rule_tac prime_coprime, simp_all add: Pi_notEq) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3708 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3709 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3710 |
lemma Pi_power_coprime: "i \<noteq> j \<Longrightarrow> coprime ((Pi i)^m) ((Pi j)^n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3711 |
by(rule_tac coprime_exp2_nat, erule_tac Pi_coprime) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3712 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3713 |
lemma coprime_dvd_mult_nat2: "\<lbrakk>coprime (k::nat) n; k dvd n * m\<rbrakk> \<Longrightarrow> k dvd m" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3714 |
apply(erule_tac coprime_dvd_mult_nat) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3715 |
apply(simp add: dvd_def, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3716 |
apply(rule_tac x = ka in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3717 |
apply(subgoal_tac "n * m = m * n", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3718 |
apply(simp add: nat_mult_commute) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3719 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3720 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3721 |
declare godel_code'.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3722 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3723 |
lemma godel_code'_butlast_last_id' : |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3724 |
"godel_code' (ys @ [y]) (Suc j) = godel_code' ys (Suc j) * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3725 |
Pi (Suc (length ys + j)) ^ y" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3726 |
proof(induct ys arbitrary: j, simp_all add: godel_code'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3727 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3728 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3729 |
lemma godel_code'_butlast_last_id: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3730 |
"xs \<noteq> [] \<Longrightarrow> godel_code' xs (Suc j) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3731 |
godel_code' (butlast xs) (Suc j) * Pi (length xs + j)^(last xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3732 |
apply(subgoal_tac "\<exists> ys y. xs = ys @ [y]") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3733 |
apply(erule_tac exE, erule_tac exE, simp add: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3734 |
godel_code'_butlast_last_id') |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3735 |
apply(rule_tac x = "butlast xs" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3736 |
apply(rule_tac x = "last xs" in exI, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3737 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3738 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3739 |
lemma godel_code'_not0: "godel_code' xs n \<noteq> 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3740 |
apply(induct xs, auto simp: godel_code'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3741 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3742 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3743 |
lemma godel_code_append_cons: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3744 |
"length xs = i \<Longrightarrow> godel_code' (xs@y#ys) (Suc 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3745 |
= godel_code' xs (Suc 0) * Pi (Suc i)^y * godel_code' ys (i + 2)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3746 |
proof(induct "length xs" arbitrary: i y ys xs, simp add: godel_code'.simps,simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3747 |
fix x xs i y ys |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3748 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3749 |
"\<And>xs i y ys. \<lbrakk>x = i; length xs = i\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3750 |
godel_code' (xs @ y # ys) (Suc 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3751 |
= godel_code' xs (Suc 0) * Pi (Suc i) ^ y * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3752 |
godel_code' ys (Suc (Suc i))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3753 |
and h: "Suc x = i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3754 |
"length (xs::nat list) = i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3755 |
have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3756 |
"godel_code' (butlast xs @ last xs # ((y::nat)#ys)) (Suc 0) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3757 |
godel_code' (butlast xs) (Suc 0) * Pi (Suc (i - 1))^(last xs) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3758 |
* godel_code' (y#ys) (Suc (Suc (i - 1)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3759 |
apply(rule_tac ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3760 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3761 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3762 |
moreover have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3763 |
"godel_code' xs (Suc 0)= godel_code' (butlast xs) (Suc 0) * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3764 |
Pi (i)^(last xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3765 |
using godel_code'_butlast_last_id[of xs] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3766 |
apply(case_tac "xs = []", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3767 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3768 |
moreover have "butlast xs @ last xs # y # ys = xs @ y # ys" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3769 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3770 |
apply(case_tac xs, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3771 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3772 |
ultimately show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3773 |
"godel_code' (xs @ y # ys) (Suc 0) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3774 |
godel_code' xs (Suc 0) * Pi (Suc i) ^ y * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3775 |
godel_code' ys (Suc (Suc i))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3776 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3777 |
apply(simp add: godel_code'_not0 Pi_not_0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3778 |
apply(simp add: godel_code'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3779 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3780 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3781 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3782 |
lemma Pi_coprime_pre: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3783 |
"length ps \<le> i \<Longrightarrow> coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3784 |
proof(induct "length ps" arbitrary: ps, simp add: godel_code'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3785 |
fix x ps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3786 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3787 |
"\<And>ps. \<lbrakk>x = length ps; length ps \<le> i\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3788 |
coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3789 |
and h: "Suc x = length ps" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3790 |
"length (ps::nat list) \<le> i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3791 |
have g: "coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3792 |
apply(rule_tac ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3793 |
using h by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3794 |
have k: "godel_code' ps (Suc 0) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3795 |
godel_code' (butlast ps) (Suc 0) * Pi (length ps)^(last ps)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3796 |
using godel_code'_butlast_last_id[of ps 0] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3797 |
by(case_tac ps, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3798 |
from g have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3799 |
"coprime (Pi (Suc i)) (godel_code' (butlast ps) (Suc 0) * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3800 |
Pi (length ps)^(last ps)) " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3801 |
proof(rule_tac coprime_mult_nat, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3802 |
show "coprime (Pi (Suc i)) (Pi (length ps) ^ last ps)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3803 |
apply(rule_tac coprime_exp_nat, rule prime_coprime, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3804 |
using Pi_notEq[of "Suc i" "length ps"] h by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3805 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3806 |
from this and k show "coprime (Pi (Suc i)) (godel_code' ps (Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3807 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3808 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3809 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3810 |
lemma Pi_coprime_suf: "i < j \<Longrightarrow> coprime (Pi i) (godel_code' ps j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3811 |
proof(induct "length ps" arbitrary: ps, simp add: godel_code'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3812 |
fix x ps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3813 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3814 |
"\<And>ps. \<lbrakk>x = length ps; i < j\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3815 |
coprime (Pi i) (godel_code' ps j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3816 |
and h: "Suc x = length (ps::nat list)" "i < j" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3817 |
have g: "coprime (Pi i) (godel_code' (butlast ps) j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3818 |
apply(rule ind) using h by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3819 |
have k: "(godel_code' ps j) = godel_code' (butlast ps) j * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3820 |
Pi (length ps + j - 1)^last ps" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3821 |
using h godel_code'_butlast_last_id[of ps "j - 1"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3822 |
apply(case_tac "ps = []", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3823 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3824 |
from g have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3825 |
"coprime (Pi i) (godel_code' (butlast ps) j * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3826 |
Pi (length ps + j - 1)^last ps)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3827 |
apply(rule_tac coprime_mult_nat, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3828 |
using Pi_power_coprime[of i "length ps + j - 1" 1 "last ps"] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3829 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3830 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3831 |
from k and this show "coprime (Pi i) (godel_code' ps j)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3832 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3833 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3834 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3835 |
lemma godel_finite: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3836 |
"finite {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3837 |
proof(rule_tac n = "godel_code' nl (Suc 0)" in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3838 |
bounded_nat_set_is_finite, auto, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3839 |
case_tac "ia < godel_code' nl (Suc 0)", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3840 |
fix ia |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3841 |
assume g1: "Pi (Suc i) ^ ia dvd godel_code' nl (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3842 |
and g2: "\<not> ia < godel_code' nl (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3843 |
from g1 have "Pi (Suc i)^ia \<le> godel_code' nl (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3844 |
apply(erule_tac dvd_imp_le) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3845 |
using godel_code'_not0[of nl "Suc 0"] by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3846 |
moreover have "ia < Pi (Suc i)^ia" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3847 |
apply(rule x_less_exp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3848 |
using Pi_gr_1 by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3849 |
ultimately show "False" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3850 |
using g2 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3851 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3852 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3853 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3854 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3855 |
lemma godel_code_in: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3856 |
"i < length nl \<Longrightarrow> nl ! i \<in> {u. Pi (Suc i) ^ u dvd |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3857 |
godel_code' nl (Suc 0)}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3858 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3859 |
assume h: "i<length nl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3860 |
hence "godel_code' (take i nl@(nl!i)#drop (Suc i) nl) (Suc 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3861 |
= godel_code' (take i nl) (Suc 0) * Pi (Suc i)^(nl!i) * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3862 |
godel_code' (drop (Suc i) nl) (i + 2)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3863 |
by(rule_tac godel_code_append_cons, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3864 |
moreover from h have "take i nl @ (nl ! i) # drop (Suc i) nl = nl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3865 |
using upd_conv_take_nth_drop[of i nl "nl ! i"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3866 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3867 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3868 |
ultimately show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3869 |
"nl ! i \<in> {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3870 |
by(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3871 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3872 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3873 |
lemma godel_code'_get_nth: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3874 |
"i < length nl \<Longrightarrow> Max {u. Pi (Suc i) ^ u dvd |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3875 |
godel_code' nl (Suc 0)} = nl ! i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3876 |
proof(rule_tac Max_eqI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3877 |
let ?gc = "godel_code' nl (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3878 |
assume h: "i < length nl" thus "finite {u. Pi (Suc i) ^ u dvd ?gc}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3879 |
by (simp add: godel_finite) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3880 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3881 |
fix y |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3882 |
let ?suf ="godel_code' (drop (Suc i) nl) (i + 2)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3883 |
let ?pref = "godel_code' (take i nl) (Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3884 |
assume h: "i < length nl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3885 |
"y \<in> {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3886 |
moreover hence |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3887 |
"godel_code' (take i nl@(nl!i)#drop (Suc i) nl) (Suc 0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3888 |
= ?pref * Pi (Suc i)^(nl!i) * ?suf" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3889 |
by(rule_tac godel_code_append_cons, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3890 |
moreover from h have "take i nl @ (nl!i) # drop (Suc i) nl = nl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3891 |
using upd_conv_take_nth_drop[of i nl "nl!i"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3892 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3893 |
ultimately show "y\<le>nl!i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3894 |
proof(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3895 |
let ?suf' = "godel_code' (drop (Suc i) nl) (Suc (Suc i))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3896 |
assume mult_dvd: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3897 |
"Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i * ?suf'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3898 |
hence "Pi (Suc i) ^ y dvd ?pref * Pi (Suc i) ^ nl ! i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3899 |
proof(rule_tac coprime_dvd_mult_nat) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3900 |
show "coprime (Pi (Suc i)^y) ?suf'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3901 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3902 |
have "coprime (Pi (Suc i) ^ y) (?suf'^(Suc 0))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3903 |
apply(rule_tac coprime_exp2_nat) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3904 |
apply(rule_tac Pi_coprime_suf, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3905 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3906 |
thus "?thesis" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3907 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3908 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3909 |
hence "Pi (Suc i) ^ y dvd Pi (Suc i) ^ nl ! i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3910 |
proof(rule_tac coprime_dvd_mult_nat2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3911 |
show "coprime (Pi (Suc i) ^ y) ?pref" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3912 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3913 |
have "coprime (Pi (Suc i)^y) (?pref^Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3914 |
apply(rule_tac coprime_exp2_nat) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3915 |
apply(rule_tac Pi_coprime_pre, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3916 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3917 |
thus "?thesis" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3918 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3919 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3920 |
hence "Pi (Suc i) ^ y \<le> Pi (Suc i) ^ nl ! i " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3921 |
apply(rule_tac dvd_imp_le, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3922 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3923 |
thus "y \<le> nl ! i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3924 |
apply(rule_tac power_le_imp_le_exp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3925 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3926 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3927 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3928 |
assume h: "i<length nl" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3929 |
thus "nl ! i \<in> {u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3930 |
by(rule_tac godel_code_in, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3931 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3932 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3933 |
lemma [simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3934 |
"{u. Pi (Suc i) ^ u dvd (Suc (Suc 0)) ^ length nl * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3935 |
godel_code' nl (Suc 0)} = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3936 |
{u. Pi (Suc i) ^ u dvd godel_code' nl (Suc 0)}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3937 |
apply(rule_tac Collect_cong, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3938 |
apply(rule_tac n = " (Suc (Suc 0)) ^ length nl" in |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3939 |
coprime_dvd_mult_nat2) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3940 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3941 |
fix u |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3942 |
show "coprime (Pi (Suc i) ^ u) ((Suc (Suc 0)) ^ length nl)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3943 |
proof(rule_tac coprime_exp2_nat) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3944 |
have "Pi 0 = (2::nat)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3945 |
apply(simp add: Pi.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3946 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3947 |
moreover have "coprime (Pi (Suc i)) (Pi 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3948 |
apply(rule_tac Pi_coprime, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3949 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3950 |
ultimately show "coprime (Pi (Suc i)) (Suc (Suc 0))" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3951 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3952 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3953 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3954 |
lemma godel_code_get_nth: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3955 |
"i < length nl \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3956 |
Max {u. Pi (Suc i) ^ u dvd godel_code nl} = nl ! i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3957 |
by(simp add: godel_code.simps godel_code'_get_nth) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3958 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3959 |
lemma "trpl l st r = godel_code' [l, st, r] 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3960 |
apply(simp add: trpl.simps godel_code'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3961 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3962 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3963 |
lemma mod_dvd_simp: "(x mod y = (0::nat)) = (y dvd x)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3964 |
by(simp add: dvd_def, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3965 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3966 |
lemma dvd_power_le: "\<lbrakk>a > Suc 0; a ^ y dvd a ^ l\<rbrakk> \<Longrightarrow> y \<le> l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3967 |
apply(case_tac "y \<le> l", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3968 |
apply(subgoal_tac "\<exists> d. y = l + d", auto simp: power_add) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3969 |
apply(rule_tac x = "y - l" in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3970 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3971 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3972 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3973 |
lemma [elim]: "Pi n = 0 \<Longrightarrow> RR" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3974 |
using Pi_not_0[of n] by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3975 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3976 |
lemma [elim]: "Pi n = Suc 0 \<Longrightarrow> RR" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3977 |
using Pi_gr_1[of n] by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3978 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3979 |
lemma finite_power_dvd: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3980 |
"\<lbrakk>(a::nat) > Suc 0; y \<noteq> 0\<rbrakk> \<Longrightarrow> finite {u. a^u dvd y}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3981 |
apply(auto simp: dvd_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3982 |
apply(rule_tac n = y in bounded_nat_set_is_finite, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3983 |
apply(case_tac k, simp,simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3984 |
apply(rule_tac trans_less_add1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3985 |
apply(erule_tac x_less_exp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3986 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3987 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3988 |
lemma conf_decode1: "\<lbrakk>m \<noteq> n; m \<noteq> k; k \<noteq> n\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3989 |
Max {u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r} = l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3990 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3991 |
let ?setx = "{u. Pi m ^ u dvd Pi m ^ l * Pi n ^ st * Pi k ^ r}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3992 |
assume g: "m \<noteq> n" "m \<noteq> k" "k \<noteq> n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3993 |
show "Max ?setx = l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3994 |
proof(rule_tac Max_eqI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3995 |
show "finite ?setx" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3996 |
apply(rule_tac finite_power_dvd, auto simp: Pi_gr_1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3997 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3998 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
3999 |
fix y |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4000 |
assume h: "y \<in> ?setx" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4001 |
have "Pi m ^ y dvd Pi m ^ l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4002 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4003 |
have "Pi m ^ y dvd Pi m ^ l * Pi n ^ st" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4004 |
using h g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4005 |
apply(rule_tac n = "Pi k^r" in coprime_dvd_mult_nat) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4006 |
apply(rule Pi_power_coprime, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4007 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4008 |
thus "Pi m^y dvd Pi m^l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4009 |
apply(rule_tac n = " Pi n ^ st" in coprime_dvd_mult_nat) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4010 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4011 |
apply(rule_tac Pi_power_coprime, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4012 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4013 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4014 |
thus "y \<le> (l::nat)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4015 |
apply(rule_tac a = "Pi m" in power_le_imp_le_exp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4016 |
apply(simp_all add: Pi_gr_1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4017 |
apply(rule_tac dvd_power_le, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4018 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4019 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4020 |
show "l \<in> ?setx" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4021 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4022 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4023 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4024 |
lemma conf_decode2: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4025 |
"\<lbrakk>m \<noteq> n; m \<noteq> k; n \<noteq> k; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4026 |
\<not> Suc 0 < Pi m ^ l * Pi n ^ st * Pi k ^ r\<rbrakk> \<Longrightarrow> l = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4027 |
apply(case_tac "Pi m ^ l * Pi n ^ st * Pi k ^ r", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4028 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4029 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4030 |
lemma [simp]: "left (trpl l st r) = l" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4031 |
apply(simp add: left.simps trpl.simps lo.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4032 |
loR.simps mod_dvd_simp, auto simp: conf_decode1) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4033 |
apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4034 |
auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4035 |
apply(erule_tac x = l in allE, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4036 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4037 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4038 |
lemma [simp]: "stat (trpl l st r) = st" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4039 |
apply(simp add: stat.simps trpl.simps lo.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4040 |
loR.simps mod_dvd_simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4041 |
apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4042 |
= Pi (Suc 0)^st * Pi 0 ^ l * Pi (Suc (Suc 0)) ^ r") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4043 |
apply(simp (no_asm_simp) add: conf_decode1, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4044 |
apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4045 |
Pi (Suc (Suc 0)) ^ r", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4046 |
apply(erule_tac x = st in allE, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4047 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4048 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4049 |
lemma [simp]: "rght (trpl l st r) = r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4050 |
apply(simp add: rght.simps trpl.simps lo.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4051 |
loR.simps mod_dvd_simp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4052 |
apply(subgoal_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4053 |
= Pi (Suc (Suc 0))^r * Pi 0 ^ l * Pi (Suc 0) ^ st") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4054 |
apply(simp (no_asm_simp) add: conf_decode1, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4055 |
apply(case_tac "Pi 0 ^ l * Pi (Suc 0) ^ st * Pi (Suc (Suc 0)) ^ r", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4056 |
auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4057 |
apply(erule_tac x = r in allE, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4058 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4059 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4060 |
lemma max_lor: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4061 |
"i < length nl \<Longrightarrow> Max {u. loR [godel_code nl, Pi (Suc i), u]} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4062 |
= nl ! i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4063 |
apply(simp add: loR.simps godel_code_get_nth mod_dvd_simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4064 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4065 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4066 |
lemma godel_decode: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4067 |
"i < length nl \<Longrightarrow> Entry (godel_code nl) i = nl ! i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4068 |
apply(auto simp: Entry.simps lo.simps max_lor) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4069 |
apply(erule_tac x = "nl!i" in allE) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4070 |
using max_lor[of i nl] godel_finite[of i nl] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4071 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4072 |
apply(drule_tac Max_in, auto simp: loR.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4073 |
godel_code.simps mod_dvd_simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4074 |
using godel_code_in[of i nl] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4075 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4076 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4077 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4078 |
lemma Four_Suc: "4 = Suc (Suc (Suc (Suc 0)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4079 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4080 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4081 |
declare numeral_2_eq_2[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4082 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4083 |
lemma modify_tprog_fetch_even: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4084 |
"\<lbrakk>st \<le> length tp div 2; st > 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4085 |
modify_tprog tp ! (4 * (st - Suc 0) ) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4086 |
action_map (fst (tp ! (2 * (st - Suc 0))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4087 |
proof(induct st arbitrary: tp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4088 |
fix tp st |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4089 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4090 |
"\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4091 |
modify_tprog tp ! (4 * (st - Suc 0)) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4092 |
action_map (fst ((tp::tprog) ! (2 * (st - Suc 0))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4093 |
and h: "Suc st \<le> length (tp::tprog) div 2" "0 < Suc st" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4094 |
thus "modify_tprog tp ! (4 * (Suc st - Suc 0)) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4095 |
action_map (fst (tp ! (2 * (Suc st - Suc 0))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4096 |
proof(cases "st = 0") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4097 |
case True thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4098 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4099 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4100 |
apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4101 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4102 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4103 |
case False |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4104 |
assume g: "st \<noteq> 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4105 |
hence "\<exists> aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4106 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4107 |
apply(case_tac tp, simp, case_tac list, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4108 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4109 |
from this obtain aa ab ba bb tp' where g1: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4110 |
"tp = (aa, ab) # (ba, bb) # tp'" by blast |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4111 |
hence g2: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4112 |
"modify_tprog tp' ! (4 * (st - Suc 0)) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4113 |
action_map (fst ((tp'::tprog) ! (2 * (st - Suc 0))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4114 |
apply(rule_tac ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4115 |
using h g by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4116 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4117 |
using g1 g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4118 |
apply(case_tac st, simp, simp add: Four_Suc) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4119 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4120 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4121 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4122 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4123 |
lemma modify_tprog_fetch_odd: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4124 |
"\<lbrakk>st \<le> length tp div 2; st > 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4125 |
modify_tprog tp ! (Suc (Suc (4 * (st - Suc 0)))) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4126 |
action_map (fst (tp ! (Suc (2 * (st - Suc 0)))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4127 |
proof(induct st arbitrary: tp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4128 |
fix tp st |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4129 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4130 |
"\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4131 |
modify_tprog tp ! Suc (Suc (4 * (st - Suc 0))) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4132 |
action_map (fst (tp ! Suc (2 * (st - Suc 0))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4133 |
and h: "Suc st \<le> length (tp::tprog) div 2" "0 < Suc st" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4134 |
thus "modify_tprog tp ! Suc (Suc (4 * (Suc st - Suc 0))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4135 |
= action_map (fst (tp ! Suc (2 * (Suc st - Suc 0))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4136 |
proof(cases "st = 0") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4137 |
case True thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4138 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4139 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4140 |
apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4141 |
apply(case_tac list, simp, case_tac ab, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4142 |
simp add: modify_tprog.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4143 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4144 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4145 |
case False |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4146 |
assume g: "st \<noteq> 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4147 |
hence "\<exists> aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4148 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4149 |
apply(case_tac tp, simp, case_tac list, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4150 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4151 |
from this obtain aa ab ba bb tp' where g1: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4152 |
"tp = (aa, ab) # (ba, bb) # tp'" by blast |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4153 |
hence g2: "modify_tprog tp' ! Suc (Suc (4 * (st - Suc 0))) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4154 |
action_map (fst (tp' ! Suc (2 * (st - Suc 0))))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4155 |
apply(rule_tac ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4156 |
using h g by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4157 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4158 |
using g1 g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4159 |
apply(case_tac st, simp, simp add: Four_Suc) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4160 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4161 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4162 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4163 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4164 |
lemma modify_tprog_fetch_action: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4165 |
"\<lbrakk>st \<le> length tp div 2; st > 0; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4166 |
modify_tprog tp ! (4 * (st - Suc 0) + 2* b) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4167 |
action_map (fst (tp ! ((2 * (st - Suc 0)) + b)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4168 |
apply(erule_tac disjE, auto elim: modify_tprog_fetch_odd |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4169 |
modify_tprog_fetch_even) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4170 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4171 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4172 |
lemma length_modify: "length (modify_tprog tp) = 2 * length tp" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4173 |
apply(induct tp, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4174 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4175 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4176 |
declare fetch.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4177 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4178 |
lemma fetch_action_eq: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4179 |
"\<lbrakk>block_map b = scan r; fetch tp st b = (nact, ns); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4180 |
st \<le> length tp div 2\<rbrakk> \<Longrightarrow> actn (code tp) st r = action_map nact" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4181 |
proof(simp add: actn.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4182 |
let ?i = "4 * (st - Suc 0) + 2 * (r mod 2)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4183 |
assume h: "block_map b = r mod 2" "fetch tp st b = (nact, ns)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4184 |
"st \<le> length tp div 2" "0 < st" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4185 |
have "?i < length (modify_tprog tp)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4186 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4187 |
have "length (modify_tprog tp) = 2 * length tp" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4188 |
by(simp add: length_modify) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4189 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4190 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4191 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4192 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4193 |
hence |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4194 |
"Entry (godel_code (modify_tprog tp))?i = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4195 |
(modify_tprog tp) ! ?i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4196 |
by(erule_tac godel_decode) |
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
4197 |
moreover have |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4198 |
"modify_tprog tp ! ?i = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4199 |
action_map (fst (tp ! (2 * (st - Suc 0) + r mod 2)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4200 |
apply(rule_tac modify_tprog_fetch_action) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4201 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4202 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4203 |
moreover have "(fst (tp ! (2 * (st - Suc 0) + r mod 2))) = nact" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4204 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4205 |
apply(simp add: fetch.simps nth_of.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4206 |
apply(case_tac b, auto simp: block_map.simps nth_of.simps split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4207 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4208 |
ultimately show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4209 |
"Entry (godel_code (modify_tprog tp)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4210 |
(4 * (st - Suc 0) + 2 * (r mod 2)) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4211 |
= action_map nact" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4212 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4213 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4214 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4215 |
lemma [simp]: "fetch tp 0 b = (nact, ns) \<Longrightarrow> ns = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4216 |
by(simp add: fetch.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4217 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4218 |
lemma Five_Suc: "5 = Suc 4" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4219 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4220 |
lemma modify_tprog_fetch_state: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4221 |
"\<lbrakk>st \<le> length tp div 2; st > 0; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4222 |
modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4223 |
(snd (tp ! (2 * (st - Suc 0) + b)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4224 |
proof(induct st arbitrary: tp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4225 |
fix st tp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4226 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4227 |
"\<And>tp. \<lbrakk>st \<le> length tp div 2; 0 < st; b = 1 \<or> b = 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4228 |
modify_tprog tp ! Suc (4 * (st - Suc 0) + 2 * b) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4229 |
snd (tp ! (2 * (st - Suc 0) + b))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4230 |
and h: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4231 |
"Suc st \<le> length (tp::tprog) div 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4232 |
"0 < Suc st" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4233 |
"b = 1 \<or> b = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4234 |
show "modify_tprog tp ! Suc (4 * (Suc st - Suc 0) + 2 * b) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4235 |
snd (tp ! (2 * (Suc st - Suc 0) + b))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4236 |
proof(cases "st = 0") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4237 |
case True |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4238 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4239 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4240 |
apply(cases tp, simp, case_tac a, simp add: modify_tprog.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4241 |
apply(case_tac list, simp, case_tac ab, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4242 |
simp add: modify_tprog.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4243 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4244 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4245 |
case False |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4246 |
assume g: "st \<noteq> 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4247 |
hence "\<exists> aa ab ba bb tp'. tp = (aa, ab) # (ba, bb) # tp'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4248 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4249 |
apply(case_tac tp, simp, case_tac list, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4250 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4251 |
from this obtain aa ab ba bb tp' where g1: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4252 |
"tp = (aa, ab) # (ba, bb) # tp'" by blast |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4253 |
hence g2: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4254 |
"modify_tprog tp' ! Suc (4 * (st - Suc 0) + 2 * b) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4255 |
snd (tp' ! (2 * (st - Suc 0) + b))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4256 |
apply(rule_tac ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4257 |
using h g by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4258 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4259 |
using g1 g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4260 |
apply(case_tac st, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4261 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4262 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4263 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4264 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4265 |
lemma fetch_state_eq: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4266 |
"\<lbrakk>block_map b = scan r; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4267 |
fetch tp st b = (nact, ns); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4268 |
st \<le> length tp div 2\<rbrakk> \<Longrightarrow> newstat (code tp) st r = ns" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4269 |
proof(simp add: newstat.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4270 |
let ?i = "Suc (4 * (st - Suc 0) + 2 * (r mod 2))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4271 |
assume h: "block_map b = r mod 2" "fetch tp st b = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4272 |
(nact, ns)" "st \<le> length tp div 2" "0 < st" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4273 |
have "?i < length (modify_tprog tp)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4274 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4275 |
have "length (modify_tprog tp) = 2 * length tp" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4276 |
apply(simp add: length_modify) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4277 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4278 |
thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4279 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4280 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4281 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4282 |
hence "Entry (godel_code (modify_tprog tp)) (?i) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4283 |
(modify_tprog tp) ! ?i" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4284 |
by(erule_tac godel_decode) |
371
48b231495281
Some illustration added together with more explanations.
zhang
parents:
370
diff
changeset
|
4285 |
moreover have |
370
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4286 |
"modify_tprog tp ! ?i = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4287 |
(snd (tp ! (2 * (st - Suc 0) + r mod 2)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4288 |
apply(rule_tac modify_tprog_fetch_state) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4289 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4290 |
by(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4291 |
moreover have "(snd (tp ! (2 * (st - Suc 0) + r mod 2))) = ns" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4292 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4293 |
apply(simp add: fetch.simps nth_of.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4294 |
apply(case_tac b, auto simp: block_map.simps nth_of.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4295 |
split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4296 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4297 |
ultimately show "Entry (godel_code (modify_tprog tp)) (?i) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4298 |
= ns" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4299 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4300 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4301 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4302 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4303 |
lemma [intro!]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4304 |
"\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> trpl a b c = trpl a' b' c'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4305 |
by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4306 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4307 |
lemma [simp]: "bl2wc [Bk] = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4308 |
by(simp add: bl2wc.simps bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4309 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4310 |
lemma bl2nat_double: "bl2nat xs (Suc n) = 2 * bl2nat xs n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4311 |
proof(induct xs arbitrary: n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4312 |
case Nil thus "?case" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4313 |
by(simp add: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4314 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4315 |
case (Cons x xs) thus "?case" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4316 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4317 |
assume ind: "\<And>n. bl2nat xs (Suc n) = 2 * bl2nat xs n " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4318 |
show "bl2nat (x # xs) (Suc n) = 2 * bl2nat (x # xs) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4319 |
proof(cases x) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4320 |
case Bk thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4321 |
apply(simp add: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4322 |
using ind[of "Suc n"] by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4323 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4324 |
case Oc thus "?thesis" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4325 |
apply(simp add: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4326 |
using ind[of "Suc n"] by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4327 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4328 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4329 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4330 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4331 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4332 |
lemma [simp]: "c \<noteq> [] \<Longrightarrow> 2 * bl2wc (tl c) = bl2wc c - bl2wc c mod 2 " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4333 |
apply(case_tac c, simp, case_tac a) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4334 |
apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4335 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4336 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4337 |
lemma [simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4338 |
"c \<noteq> [] \<Longrightarrow> bl2wc (Oc # tl c) = Suc (bl2wc c) - bl2wc c mod 2 " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4339 |
apply(case_tac c, simp, case_tac a) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4340 |
apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4341 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4342 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4343 |
lemma [simp]: "bl2wc (Bk # c) = 2*bl2wc (c)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4344 |
apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4345 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4346 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4347 |
lemma [simp]: "bl2wc [Oc] = Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4348 |
by(simp add: bl2wc.simps bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4349 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4350 |
lemma [simp]: "b \<noteq> [] \<Longrightarrow> bl2wc (tl b) = bl2wc b div 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4351 |
apply(case_tac b, simp, case_tac a) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4352 |
apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4353 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4354 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4355 |
lemma [simp]: "b \<noteq> [] \<Longrightarrow> bl2wc ([hd b]) = bl2wc b mod 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4356 |
apply(case_tac b, simp, case_tac a) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4357 |
apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4358 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4359 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4360 |
lemma [simp]: "\<lbrakk>b \<noteq> []; c \<noteq> []\<rbrakk> \<Longrightarrow> bl2wc (hd b # c) = 2 * bl2wc c + bl2wc b mod 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4361 |
apply(case_tac b, simp, case_tac a) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4362 |
apply(auto simp: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4363 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4364 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4365 |
lemma [simp]: " 2 * (bl2wc c div 2) = bl2wc c - bl2wc c mod 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4366 |
by(simp add: mult_div_cancel) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4367 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4368 |
lemma [simp]: "bl2wc (Oc # list) mod 2 = Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4369 |
by(simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4370 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4371 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4372 |
declare code.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4373 |
declare nth_of.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4374 |
declare new_tape.simps[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4375 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4376 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4377 |
The lemma relates the one step execution of TMs with the interpreter function @{text "rec_newconf"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4378 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4379 |
lemma rec_t_eq_step: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4380 |
"(\<lambda> (s, l, r). s \<le> length tp div 2) c \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4381 |
trpl_code (tstep c tp) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4382 |
rec_exec rec_newconf [code tp, trpl_code c]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4383 |
apply(cases c, auto simp: tstep.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4384 |
proof(case_tac "fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4385 |
simp add: newconf.simps trpl_code.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4386 |
fix a b c aa ba |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4387 |
assume h: "(a::nat) \<le> length tp div 2" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4388 |
"fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, ba)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4389 |
moreover hence "actn (code tp) a (bl2wc c) = action_map aa" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4390 |
apply(rule_tac b = "(case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4391 |
in fetch_action_eq, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4392 |
apply(auto split: list.splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4393 |
apply(case_tac ab, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4394 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4395 |
moreover from h have "(newstat (code tp) a (bl2wc c)) = ba" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4396 |
apply(rule_tac b = "(case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4397 |
in fetch_state_eq, auto split: list.splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4398 |
apply(case_tac ab, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4399 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4400 |
ultimately show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4401 |
"trpl_code (ba, new_tape aa (b, c)) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4402 |
trpl (newleft (bl2wc b) (bl2wc c) (actn (code tp) a (bl2wc c))) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4403 |
(newstat (code tp) a (bl2wc c)) (newrght (bl2wc b) (bl2wc c) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4404 |
(actn (code tp) a (bl2wc c)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4405 |
by(auto simp: new_tape.simps trpl_code.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4406 |
newleft.simps newrght.simps split: taction.splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4407 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4408 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4409 |
lemma [simp]: "a\<^bsup>0\<^esup> = []" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4410 |
apply(simp add: exp_zero) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4411 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4412 |
lemma [simp]: "bl2nat (Oc # Oc\<^bsup>x\<^esup>) 0 = (2 * 2 ^ x - Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4413 |
apply(induct x) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4414 |
apply(simp add: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4415 |
apply(simp add: bl2nat.simps bl2nat_double exp_ind_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4416 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4417 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4418 |
lemma [simp]: "bl2nat (Oc\<^bsup>y\<^esup>) 0 = 2^y - Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4419 |
apply(induct y, auto simp: bl2nat.simps exp_ind_def bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4420 |
apply(case_tac "(2::nat)^y", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4421 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4422 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4423 |
lemma [simp]: "bl2nat (Bk\<^bsup>l\<^esup>) n = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4424 |
apply(induct l, auto simp: bl2nat.simps bl2nat_double exp_ind_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4425 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4426 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4427 |
lemma bl2nat_cons_bk: "bl2nat (ks @ [Bk]) 0 = bl2nat ks 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4428 |
apply(induct ks, auto simp: bl2nat.simps split: block.splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4429 |
apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4430 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4431 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4432 |
lemma bl2nat_cons_oc: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4433 |
"bl2nat (ks @ [Oc]) 0 = bl2nat ks 0 + 2 ^ length ks" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4434 |
apply(induct ks, auto simp: bl2nat.simps split: block.splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4435 |
apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4436 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4437 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4438 |
lemma bl2nat_append: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4439 |
"bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs) " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4440 |
proof(induct "length xs" arbitrary: xs ys, simp add: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4441 |
fix x xs ys |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4442 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4443 |
"\<And>xs ys. x = length xs \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4444 |
bl2nat (xs @ ys) 0 = bl2nat xs 0 + bl2nat ys (length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4445 |
and h: "Suc x = length (xs::block list)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4446 |
have "\<exists> ks k. xs = ks @ [k]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4447 |
apply(rule_tac x = "butlast xs" in exI, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4448 |
rule_tac x = "last xs" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4449 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4450 |
apply(case_tac xs, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4451 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4452 |
from this obtain ks k where "xs = ks @ [k]" by blast |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4453 |
moreover hence |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4454 |
"bl2nat (ks @ (k # ys)) 0 = bl2nat ks 0 + |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4455 |
bl2nat (k # ys) (length ks)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4456 |
apply(rule_tac ind) using h by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4457 |
ultimately show "bl2nat (xs @ ys) 0 = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4458 |
bl2nat xs 0 + bl2nat ys (length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4459 |
apply(case_tac k, simp_all add: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4460 |
apply(simp_all only: bl2nat_cons_bk bl2nat_cons_oc) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4461 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4462 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4463 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4464 |
lemma bl2nat_exp: "n \<noteq> 0 \<Longrightarrow> bl2nat bl n = 2^n * bl2nat bl 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4465 |
apply(induct bl) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4466 |
apply(auto simp: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4467 |
apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4468 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4469 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4470 |
lemma nat_minus_eq: "\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> a - c = b - d" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4471 |
by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4472 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4473 |
lemma tape_of_nat_list_butlast_last: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4474 |
"ys \<noteq> [] \<Longrightarrow> <ys @ [y]> = <ys> @ Bk # Oc\<^bsup>Suc y\<^esup>" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4475 |
apply(induct ys, simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4476 |
apply(case_tac "ys = []", simp add: tape_of_nl_abv |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4477 |
tape_of_nat_list.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4478 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4479 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4480 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4481 |
lemma listsum2_append: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4482 |
"\<lbrakk>n \<le> length xs\<rbrakk> \<Longrightarrow> listsum2 (xs @ ys) n = listsum2 xs n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4483 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4484 |
apply(auto simp: listsum2.simps nth_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4485 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4486 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4487 |
lemma strt'_append: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4488 |
"\<lbrakk>n \<le> length xs\<rbrakk> \<Longrightarrow> strt' xs n = strt' (xs @ ys) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4489 |
proof(induct n arbitrary: xs ys) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4490 |
fix xs ys |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4491 |
show "strt' xs 0 = strt' (xs @ ys) 0" by(simp add: strt'.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4492 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4493 |
fix n xs ys |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4494 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4495 |
"\<And> xs ys. n \<le> length xs \<Longrightarrow> strt' xs n = strt' (xs @ ys) n" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4496 |
and h: "Suc n \<le> length (xs::nat list)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4497 |
show "strt' xs (Suc n) = strt' (xs @ ys) (Suc n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4498 |
using ind[of xs ys] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4499 |
apply(simp add: strt'.simps nth_append listsum2_append) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4500 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4501 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4502 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4503 |
lemma length_listsum2_eq: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4504 |
"\<lbrakk>length (ys::nat list) = k\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4505 |
\<Longrightarrow> length (<ys>) = listsum2 (map Suc ys) k + k - 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4506 |
apply(induct k arbitrary: ys, simp_all add: listsum2.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4507 |
apply(subgoal_tac "\<exists> xs x. ys = xs @ [x]", auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4508 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4509 |
fix xs x |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4510 |
assume ind: "\<And>ys. length ys = length xs \<Longrightarrow> length (<ys>) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4511 |
= listsum2 (map Suc ys) (length xs) + |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4512 |
length (xs::nat list) - Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4513 |
have "length (<xs>) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4514 |
= listsum2 (map Suc xs) (length xs) + length xs - Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4515 |
apply(rule_tac ind, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4516 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4517 |
thus "length (<xs @ [x]>) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4518 |
Suc (listsum2 (map Suc xs @ [Suc x]) (length xs) + x + length xs)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4519 |
apply(case_tac "xs = []") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4520 |
apply(simp add: tape_of_nl_abv listsum2.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4521 |
tape_of_nat_list.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4522 |
apply(simp add: tape_of_nat_list_butlast_last) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4523 |
using listsum2_append[of "length xs" "map Suc xs" "[Suc x]"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4524 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4525 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4526 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4527 |
fix k ys |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4528 |
assume "length ys = Suc k" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4529 |
thus "\<exists>xs x. ys = xs @ [x]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4530 |
apply(rule_tac x = "butlast ys" in exI, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4531 |
rule_tac x = "last ys" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4532 |
apply(case_tac ys, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4533 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4534 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4535 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4536 |
lemma tape_of_nat_list_length: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4537 |
"length (<(ys::nat list)>) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4538 |
listsum2 (map Suc ys) (length ys) + length ys - 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4539 |
using length_listsum2_eq[of ys "length ys"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4540 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4541 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4542 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4543 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4544 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4545 |
lemma [simp]: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4546 |
"trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp 0) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4547 |
rec_exec rec_conf [code tp, bl2wc (<lm>), 0]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4548 |
apply(simp add: steps.simps rec_exec.simps conf_lemma conf.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4549 |
inpt.simps trpl_code.simps bl2wc.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4550 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4551 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4552 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4553 |
The following lemma relates the multi-step interpreter function @{text "rec_conf"} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4554 |
with the multi-step execution of TMs. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4555 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4556 |
lemma rec_t_eq_steps: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4557 |
"turing_basic.t_correct tp \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4558 |
trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4559 |
rec_exec rec_conf [code tp, bl2wc (<lm>), stp]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4560 |
proof(induct stp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4561 |
case 0 thus "?case" by(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4562 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4563 |
case (Suc n) thus "?case" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4564 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4565 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4566 |
"t_correct tp \<Longrightarrow> trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4567 |
= rec_exec rec_conf [code tp, bl2wc (<lm>), n]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4568 |
and h: "t_correct tp" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4569 |
show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4570 |
"trpl_code (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp (Suc n)) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4571 |
rec_exec rec_conf [code tp, bl2wc (<lm>), Suc n]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4572 |
proof(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp n", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4573 |
simp only: tstep_red conf_lemma conf.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4574 |
fix a b c |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4575 |
assume g: "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp n = (a, b, c) " |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4576 |
hence "conf (code tp) (bl2wc (<lm>)) n= trpl_code (a, b, c)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4577 |
using ind h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4578 |
apply(simp add: conf_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4579 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4580 |
moreover hence |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4581 |
"trpl_code (tstep (a, b, c) tp) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4582 |
rec_exec rec_newconf [code tp, trpl_code (a, b, c)]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4583 |
apply(rule_tac rec_t_eq_step) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4584 |
using h g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4585 |
apply(simp add: s_keep) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4586 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4587 |
ultimately show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4588 |
"trpl_code (tstep (a, b, c) tp) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4589 |
newconf (code tp) (conf (code tp) (bl2wc (<lm>)) n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4590 |
by(simp add: newconf_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4591 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4592 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4593 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4594 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4595 |
lemma [simp]: "bl2wc (Bk\<^bsup>m\<^esup>) = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4596 |
apply(induct m) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4597 |
apply(simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4598 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4599 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4600 |
lemma [simp]: "bl2wc (Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>) = bl2wc (Oc\<^bsup>rs\<^esup>)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4601 |
apply(induct rs, simp, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4602 |
simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4603 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4604 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4605 |
lemma lg_power: "x > Suc 0 \<Longrightarrow> lg (x ^ rs) x = rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4606 |
proof(simp add: lg.simps, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4607 |
fix xa |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4608 |
assume h: "Suc 0 < x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4609 |
show "Max {ya. ya \<le> x ^ rs \<and> lgR [x ^ rs, x, ya]} = rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4610 |
apply(rule_tac Max_eqI, simp_all add: lgR.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4611 |
apply(simp add: h) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4612 |
using x_less_exp[of x rs] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4613 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4614 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4615 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4616 |
assume "\<not> Suc 0 < x ^ rs" "Suc 0 < x" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4617 |
thus "rs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4618 |
apply(case_tac "x ^ rs", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4619 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4620 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4621 |
assume "Suc 0 < x" "\<forall>xa. \<not> lgR [x ^ rs, x, xa]" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4622 |
thus "rs = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4623 |
apply(simp only:lgR.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4624 |
apply(erule_tac x = rs in allE, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4625 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4626 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4627 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4628 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4629 |
The following lemma relates execution of TMs with |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4630 |
the multi-step interpreter function @{text "rec_nonstop"}. Note, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4631 |
@{text "rec_nonstop"} is constructed using @{text "rec_conf"}. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4632 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4633 |
lemma nonstop_t_eq: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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>); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4635 |
turing_basic.t_correct tp; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4636 |
rs > 0\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4637 |
\<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4638 |
proof(simp add: nonstop_lemma nonstop.simps nstd.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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>)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4640 |
and tc_t: "turing_basic.t_correct tp" "rs > 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4641 |
have g: "rec_exec rec_conf [code tp, bl2wc (<lm>), stp] = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4642 |
trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4643 |
using rec_t_eq_steps[of tp l lm stp] tc_t h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4644 |
by(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4645 |
thus "\<not> NSTD (conf (code tp) (bl2wc (<lm>)) stp)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4646 |
proof(auto simp: NSTD.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4647 |
show "stat (conf (code tp) (bl2wc (<lm>)) stp) = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4648 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4649 |
by(auto simp: conf_lemma trpl_code.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4650 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4651 |
show "left (conf (code tp) (bl2wc (<lm>)) stp) = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4652 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4653 |
by(simp add: conf_lemma trpl_code.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4654 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4655 |
show "rght (conf (code tp) (bl2wc (<lm>)) stp) = |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4656 |
2 ^ lg (Suc (rght (conf (code tp) (bl2wc (<lm>)) stp))) 2 - Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4657 |
using g h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4658 |
proof(simp add: conf_lemma trpl_code.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4659 |
have "2 ^ lg (Suc (bl2wc (Oc\<^bsup>rs\<^esup>))) 2 = Suc (bl2wc (Oc\<^bsup>rs\<^esup>))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4660 |
apply(simp add: bl2wc.simps lg_power) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4661 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4662 |
thus "bl2wc (Oc\<^bsup>rs\<^esup>) = 2 ^ lg (Suc (bl2wc (Oc\<^bsup>rs\<^esup>))) 2 - Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4663 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4664 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4665 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4666 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4667 |
show "0 < rght (conf (code tp) (bl2wc (<lm>)) stp)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4668 |
using g h tc_t |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4669 |
apply(simp add: conf_lemma trpl_code.simps bl2wc.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4670 |
bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4671 |
apply(case_tac rs, simp, simp add: bl2nat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4672 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4673 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4674 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4675 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4676 |
lemma [simp]: "actn m 0 r = 4" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4677 |
by(simp add: actn.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4678 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4679 |
lemma [simp]: "newstat m 0 r = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4680 |
by(simp add: newstat.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4681 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4682 |
declare exp_def[simp del] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4683 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4684 |
lemma halt_least_step: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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>); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4686 |
turing_basic.t_correct tp; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4687 |
0<rs\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4688 |
\<exists> stp. (nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4689 |
(\<forall> stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp'))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4690 |
proof(induct stp, simp add: steps.simps, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4691 |
fix stp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4692 |
assume ind: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4694 |
\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4695 |
(\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4696 |
and h: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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>)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4698 |
"turing_basic.t_correct tp" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4699 |
"0 < rs" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4700 |
from h show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4701 |
"\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4702 |
\<and> (\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4703 |
proof(simp add: tstep_red, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4704 |
case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp", simp, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4705 |
case_tac a, simp add: tstep_0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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>)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4707 |
thus "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4708 |
(\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4709 |
apply(erule_tac ind) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4710 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4711 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4712 |
fix a b c nat |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4713 |
assume "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (a, b, c)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4714 |
"a = Suc nat" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4715 |
thus "\<exists>stp. nonstop (code tp) (bl2wc (<lm>)) stp = 0 \<and> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4716 |
(\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stp \<le> stp')" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4717 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4718 |
apply(rule_tac x = "Suc stp" in exI, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4719 |
apply(drule_tac nonstop_t_eq, simp_all add: nonstop_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4720 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4721 |
fix stp' |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4722 |
assume g:"steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (Suc nat, b, c)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4723 |
"nonstop (code tp) (bl2wc (<lm>)) stp' = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4724 |
thus "Suc stp \<le> stp'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4725 |
proof(case_tac "Suc stp \<le> stp'", simp, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4726 |
assume "\<not> Suc stp \<le> stp'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4727 |
hence "stp' \<le> stp" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4728 |
hence "\<not> isS0 (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp')" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4729 |
using g |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4730 |
apply(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp'",auto, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4731 |
simp add: isS0_def) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4732 |
apply(subgoal_tac "\<exists> n. stp = stp' + n", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4733 |
auto simp: steps_add steps_0) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4734 |
apply(rule_tac x = "stp - stp'" in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4735 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4736 |
hence "nonstop (code tp) (bl2wc (<lm>)) stp' = 1" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4737 |
proof(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp'", |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4738 |
simp add: isS0_def nonstop.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4739 |
fix a b c |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4740 |
assume k: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4741 |
"0 < a" "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp' = (a, b, c)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4742 |
thus " NSTD (conf (code tp) (bl2wc (<lm>)) stp')" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4743 |
using rec_t_eq_steps[of tp l lm stp'] h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4744 |
proof(simp add: conf_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4745 |
assume "trpl_code (a, b, c) = conf (code tp) (bl2wc (<lm>)) stp'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4746 |
moreover have "NSTD (trpl_code (a, b, c))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4747 |
using k |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4748 |
apply(auto simp: trpl_code.simps NSTD.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4749 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4750 |
ultimately show "NSTD (conf (code tp) (bl2wc (<lm>)) stp')" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4751 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4752 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4753 |
thus "False" using g by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4754 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4755 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4756 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4757 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4758 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4759 |
lemma conf_trpl_ex: "\<exists> p q r. conf m (bl2wc (<lm>)) stp = trpl p q r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4760 |
apply(induct stp, auto simp: conf.simps inpt.simps trpl.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4761 |
newconf.simps) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4762 |
apply(rule_tac x = 0 in exI, rule_tac x = 1 in exI, |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4763 |
rule_tac x = "bl2wc (<lm>)" in exI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4764 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4765 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4766 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4767 |
lemma nonstop_rgt_ex: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4768 |
"nonstop m (bl2wc (<lm>)) stpa = 0 \<Longrightarrow> \<exists> r. conf m (bl2wc (<lm>)) stpa = trpl 0 0 r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4769 |
apply(auto simp: nonstop.simps NSTD.simps split: if_splits) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4770 |
using conf_trpl_ex[of m lm stpa] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4771 |
apply(auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4772 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4773 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4774 |
lemma [elim]: "x > Suc 0 \<Longrightarrow> Max {u. x ^ u dvd x ^ r} = r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4775 |
proof(rule_tac Max_eqI) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4776 |
assume "x > Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4777 |
thus "finite {u. x ^ u dvd x ^ r}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4778 |
apply(rule_tac finite_power_dvd, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4779 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4780 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4781 |
fix y |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4782 |
assume "Suc 0 < x" "y \<in> {u. x ^ u dvd x ^ r}" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4783 |
thus "y \<le> r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4784 |
apply(case_tac "y\<le> r", simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4785 |
apply(subgoal_tac "\<exists> d. y = r + d") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4786 |
apply(auto simp: power_add) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4787 |
apply(rule_tac x = "y - r" in exI, simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4788 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4789 |
next |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4790 |
show "r \<in> {u. x ^ u dvd x ^ r}" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4791 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4792 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4793 |
lemma lo_power: "x > Suc 0 \<Longrightarrow> lo (x ^ r) x = r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4794 |
apply(auto simp: lo.simps loR.simps mod_dvd_simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4795 |
apply(case_tac "x^r", simp_all) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4796 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4797 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4798 |
lemma lo_rgt: "lo (trpl 0 0 r) (Pi 2) = r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4799 |
apply(simp add: trpl.simps lo_power) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4800 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4801 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4802 |
lemma conf_keep: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4803 |
"conf m lm stp = trpl 0 0 r \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4804 |
conf m lm (stp + n) = trpl 0 0 r" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4805 |
apply(induct n) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4806 |
apply(auto simp: conf.simps newconf.simps newleft.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4807 |
newrght.simps rght.simps lo_rgt) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4808 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4809 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4810 |
lemma halt_state_keep_steps_add: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4811 |
"\<lbrakk>nonstop m (bl2wc (<lm>)) stpa = 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4812 |
conf m (bl2wc (<lm>)) stpa = conf m (bl2wc (<lm>)) (stpa + n)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4813 |
apply(drule_tac nonstop_rgt_ex, auto simp: conf_keep) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4814 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4815 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4816 |
lemma halt_state_keep: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4817 |
"\<lbrakk>nonstop m (bl2wc (<lm>)) stpa = 0; nonstop m (bl2wc (<lm>)) stpb = 0\<rbrakk> \<Longrightarrow> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4818 |
conf m (bl2wc (<lm>)) stpa = conf m (bl2wc (<lm>)) stpb" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4819 |
apply(case_tac "stpa > stpb") |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4820 |
using halt_state_keep_steps_add[of m lm stpb "stpa - stpb"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4821 |
apply simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4822 |
using halt_state_keep_steps_add[of m lm stpa "stpb - stpa"] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4823 |
apply(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4824 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4825 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4826 |
text {* |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4827 |
The correntess of @{text "rec_F"} which relates the interpreter function @{text "rec_F"} with the |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4828 |
execution of of TMs. |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4829 |
*} |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4830 |
lemma F_t_halt_eq: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
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>); |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4832 |
turing_basic.t_correct tp; |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4833 |
0<rs\<rbrakk> |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4834 |
\<Longrightarrow> rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] (rs - Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4835 |
apply(frule_tac halt_least_step, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4836 |
apply(frule_tac nonstop_t_eq, auto simp: nonstop_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4837 |
using rec_t_eq_steps[of tp l lm stp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4838 |
apply(simp add: conf_lemma) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4839 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4840 |
fix stpa |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4841 |
assume h: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4842 |
"nonstop (code tp) (bl2wc (<lm>)) stpa = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4843 |
"\<forall>stp'. nonstop (code tp) (bl2wc (<lm>)) stp' = 0 \<longrightarrow> stpa \<le> stp'" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4844 |
"nonstop (code tp) (bl2wc (<lm>)) stp = 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4845 |
"trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) = conf (code tp) (bl2wc (<lm>)) stp" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4846 |
"steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4847 |
hence g1: "conf (code tp) (bl2wc (<lm>)) stpa = trpl_code (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4848 |
using halt_state_keep[of "code tp" lm stpa stp] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4849 |
by(simp) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4850 |
moreover have g2: |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4851 |
"rec_calc_rel rec_halt [code tp, (bl2wc (<lm>))] stpa" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4852 |
using h |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4853 |
apply(simp add: halt_lemma nonstop_lemma, auto) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4854 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4855 |
show |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4856 |
"rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] (rs - Suc 0)" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4857 |
proof - |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4858 |
have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4859 |
"rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4860 |
(valu (rght (conf (code tp) (bl2wc (<lm>)) stpa)))" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4861 |
apply(rule F_lemma) using g2 h by auto |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4862 |
moreover have |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4863 |
"valu (rght (conf (code tp) (bl2wc (<lm>)) stpa)) = rs - Suc 0" |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4864 |
using g1 |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4865 |
apply(simp add: valu.simps trpl_code.simps |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4866 |
bl2wc.simps bl2nat_append lg_power) |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4867 |
done |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4868 |
ultimately show "?thesis" by simp |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4869 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4870 |
qed |
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4871 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4872 |
|
1ce04eb1c8ad
Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff
changeset
|
4873 |
end |