|
1 (***************************************************************** |
|
2 |
|
3 Isabelle Tutorial |
|
4 ----------------- |
|
5 |
|
6 1st June 2009, Beijing |
|
7 |
|
8 *) |
|
9 |
|
10 theory Lec2 |
|
11 imports "Main" |
|
12 begin |
|
13 |
|
14 text {*********************************************************** |
|
15 |
|
16 Automatic proofs again: |
|
17 |
|
18 *} |
|
19 |
|
20 inductive |
|
21 even :: "nat \<Rightarrow> bool" |
|
22 where |
|
23 eZ[intro]: "even 0" |
|
24 | eSS[intro]: "even n \<Longrightarrow> even (Suc (Suc n))" |
|
25 |
|
26 text {* |
|
27 In the following two lemmas we have to indicate the induction, |
|
28 but the routine calculations can be done completely automatically |
|
29 by Isabelle's internal tools. |
|
30 *} |
|
31 |
|
32 lemma even_twice: |
|
33 shows "even (n + n)" |
|
34 by (induct n) (auto) |
|
35 |
|
36 lemma even_add: |
|
37 assumes a: "even n" |
|
38 and b: "even m" |
|
39 shows "even (n + m)" |
|
40 using a b by (induct) (auto) |
|
41 |
|
42 text {* |
|
43 In the next proof it was crucial that we introduced |
|
44 the equation |
|
45 |
|
46 (Suc (Suc n) * m) = (m + m) + (n * m) |
|
47 |
|
48 because this allowed us to use the induction hypothesis |
|
49 and the previous two lemmas. *} |
|
50 |
|
51 lemma even_mul: |
|
52 assumes a: "even n" |
|
53 shows "even (n * m)" |
|
54 using a |
|
55 proof (induct) |
|
56 case eZ |
|
57 show "even (0 * m)" by auto |
|
58 next |
|
59 case (eSS n) |
|
60 have as: "even n" by fact |
|
61 have ih: "even (n * m)" by fact |
|
62 have "(Suc (Suc n) * m) = (m + m) + (n * m)" by simp |
|
63 moreover |
|
64 have "even (m + m)" using even_twice by simp |
|
65 ultimately |
|
66 show "even (Suc (Suc n) * m)" using ih even_add by (simp only:) |
|
67 qed |
|
68 |
|
69 text {* |
|
70 This proof cannot be found by the internal tools, but |
|
71 Isabelle has an interface to external provers. This |
|
72 interface is called sledgehammer and it finds the |
|
73 proof more or less automatically. |
|
74 |
|
75 *} |
|
76 |
|
77 |
|
78 lemma even_mul_auto: |
|
79 assumes a: "even n" |
|
80 shows "even (n * m)" |
|
81 using a |
|
82 apply(induct) |
|
83 apply(metis eZ mult_is_0) |
|
84 apply(metis even_add even_twice mult_Suc_right nat_add_assoc nat_mult_commute) |
|
85 done |
|
86 |
|
87 text {* |
|
88 The disadvantage of such proofs is that you do |
|
89 not know why they are true. |
|
90 *} |
|
91 |
|
92 |
|
93 text {***************************************************************** |
|
94 |
|
95 Function Definitions |
|
96 -------------------- |
|
97 |
|
98 Many functions over datatypes can be defined by recursion on the |
|
99 structure. For this purpose, Isabelle provides "fun". To use it one needs |
|
100 to give a name for the function, its type, optionally some pretty-syntax |
|
101 and then some equations defining the function. Like in "inductive", |
|
102 "fun" expects that more than one such equation is separated by "|". |
|
103 |
|
104 The Fibonacci function: |
|
105 |
|
106 *} |
|
107 |
|
108 fun |
|
109 fib :: "nat \<Rightarrow> nat" |
|
110 where |
|
111 "fib 0 = 1" |
|
112 | "fib (Suc 0) = 1" |
|
113 | "fib (Suc (Suc n)) = fib n + fib (Suc n)" |
|
114 |
|
115 text {* The Ackermann function: *} |
|
116 |
|
117 fun |
|
118 ack :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
119 where |
|
120 "ack 0 m = Suc m" |
|
121 | "ack (Suc n) 0 = ack n (Suc 0)" |
|
122 | "ack (Suc n) (Suc m) = ack n (ack (Suc n) m)" |
|
123 |
|
124 text {* |
|
125 Non-terminating functions can lead to inconsistencies. |
|
126 *} |
|
127 |
|
128 lemma |
|
129 assumes a: "f x = f x + (1::nat)" |
|
130 shows "False" |
|
131 proof - |
|
132 from a have "0 = (1::nat)" by simp |
|
133 then show "False" by simp |
|
134 qed |
|
135 |
|
136 text {* |
|
137 If Isabelle cannot automatically prove termination, |
|
138 then you can give an explicit measure that establishes |
|
139 termination. For example a generalisation of the |
|
140 Fibonacci function to integers can be shown terminating |
|
141 |
|
142 *} |
|
143 |
|
144 function |
|
145 fib' :: "int \<Rightarrow> int" |
|
146 where |
|
147 "n< -1 \<Longrightarrow> fib' n = fib' (n + 2) - fib' (n + 1)" |
|
148 | "fib' -1 = (1::int)" |
|
149 | "fib' 0 = (0::int)" |
|
150 | "fib' 1 = (1::int)" |
|
151 | "n > 1 \<Longrightarrow> fib' n = fib' (n - 1) + fib' (n - 2)" |
|
152 by (atomize_elim, presburger) (auto) |
|
153 |
|
154 termination |
|
155 by (relation "measure (\<lambda>x. nat (\<bar>x\<bar>))") |
|
156 (simp_all add: zabs_def) |
|
157 |
|
158 |
|
159 text {***************************************************************** |
|
160 |
|
161 Datatypes |
|
162 --------- |
|
163 |
|
164 Datatypes can be defined in Isabelle as follows: we have to provide the name |
|
165 of the datatype and list its type-constructors. Each type-constructor needs |
|
166 to have the information about the types of its arguments, and optionally |
|
167 can also contain some information about pretty syntax. For example, we write |
|
168 "[]" for Nil and ":::" for Cons. |
|
169 |
|
170 *} |
|
171 |
|
172 datatype 'a mylist = |
|
173 MyNil ("[]") |
|
174 | MyCons "'a" "'a mylist" ("_ ::: _" 65) |
|
175 |
|
176 text {* |
|
177 You can easily define functions over the structure of datatypes. |
|
178 *} |
|
179 |
|
180 fun |
|
181 myappend :: "'a mylist \<Rightarrow> 'a mylist \<Rightarrow> 'a mylist" ("_ @@ _" 65) |
|
182 where |
|
183 "[] @@ xs = xs" |
|
184 | "(y:::ys) @@ xs = y:::(ys @@ xs)" |
|
185 |
|
186 fun |
|
187 myrev :: "'a mylist \<Rightarrow> 'a mylist" |
|
188 where |
|
189 "myrev [] = []" |
|
190 | "myrev (x:::xs) = (myrev xs) @@ (x:::[])" |
|
191 |
|
192 text {***************************************************************** |
|
193 |
|
194 Exercise: Prove the following fact about append and reversal. |
|
195 |
|
196 *} |
|
197 |
|
198 lemma myrev_append: |
|
199 shows "myrev (xs @@ ys) = (myrev ys) @@ (myrev xs)" |
|
200 proof (induct xs) |
|
201 case MyNil |
|
202 show "myrev ([] @@ ys) = myrev ys @@ myrev []" sorry |
|
203 next |
|
204 case (MyCons x xs) |
|
205 have ih: "myrev (xs @@ ys) = myrev ys @@ myrev xs" by fact |
|
206 |
|
207 show "myrev ((x:::xs) @@ ys) = myrev ys @@ myrev (x:::xs)" sorry |
|
208 qed |
|
209 |
|
210 |
|
211 |
|
212 |
|
213 |
|
214 |
|
215 |
|
216 |
|
217 |
|
218 |
|
219 |
|
220 |
|
221 |
|
222 |
|
223 |
|
224 |
|
225 text {* auxiliary lemmas *} |
|
226 |
|
227 lemma mynil_append[simp]: |
|
228 shows "xs @@ [] = xs" |
|
229 by (induct xs) (auto) |
|
230 |
|
231 lemma myappend_assoc[simp]: |
|
232 "(xs @@ ys) @@ zs = xs @@ (ys @@ zs)" |
|
233 by (induct xs) (auto) |
|
234 |
|
235 lemma |
|
236 shows "myrev (xs @@ ys) = (myrev ys) @@ (myrev xs)" |
|
237 by (induct xs) (auto) |
|
238 |
|
239 text {***************************************************************** |
|
240 |
|
241 A WHILE Language: |
|
242 |
|
243 Arithmetic Expressions, Boolean Expressions, Commands |
|
244 ----------------------------------------------------- |
|
245 |
|
246 *} |
|
247 |
|
248 types memory = "nat \<Rightarrow> nat" |
|
249 |
|
250 datatype aexp = |
|
251 C nat |
|
252 | X nat |
|
253 | Op1 "nat \<Rightarrow> nat" aexp |
|
254 | Op2 "nat \<Rightarrow> nat \<Rightarrow> nat" aexp aexp |
|
255 |
|
256 datatype bexp = |
|
257 TRUE |
|
258 | FALSE |
|
259 | ROp "nat \<Rightarrow> nat \<Rightarrow> bool" aexp aexp |
|
260 | NOT bexp |
|
261 | AND bexp bexp |
|
262 | OR bexp bexp |
|
263 |
|
264 datatype cmd = |
|
265 SKIP |
|
266 | ASSIGN nat aexp ("_ ::= _ " 60) |
|
267 | SEQ cmd cmd ("_; _" [60, 60] 10) |
|
268 | COND bexp cmd cmd ("IF _ THEN _ ELSE _" 60) |
|
269 | WHILE bexp cmd ("WHILE _ DO _" 60) |
|
270 |
|
271 text {***************************************************************** |
|
272 |
|
273 Big-step Evaluation Semantics |
|
274 ----------------------------- |
|
275 |
|
276 *} |
|
277 |
|
278 inductive |
|
279 evala :: "aexp \<Rightarrow> memory \<Rightarrow> nat \<Rightarrow> bool" ("'(_,_') \<longrightarrow>a _" [100,100] 50) |
|
280 where |
|
281 C: "(C n, m) \<longrightarrow>a n" |
|
282 | X: "(X i, m) \<longrightarrow>a m i" |
|
283 | Op1: "(e, m) \<longrightarrow>a n \<Longrightarrow> (Op1 f e, m) \<longrightarrow>a (f n)" |
|
284 | Op2: "\<lbrakk>(e0, m) \<longrightarrow>a n0; (e1, m) \<longrightarrow>a n1\<rbrakk> \<Longrightarrow> (Op2 f e0 e1, m) \<longrightarrow>a f n0 n1" |
|
285 |
|
286 inductive |
|
287 evalb :: "bexp \<Rightarrow> memory \<Rightarrow> bool \<Rightarrow> bool" ("'(_,_') \<longrightarrow>b _" [100,100] 50) |
|
288 where |
|
289 TRUE: "(TRUE, m) \<longrightarrow>b True" |
|
290 | FALSE: "(FALSE, m) \<longrightarrow>b False" |
|
291 | ROp: "\<lbrakk>(e1, m) \<longrightarrow>a n1; (e2, m) \<longrightarrow>a n2\<rbrakk> \<Longrightarrow> (ROp f e1 e2, m) \<longrightarrow>b (f n1 n2)" |
|
292 | NOT: "(e, m) \<longrightarrow>b b \<Longrightarrow> (NOT e, m) \<longrightarrow>b (Not b)" |
|
293 | AND: "\<lbrakk>(e1, m) \<longrightarrow>b b1; (e2, m) \<longrightarrow>b b2\<rbrakk> \<Longrightarrow> (AND e1 e2, m) \<longrightarrow>b (b1 \<and> b2)" |
|
294 | OR: "\<lbrakk>(e1, m) \<longrightarrow>b b1; (e2, m) \<longrightarrow>b b2\<rbrakk> \<Longrightarrow> (OR e1 e2, m) \<longrightarrow>b (b1 \<or> b2)" |
|
295 |
|
296 inductive |
|
297 evalc :: "cmd \<Rightarrow> memory \<Rightarrow> memory \<Rightarrow> bool" ("'(_,_') \<longrightarrow>c _" [0,0,60] 60) |
|
298 where |
|
299 SKIP: "(SKIP, m) \<longrightarrow>c m" |
|
300 | ASSIGN: "(a, m) \<longrightarrow>a n \<Longrightarrow> (x::=a, m) \<longrightarrow>c m(x:=n)" |
|
301 | SEQ: "\<lbrakk>(c0, m) \<longrightarrow>c m'; (c1, m') \<longrightarrow>c m''\<rbrakk> \<Longrightarrow> (c0; c1, m) \<longrightarrow>c m''" |
|
302 | IFTrue: "\<lbrakk>(e, m) \<longrightarrow>b True; (c0, m) \<longrightarrow>c m'\<rbrakk> \<Longrightarrow> (IF e THEN c0 ELSE c1, m) \<longrightarrow>c m'" |
|
303 | IFFalse: "\<lbrakk>(e, m) \<longrightarrow>b False; (c1, m) \<longrightarrow>c m'\<rbrakk> \<Longrightarrow> (IF e THEN c0 ELSE c1, m) \<longrightarrow>c m'" |
|
304 | WHILEFalse: "(e, m) \<longrightarrow>b False \<Longrightarrow> (WHILE e DO c,m) \<longrightarrow>c m" |
|
305 | WHILETrue: "\<lbrakk>(e, m) \<longrightarrow>b True; (c,m) \<longrightarrow>c m'; (WHILE e DO c, m') \<longrightarrow>c m''\<rbrakk> |
|
306 \<Longrightarrow> (WHILE e DO c, m) \<longrightarrow>c m''" |
|
307 |
|
308 lemmas evala.intros[intro] evalb.intros[intro] evalc.intros[intro] |
|
309 |
|
310 text {***************************************************************** |
|
311 |
|
312 Machine Instructions |
|
313 -------------------- |
|
314 |
|
315 *} |
|
316 |
|
317 datatype instr = |
|
318 JPFZ "nat" (* JUMP forward n steps, if top of stack is False *) |
|
319 | JPB "nat" (* JUMP backward n steps *) |
|
320 | FETCH "nat" (* move memory to top of stack *) |
|
321 | STORE "nat" (* move top of stack to memory *) |
|
322 | PUSH "nat" (* push to stack *) |
|
323 | OPU "nat \<Rightarrow> nat" (* pop one from stack and apply f *) |
|
324 | OPB "nat \<Rightarrow> nat \<Rightarrow> nat" (* pop two from stack and apply f *) |
|
325 |
|
326 text {***************************************************************** |
|
327 |
|
328 Booleans 'as' Zero and One |
|
329 -------------------------- |
|
330 |
|
331 *} |
|
332 |
|
333 fun |
|
334 WRAP::"bool\<Rightarrow>nat" |
|
335 where |
|
336 "WRAP True = Suc 0" |
|
337 | "WRAP False = 0" |
|
338 |
|
339 fun |
|
340 MNot::"nat \<Rightarrow> nat" |
|
341 where |
|
342 "MNot 0 = 1" |
|
343 | "MNot (Suc 0) = 0" |
|
344 |
|
345 fun |
|
346 MAnd::"nat \<Rightarrow> nat \<Rightarrow> nat" |
|
347 where |
|
348 "MAnd 0 x = 0" |
|
349 | "MAnd x 0 = 0" |
|
350 | "MAnd (Suc 0) (Suc 0) = 1" |
|
351 |
|
352 fun |
|
353 MOr::"nat \<Rightarrow> nat \<Rightarrow> nat" |
|
354 where |
|
355 "MOr 0 0 = 0" |
|
356 | "MOr (Suc 0) x = 1" |
|
357 | "MOr x (Suc 0) = 1" |
|
358 |
|
359 lemma MNot_WRAP: |
|
360 "MNot (WRAP b) = WRAP (Not b)" |
|
361 by (cases b) (auto) |
|
362 |
|
363 lemma MAnd_WRAP: |
|
364 "MAnd (WRAP b1) (WRAP b2) = WRAP (b1 \<and> b2)" |
|
365 by (cases b1, auto) (cases b2, simp_all) |
|
366 |
|
367 |
|
368 lemma MOr_WRAP: |
|
369 "MOr (WRAP b1) (WRAP b2) = WRAP (b1 \<or> b2)" |
|
370 by (cases b1, auto) (cases b2, simp_all) |
|
371 |
|
372 lemmas WRAP_lems = MNot_WRAP MAnd_WRAP MOr_WRAP |
|
373 |
|
374 text {***************************************************************** |
|
375 |
|
376 Compiler Functions |
|
377 ------------------ |
|
378 |
|
379 *} |
|
380 |
|
381 fun |
|
382 compa |
|
383 where |
|
384 "compa (C n) = [PUSH n]" |
|
385 | "compa (X l) = [FETCH l]" |
|
386 | "compa (Op1 f e) = (compa e) @ [OPU f]" |
|
387 | "compa (Op2 f e1 e2) = (compa e1) @ (compa e2) @ [OPB f]" |
|
388 |
|
389 fun |
|
390 compb |
|
391 where |
|
392 "compb (TRUE) = [PUSH 1]" |
|
393 | "compb (FALSE) = [PUSH 0]" |
|
394 | "compb (ROp f e1 e2) = (compa e1) @ (compa e2) @ [OPB (\<lambda>x y. WRAP (f x y))]" |
|
395 | "compb (NOT e) = (compb e) @ [OPU MNot]" |
|
396 | "compb (AND e1 e2) = (compb e1) @ (compb e2) @ [OPB MAnd]" |
|
397 | "compb (OR e1 e2) = (compb e1) @ (compb e2) @ [OPB MOr]" |
|
398 |
|
399 fun |
|
400 compc :: "cmd \<Rightarrow> instr list" |
|
401 where |
|
402 "compc SKIP = []" |
|
403 | "compc (x::=a) = (compa a) @ [STORE x]" |
|
404 | "compc (c1;c2) = compc c1 @ compc c2" |
|
405 | "compc (IF b THEN c1 ELSE c2) = |
|
406 (compb b) @ [JPFZ (length(compc c1) + 2)] @ compc c1 @ |
|
407 [PUSH 0, JPFZ (length(compc c2))] @ compc c2" |
|
408 | "compc (WHILE b DO c) = |
|
409 (compb b) @ |
|
410 [JPFZ (length(compc c) + 1)] @ compc c @ |
|
411 [JPB (length(compc c) + length(compb b)+1)]" |
|
412 |
|
413 value "compc SKIP" |
|
414 value "compc (SKIP; SKIP)" |
|
415 value "compa (Op2 (op +) (C 2) (C 3))" |
|
416 value "compb (ROp (op <) (C 2) (C 3))" |
|
417 value "compc (IF (ROp (op <) (C 2) (C 3)) THEN SKIP ELSE SKIP)" |
|
418 value "compc (WHILE (ROp (op <) (C 2) (C 3)) DO SKIP)" |
|
419 value "compc (x::= C 0; WHILE (ROp (op <) (X x) (C 3)) DO (x::= (Op2 (op +) (X x) (C 1))))" |
|
420 |
|
421 text {***************************************************************** |
|
422 |
|
423 Machine Execution |
|
424 ----------------- |
|
425 |
|
426 *} |
|
427 |
|
428 types instrs = "instr list" |
|
429 types stack = "nat list" |
|
430 |
|
431 abbreviation |
|
432 "rtake i xs \<equiv> rev (take i xs)" |
|
433 |
|
434 inductive |
|
435 step :: "instrs \<Rightarrow> instrs \<Rightarrow> stack \<Rightarrow> memory \<Rightarrow> |
|
436 instrs \<Rightarrow> instrs \<Rightarrow> stack \<Rightarrow> memory \<Rightarrow> bool" ("'(_,_,_,_') \<longrightarrow>m '(_,_,_,_')") |
|
437 where |
|
438 "(PUSH n#p, q, s, m) \<longrightarrow>m (p, PUSH n#q, n#s, m)" |
|
439 | "(FETCH i#p, q, s, m) \<longrightarrow>m (p, FETCH i#q, m i#s, m)" |
|
440 | "(OPU f#p, q, n#s, m) \<longrightarrow>m (p, OPU f#q, f n#s, m)" |
|
441 | "(OPB f#p, q, n1#n2#s, m) \<longrightarrow>m (p, OPB f#q, f n2 n1#s, m)" |
|
442 | "(STORE i#p, q, n#s, m) \<longrightarrow>m (p, STORE i#q, s, m(i:=n))" |
|
443 | "(JPFZ i#p, q, Suc 0#s, m) \<longrightarrow>m (p, JPFZ i#q, s, m)" |
|
444 | "i\<le>length p\<Longrightarrow>(JPFZ i#p, q, 0#s, m) \<longrightarrow>m (drop i p, (rtake i p) @ (JPFZ i#q), s, m)" |
|
445 | "i\<le>length q\<Longrightarrow>(JPB i#p, q, s, m) \<longrightarrow>m ((rtake i q) @ (JPB i#p), drop i q, s, m)" |
|
446 |
|
447 lemmas step.intros[intro] |
|
448 |
|
449 inductive_cases step_elim[elim]: |
|
450 "(p,q,s,m) \<longrightarrow>m (p',q',s',m')" |
|
451 |
|
452 lemma exec_simp: |
|
453 shows "(instr#p, q, s, m) \<longrightarrow>m (p', q', s', m') = |
|
454 (case instr of |
|
455 PUSH n \<Rightarrow> (p' = p \<and> q' = instr#q \<and> s' = n#s \<and> m' = m) |
|
456 | FETCH i \<Rightarrow> (p' = p \<and> q' = instr#q \<and> s' = m i#s \<and> m' = m) |
|
457 | OPU f \<Rightarrow> (\<exists>n si. p' = p \<and> q' = instr#q \<and> s = n#si \<and> s' = f n#si \<and> m' = m) |
|
458 | OPB f \<Rightarrow> (\<exists>n1 n2 si. p' = p \<and> q' = instr#q \<and> s = n1#n2#si \<and> s' = f n2 n1#si \<and> m' = m) |
|
459 | STORE i \<Rightarrow> (\<exists>n si. p' = p \<and> q' = instr#q \<and> s = n#si \<and> s' = si \<and> m' = m(i:=n)) |
|
460 | JPFZ i \<Rightarrow> (\<exists>si. (p' = p \<and> q' = instr#q \<and> s = Suc 0#si \<and> s' = si \<and> m' = m) \<or> |
|
461 (i\<le> length p \<and> p' = (drop i p) \<and> q' = (rev (take i p))@(instr#q) |
|
462 \<and> s = 0#si \<and> s' = si \<and> m' = m )) |
|
463 | JPB i \<Rightarrow> (i \<le> length q \<and> p' = (rtake i q)@(JPB i#p) \<and> q' = drop i q \<and> s' = s \<and> m' = m))" |
|
464 by (auto split: instr.split_asm split_if_asm) |
|
465 |
|
466 inductive |
|
467 steps :: "instrs \<Rightarrow> instrs \<Rightarrow> stack \<Rightarrow> memory \<Rightarrow> |
|
468 instrs \<Rightarrow> instrs \<Rightarrow> stack \<Rightarrow> memory \<Rightarrow> bool" ("'(_,_,_,_') \<longrightarrow>m* '(_,_,_,_')") |
|
469 where |
|
470 refl: "(p,q,s,m) \<longrightarrow>m* (p,q,s,m)" |
|
471 | step: "\<lbrakk>(p1,q1,s1,m1) \<longrightarrow>m (p2,q2,s2,m2); (p2,q2,s2,m2) \<longrightarrow>m* (p3,q3,s3,m3)\<rbrakk> |
|
472 \<Longrightarrow> (p1,q1,s1,m1) \<longrightarrow>m* (p3,q3,s3,m3)" |
|
473 |
|
474 lemmas steps.step[intro] |
|
475 lemmas steps.refl[simp] |
|
476 |
|
477 inductive_cases steps_elim[elim]: |
|
478 "(p,q,s,m) \<longrightarrow>m* (p',q',s',m')" |
|
479 |
|
480 lemma steps_trans: |
|
481 assumes a: "(p1,q1,s1,m1) \<longrightarrow>m* (p2,q2,s2,m2)" |
|
482 and b: "(p2,q2,s2,m2) \<longrightarrow>m* (p3,q3,s3,m3)" |
|
483 shows "(p1,q1,s1,m1) \<longrightarrow>m* (p3,q3,s3,m3)" |
|
484 using a b by (induct) (auto) |
|
485 |
|
486 lemma steps_simp: |
|
487 shows "(i#p,q,s,m) \<longrightarrow>m* (p',q',s',m') = |
|
488 ((i#p,q,s,m) = (p',q',s',m') \<or> |
|
489 (\<exists>pi qi si mi. (i#p,q,s,m) \<longrightarrow>m (pi,qi,si,mi) \<and> (pi,qi,si,mi) \<longrightarrow>m* (p',q',s',m')))" |
|
490 by (auto) |
|
491 |
|
492 lemma compa_first_attempt: |
|
493 assumes a: "(e, m) \<longrightarrow>a n" |
|
494 shows "(compa e, [], [], m) \<longrightarrow>m* ([], rev (compa e), [n], m)" |
|
495 using a |
|
496 proof (induct) |
|
497 case (C n m) |
|
498 show "(compa (C n),[],[],m) \<longrightarrow>m* ([],rev (compa (C n)),[n],m)" |
|
499 sorry |
|
500 next |
|
501 case (X i m) |
|
502 show "(compa (X i),[],[],m) \<longrightarrow>m* ([],rev (compa (X i)),[m i],m)" |
|
503 sorry |
|
504 next |
|
505 case (Op1 e m n f) |
|
506 have as: "(e, m) \<longrightarrow>a n" by fact |
|
507 have ih: "(compa e,[],[],m) \<longrightarrow>m* ([],rev (compa e),[n],m)" by fact |
|
508 show "(compa (Op1 f e),[],[],m) \<longrightarrow>m* ([],rev (compa (Op1 f e)),[f n],m)" |
|
509 using ih sorry |
|
510 next |
|
511 case (Op2 e0 m n0 e1 n1 f) |
|
512 have as1: "(e0, m) \<longrightarrow>a n0" by fact |
|
513 have as2: "(e1, m) \<longrightarrow>a n1" by fact |
|
514 have ih1: "(compa e0,[],[],m) \<longrightarrow>m* ([],rev (compa e0),[n0],m)" by fact |
|
515 have ih2: "(compa e1,[],[],m) \<longrightarrow>m* ([],rev (compa e1),[n1],m)" by fact |
|
516 show "(compa (Op2 f e0 e1),[],[],m) \<longrightarrow>m* ([],rev (compa (Op2 f e0 e1)),[f n0 n1],m)" |
|
517 using ih1[THEN steps_trans] ih2[THEN steps_trans] |
|
518 sorry |
|
519 qed |
|
520 |
|
521 |
|
522 |
|
523 |
|
524 |
|
525 |
|
526 |
|
527 |
|
528 |
|
529 |
|
530 |
|
531 |
|
532 |
|
533 |
|
534 |
|
535 lemma compa_aux: |
|
536 assumes a: "(e, m) \<longrightarrow>a n" |
|
537 shows "(compa e@p, q, s, m) \<longrightarrow>m* (p, rev (compa e)@q, n#s, m)" |
|
538 using a |
|
539 proof (induct arbitrary: p q s) |
|
540 case (C n m p q s) |
|
541 show "(compa (C n)@p,q,s,m) \<longrightarrow>m* (p,rev (compa (C n))@q,n#s,m)" |
|
542 by (simp add: steps_simp exec_simp) |
|
543 next |
|
544 case (X i m p q s) |
|
545 show "(compa (X i)@p,q,s,m) \<longrightarrow>m* (p,rev (compa (X i))@q,m i#s,m)" |
|
546 by (simp add: steps_simp exec_simp) |
|
547 next |
|
548 case (Op1 e m n f p q s) |
|
549 have as: "(e, m) \<longrightarrow>a n" by fact |
|
550 have ih: "\<And>p q s. (compa e@p,q,s,m) \<longrightarrow>m* (p,rev (compa e)@q,n#s,m)" by fact |
|
551 show "(compa (Op1 f e)@p,q,s,m) \<longrightarrow>m* (p,rev (compa (Op1 f e))@q,f n#s,m)" |
|
552 using ih[THEN steps_trans] by (simp add: steps_simp exec_simp) |
|
553 next |
|
554 case (Op2 e0 m n0 e1 n1 f p q s) |
|
555 have as1: "(e0, m) \<longrightarrow>a n0" by fact |
|
556 have as2: "(e1, m) \<longrightarrow>a n1" by fact |
|
557 have ih1: "\<And>p q s. (compa e0@p,q,s,m) \<longrightarrow>m* (p,rev (compa e0)@q,n0#s,m)" by fact |
|
558 have ih2: "\<And>p q s. (compa e1@p,q,s,m) \<longrightarrow>m* (p,rev (compa e1)@q,n1#s,m)" by fact |
|
559 show "(compa (Op2 f e0 e1)@p,q,s,m) \<longrightarrow>m* (p,rev (compa (Op2 f e0 e1))@q,f n0 n1#s,m)" |
|
560 using ih1[THEN steps_trans] ih2[THEN steps_trans] |
|
561 by (simp add: steps_simp exec_simp) |
|
562 qed |
|
563 |
|
564 |
|
565 text {* |
|
566 After the fact I constructed the detailed proof, I |
|
567 could guide Isabelle to find the proof automatically. |
|
568 |
|
569 *B U T T H A T I S C H E A T I N G!!!* |
|
570 *} |
|
571 |
|
572 lemma compa_aux_cheating: |
|
573 assumes a: "(e,m) \<longrightarrow>a n" |
|
574 shows "(compa e@p,q,s,m) \<longrightarrow>m* (p,rev (compa e)@q,n#s,m)" |
|
575 using a |
|
576 by (induct arbitrary: p q s) |
|
577 (force intro: steps_trans simp add: steps_simp exec_simp)+ |
|
578 |
|
579 |
|
580 lemma compb_aux_cheating: |
|
581 assumes a: "(e,m) \<longrightarrow>b b" |
|
582 shows "(compb e@p,q,s,m) \<longrightarrow>m* (p,rev (compb e)@q,(WRAP b)#s,m)" |
|
583 using a |
|
584 by (induct arbitrary: p q s) |
|
585 (force intro: compa_aux steps_trans |
|
586 simp add: steps_simp exec_simp WRAP_lems)+ |
|
587 |
|
588 text {* |
|
589 |
|
590 Exercise: After you know the lemma is provable, |
|
591 try to give all details. |
|
592 |
|
593 *} |
|
594 |
|
595 lemma compb_aux: |
|
596 assumes a: "(e, m) \<longrightarrow>b b" |
|
597 shows "(compb e@p, q, s, m) \<longrightarrow>m* (p, rev (compb e)@q, WRAP b#s, m)" |
|
598 using a |
|
599 proof (induct arbitrary: p q s) |
|
600 case (TRUE m p q s) |
|
601 show "(compb TRUE@p,q,s,m) \<longrightarrow>m* (p,rev (compb TRUE)@q,WRAP True#s,m)" |
|
602 sorry |
|
603 next |
|
604 case (FALSE m p q s) |
|
605 show "(compb FALSE@p,q,s,m) \<longrightarrow>m* (p,rev (compb FALSE)@q,WRAP False#s,m)" |
|
606 sorry |
|
607 next |
|
608 case (ROp e1 m n1 e2 n2 f p q s) |
|
609 have as1: "(e1, m) \<longrightarrow>a n1" by fact |
|
610 have as2: "(e2, m) \<longrightarrow>a n2" by fact |
|
611 show "(compb (ROp f e1 e2)@p,q,s,m) |
|
612 \<longrightarrow>m* (p,rev (compb (ROp f e1 e2))@q,WRAP (f n1 n2)#s,m)" |
|
613 using as1 as2 sorry |
|
614 next |
|
615 case (NOT e m b p q s) |
|
616 have ih: "\<And>p q s. (compb e@p,q,s,m) \<longrightarrow>m* (p,rev (compb e)@q,WRAP b#s,m)" by fact |
|
617 show "(compb (NOT e)@p,q,s,m) \<longrightarrow>m* (p,rev (compb (NOT e))@q,WRAP (\<not>b)#s,m)" |
|
618 using ih[THEN steps_trans] |
|
619 sorry |
|
620 next |
|
621 case (AND e1 m b1 e2 b2 p q s) |
|
622 have ih1: "\<And>p q s. (compb e1@p,q,s,m) \<longrightarrow>m* (p,rev (compb e1)@q,WRAP b1#s,m)" by fact |
|
623 have ih2: "\<And>p q s. (compb e2@p,q,s,m) \<longrightarrow>m* (p,rev (compb e2)@q,WRAP b2#s,m)" by fact |
|
624 show "(compb (AND e1 e2)@p,q,s,m) \<longrightarrow>m* (p,rev (compb (AND e1 e2))@q,WRAP (b1 \<and> b2)#s,m)" |
|
625 using ih1[THEN steps_trans] ih2[THEN steps_trans] |
|
626 sorry |
|
627 next |
|
628 case (OR e1 m b1 e2 b2 p q s) |
|
629 have ih1: "\<And>p q s. (compb e1@p,q,s,m) \<longrightarrow>m* (p,rev (compb e1)@q,WRAP b1#s,m)" by fact |
|
630 have ih2: "\<And>p q s. (compb e2@p,q,s,m) \<longrightarrow>m* (p,rev (compb e2)@q,WRAP b2#s,m)" by fact |
|
631 show "(compb (OR e1 e2)@p,q,s,m) \<longrightarrow>m* (p,rev (compb (OR e1 e2))@q,WRAP (b1 \<or> b2)#s,m)" |
|
632 using ih1[THEN steps_trans] ih2[THEN steps_trans] |
|
633 sorry |
|
634 qed |
|
635 |
|
636 text {* |
|
637 |
|
638 Also the detailed proof of this lemma is left as an exercise. |
|
639 |
|
640 *} |
|
641 |
|
642 lemma compc_aux_cheating: |
|
643 assumes a: "(c,m) \<longrightarrow>c m'" |
|
644 shows "(compc c@p,q,s,m) \<longrightarrow>m* (p,rev (compc c)@q,s,m')" |
|
645 using a |
|
646 by (induct arbitrary: p q) |
|
647 (force intro: compa_aux compb_aux_cheating steps_trans |
|
648 simp add: steps_simp exec_simp)+ |
|
649 |
|
650 end |
|
651 |
|
652 |
|
653 |
|
654 |
|
655 |
|
656 |
|
657 |
|
658 |
|
659 |
|
660 |