|
1 (***************************************************************** |
|
2 |
|
3 Nominal Isabelle Tutorial |
|
4 ------------------------- |
|
5 |
|
6 11th August 2008, Sydney |
|
7 |
|
8 This file contains most of the material that will be covered in the |
|
9 tutorial. The file can be "stepped through"; though it contains much |
|
10 commented code (purple text). |
|
11 |
|
12 *) |
|
13 |
|
14 (***************************************************************** |
|
15 |
|
16 Proof General |
|
17 ------------- |
|
18 |
|
19 Proof General is the front end for Isabelle. |
|
20 |
|
21 To run Nominal Isabelle proof-scripts you must have HOL-Nominal enabled in |
|
22 the menu Isabelle -> Logics. You also need to enable X-Symbols in the menu |
|
23 Proof-General -> Options (make sure to save this option once enabled). |
|
24 |
|
25 Proof General "paints" blue the part of the proof-script that has already |
|
26 been processed by Isabelle. You can advance or retract the "frontier" of |
|
27 this processed part using the "Next" and "Undo" buttons in the |
|
28 menubar. "Goto" will process everything up to the current cursor position, |
|
29 or retract if the cursor is inside the blue part. The key-combination |
|
30 control-c control-return is a short-cut for the "Goto"-operation. |
|
31 |
|
32 Proof General gives feedback inside the "Response" and "Goals" buffers. |
|
33 The response buffer will contain warning messages (in yellow) and |
|
34 error messages (in red). Warning messages can generally be ignored. Error |
|
35 messages must be dealt with in order to advance the proof script. The goal |
|
36 buffer shows which goals need to be proved. The sole idea of interactive |
|
37 theorem proving is to get the message "No subgoals." ;o) |
|
38 |
|
39 *) |
|
40 |
|
41 (***************************************************************** |
|
42 |
|
43 X-Symbols |
|
44 --------- |
|
45 |
|
46 X-symbols provide a nice way to input non-ascii characters. For example: |
|
47 |
|
48 \<forall>, \<exists>, \<Down>, \<sharp>, \<And>, \<Gamma>, \<times>, \<noteq>, \<in>, \<subseteq>, \<dots> |
|
49 |
|
50 They need to be input via the combination \<name-of-x-symbol> |
|
51 where name-of-x-symbol coincides with the usual latex name of |
|
52 that symbol. |
|
53 |
|
54 However, there are some handy short-cuts for frequently used X-symbols. |
|
55 For example |
|
56 |
|
57 [| \<dots> \<lbrakk> |
|
58 |] \<dots> \<rbrakk> |
|
59 ==> \<dots> \<Longrightarrow> |
|
60 => \<dots> \<Rightarrow> |
|
61 --> \<dots> \<longrightarrow> |
|
62 /\ \<dots> \<and> |
|
63 \/ \<dots> \<or> |
|
64 |-> \<dots> \<mapsto> |
|
65 =_ \<dots> \<equiv> |
|
66 |
|
67 *) |
|
68 |
|
69 (***************************************************************** |
|
70 |
|
71 Theories |
|
72 -------- |
|
73 |
|
74 Every Isabelle proof-script needs to have a name and must import |
|
75 some pre-existing theory. For Nominal Isabelle proof-scripts this will |
|
76 normally be the theory Nominal, but we use here the theory Lambda.thy, |
|
77 which extends Nominal with a definition for lambda-terms and capture- |
|
78 avoiding substitution. |
|
79 |
|
80 BTW, the Nominal theory builds directly on Isabelle/HOL and extends it |
|
81 only with some definitions and some reasoning infrastructure. It does not |
|
82 add any new axiom to Isabelle/HOL. So you can trust what you are doing. ;o) |
|
83 |
|
84 *) |
|
85 |
|
86 theory CK_Machine |
|
87 imports "Lambda" |
|
88 begin |
|
89 |
|
90 text {***************************************************************** |
|
91 |
|
92 Types |
|
93 ----- |
|
94 |
|
95 Isabelle is based, roughly, on the theory of simple types including some |
|
96 polymorphism. It also includes some overloading, which means that sometimes |
|
97 explicit type annotations need to be given. |
|
98 |
|
99 - Base types include: nat, bool, string, lam |
|
100 |
|
101 - Type formers include: 'a list, ('a\<times>'b), 'c set |
|
102 |
|
103 - Type variables are written like in ML with an apostrophe: 'a, 'b, \<dots> |
|
104 |
|
105 Types known to Isabelle can be queried using: |
|
106 |
|
107 *} |
|
108 |
|
109 typ nat |
|
110 typ bool |
|
111 typ lam (* the type for alpha-equated lambda-terms *) |
|
112 typ "('a \<times> 'b)" (* product type *) |
|
113 typ "'c set" (* set type *) |
|
114 typ "nat \<Rightarrow> bool" (* the type for functions from nat to bool *) |
|
115 |
|
116 (* These give errors: *) |
|
117 (*typ boolean *) |
|
118 (*typ set*) |
|
119 |
|
120 |
|
121 text {***************************************************************** |
|
122 |
|
123 Terms |
|
124 ----- |
|
125 |
|
126 Every term in Isabelle needs to be well-typed (however they can have polymorphic |
|
127 type). Whether a term is accepted can be queried using: |
|
128 |
|
129 *} |
|
130 |
|
131 term c (* a variable of polymorphic type *) |
|
132 term "1::nat" (* the constant 1 in natural numbers *) |
|
133 term 1 (* the constant 1 with polymorphic type *) |
|
134 term "{1, 2, 3::nat}" (* the set containing natural numbers 1, 2 and 3 *) |
|
135 term "[1, 2, 3]" (* the list containing the polymorphic numbers 1, 2 and 3 *) |
|
136 term "Lam [x].(Var x)" (* a lambda-term *) |
|
137 term "App t1 t2" (* another lambda-term *) |
|
138 |
|
139 text {* |
|
140 Isabelle provides some useful colour feedback about what are constants (in black), |
|
141 free variables (in blue) and bound variables (in green). *} |
|
142 |
|
143 term "True" (* a constant that is defined in HOL *) |
|
144 term "true" (* not recognised as a constant, therefore it is interpreted as a variable *) |
|
145 term "\<forall>x. P x" (* x is bound, P is free *) |
|
146 |
|
147 text {* Every formula in Isabelle needs to have type bool *} |
|
148 |
|
149 term "True" |
|
150 term "True \<and> False" |
|
151 term "{1,2,3} = {3,2,1}" |
|
152 term "\<forall>x. P x" |
|
153 term "A \<longrightarrow> B" |
|
154 |
|
155 text {* |
|
156 When working with Isabelle, one is confronted with an object logic (that is HOL) |
|
157 and Isabelle's meta-logic (called Pure). Sometimes one has to pay attention |
|
158 to this fact. *} |
|
159 |
|
160 term "A \<longrightarrow> B" (* versus *) |
|
161 term "A \<Longrightarrow> B" |
|
162 |
|
163 term "\<forall>x. P x" (* versus *) |
|
164 term "\<And>x. P x" |
|
165 |
|
166 term "A \<Longrightarrow> B \<Longrightarrow> C" (* is synonymous with *) |
|
167 term "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> C" |
|
168 |
|
169 |
|
170 text {***************************************************************** |
|
171 |
|
172 Inductive Definitions: The Evaluation Judgement and the Value Predicate |
|
173 ----------------------------------------------------------------------- |
|
174 |
|
175 Inductive definitions start with the keyword "inductive" and a predicate name. |
|
176 One also needs to provide a type for the predicate. Clauses of the inductive |
|
177 predicate are introduced by "where" and more than two clauses need to be |
|
178 separated by "|". Optionally one can give a name to each clause and indicate |
|
179 that it should be added to the hints database ("[intro]"). A typical clause |
|
180 has some premises and a conclusion. This is written in Isabelle as: |
|
181 |
|
182 "premise \<Longrightarrow> conclusion" |
|
183 "\<lbrakk>premise1; premise2; \<dots> \<rbrakk> \<Longrightarrow> conclusion" |
|
184 |
|
185 If no premise is present, then one just writes |
|
186 |
|
187 "conclusion" |
|
188 |
|
189 Below we define the evaluation judgement for lambda-terms. This definition |
|
190 introduces the predicate named "eval". After giving its type, we declare |
|
191 the usual pretty syntax _ \<Down> _. In this declaration _ stands for an argument |
|
192 of eval. |
|
193 |
|
194 *} |
|
195 |
|
196 inductive |
|
197 eval :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<Down> _") |
|
198 where |
|
199 e_Lam: "Lam [x].t \<Down> Lam [x].t" |
|
200 | e_App: "\<lbrakk>t1 \<Down> Lam [x].t; t2 \<Down> v'; t[x::=v'] \<Down> v\<rbrakk> \<Longrightarrow> App t1 t2 \<Down> v" |
|
201 |
|
202 declare eval.intros[intro] |
|
203 |
|
204 text {* |
|
205 Values are also defined using inductive. In our case values |
|
206 are just lambda-abstractions. *} |
|
207 |
|
208 inductive |
|
209 val :: "lam \<Rightarrow> bool" |
|
210 where |
|
211 v_Lam[intro]: "val (Lam [x].t)" |
|
212 |
|
213 |
|
214 text {***************************************************************** |
|
215 |
|
216 Theorems |
|
217 -------- |
|
218 |
|
219 A central concept in Isabelle is that of theorems. Isabelle's theorem |
|
220 database can be queried using |
|
221 |
|
222 *} |
|
223 |
|
224 thm e_Lam |
|
225 thm e_App |
|
226 thm conjI |
|
227 thm conjunct1 |
|
228 |
|
229 text {* |
|
230 Notice that theorems usually contain schematic variables (e.g. ?P, ?Q, \<dots>). |
|
231 These schematic variables can be substituted with any term (of the right type |
|
232 of course). It is sometimes useful to suppress the "?" from the schematic |
|
233 variables. This can be achieved by using the attribute "[no_vars]". *} |
|
234 |
|
235 thm e_Lam[no_vars] |
|
236 thm e_App[no_vars] |
|
237 thm conjI[no_vars] |
|
238 thm conjunct1[no_vars] |
|
239 |
|
240 |
|
241 text {* |
|
242 When defining the predicate eval, Isabelle provides us automatically with |
|
243 the following theorems that state how evaluation judgments can be introduced |
|
244 and what constitutes an induction over the predicate eval. *} |
|
245 |
|
246 thm eval.intros[no_vars] |
|
247 thm eval.induct[no_vars] |
|
248 |
|
249 text {***************************************************************** |
|
250 |
|
251 Lemma / Theorem / Corollary Statements |
|
252 -------------------------------------- |
|
253 |
|
254 Whether to use lemma, theorem or corollary makes no semantic difference |
|
255 in Isabelle. A lemma starts with "lemma" and consists of a statement |
|
256 ("shows \<dots>") and optionally a lemma name, some type-information for |
|
257 variables ("fixes \<dots>") and some assumptions ("assumes \<dots>"). Lemmas |
|
258 also need to have a proof, but ignore this 'detail' for the moment. |
|
259 |
|
260 *} |
|
261 |
|
262 lemma alpha_equ: |
|
263 shows "Lam [x].Var x = Lam [y].Var y" |
|
264 by (simp add: lam.inject alpha swap_simps fresh_atm) |
|
265 |
|
266 lemma Lam_freshness: |
|
267 assumes a: "x\<noteq>y" |
|
268 shows "y\<sharp>Lam [x].t \<Longrightarrow> y\<sharp>t" |
|
269 using a by (simp add: abs_fresh) |
|
270 |
|
271 lemma neutral_element: |
|
272 fixes x::"nat" |
|
273 shows "x + 0 = x" |
|
274 by simp |
|
275 |
|
276 |
|
277 text {***************************************************************** |
|
278 |
|
279 Datatypes: Evaluation Contexts |
|
280 ------------------------------ |
|
281 |
|
282 Datatypes can be defined in Isabelle as follows: we have to provide the name |
|
283 of the datatype and list its type-constructors. Each type-constructor needs |
|
284 to have the information about the types of its arguments, and optionally |
|
285 can also contain some information about pretty syntax. For example, we like |
|
286 to write "\<box>" for holes. |
|
287 |
|
288 *} |
|
289 |
|
290 datatype ctx = |
|
291 Hole ("\<box>") |
|
292 | CAppL "ctx" "lam" |
|
293 | CAppR "lam" "ctx" |
|
294 |
|
295 text {* Now Isabelle knows about: *} |
|
296 |
|
297 typ ctx |
|
298 term "\<box>" |
|
299 term "CAppL" |
|
300 term "CAppL \<box> (Var x)" |
|
301 |
|
302 text {* |
|
303 |
|
304 1.) MINI EXERCISE |
|
305 ----------------- |
|
306 |
|
307 Try and see what happens if you apply a Hole to a Hole? |
|
308 |
|
309 *} |
|
310 |
|
311 |
|
312 text {***************************************************************** |
|
313 |
|
314 The CK Machine |
|
315 -------------- |
|
316 |
|
317 The CK machine is also defined using an inductive predicate with four |
|
318 arguments. The idea behind this abstract machine is to transform, or reduce, |
|
319 a configuration consisting of a lambda-term and a framestack (a list of |
|
320 contexts), to a new configuration. |
|
321 |
|
322 We use the type abbreviation "ctxs" for the type for framestacks. |
|
323 |
|
324 The pretty syntax for the CK machine is <_,_> \<mapsto> <_,_>. |
|
325 |
|
326 *} |
|
327 |
|
328 types ctxs = "ctx list" |
|
329 |
|
330 inductive |
|
331 machine :: "lam \<Rightarrow> ctxs \<Rightarrow>lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto> <_,_>") |
|
332 where |
|
333 m1[intro]: "<App t1 t2,Es> \<mapsto> <t1,(CAppL \<box> t2)#Es>" |
|
334 | m2[intro]: "val v \<Longrightarrow> <v,(CAppL \<box> t2)#Es> \<mapsto> <t2,(CAppR v \<box>)#Es>" |
|
335 | m3[intro]: "val v \<Longrightarrow> <v,(CAppR (Lam [x].t) \<box>)#Es> \<mapsto> <t[x::=v],Es>" |
|
336 |
|
337 |
|
338 text {* |
|
339 Since the machine defined above only performs a single reduction, |
|
340 we need to define the transitive closure of this machine. *} |
|
341 |
|
342 inductive |
|
343 machines :: "lam \<Rightarrow> ctxs \<Rightarrow> lam \<Rightarrow> ctxs \<Rightarrow> bool" ("<_,_> \<mapsto>* <_,_>") |
|
344 where |
|
345 ms1[intro]: "<t,Es> \<mapsto>* <t,Es>" |
|
346 | ms2[intro]: "\<lbrakk><t1,Es1> \<mapsto> <t2,Es2>; <t2,Es2> \<mapsto>* <t3,Es3>\<rbrakk> \<Longrightarrow> <t1,Es1> \<mapsto>* <t3,Es3>" |
|
347 |
|
348 |
|
349 text {***************************************************************** |
|
350 |
|
351 Isar Proofs |
|
352 ----------- |
|
353 |
|
354 Isar is a language for writing down proofs that can be understood by humans |
|
355 and by Isabelle. An Isar proof can be thought of as a sequence of 'stepping stones' |
|
356 that start with the assumptions and lead to the goal to be established. Every such |
|
357 stepping stone is introduced by "have" followed by the statement of the stepping |
|
358 stone. An exception is the goal to be proved, which need to be introduced with "show". |
|
359 |
|
360 Since proofs usually do not proceed in a linear fashion, a label can be given |
|
361 to each stepping stone and then used later to refer to this stepping stone |
|
362 ("using"). |
|
363 |
|
364 Each stepping stone (or have-statement) needs to have a justification. The |
|
365 simplest justification is "sorry" which admits any stepping stone, even false |
|
366 ones (this is good during the development of proofs). Assumption can be |
|
367 "justified" using "by fact". Derived facts can be justified using |
|
368 |
|
369 - by simp (* simplification *) |
|
370 - by auto (* proof search and simplification *) |
|
371 - by blast (* only proof search *) |
|
372 |
|
373 If facts or lemmas are needed in order to justify a have-statement, then |
|
374 one can feed these facts into the proof by using "using label \<dots>" or |
|
375 "using theorem-name \<dots>". More than one label at the time is allowed. |
|
376 |
|
377 Induction proofs in Isar are set up by indicating over which predicate(s) |
|
378 the induction proceeds ("using a b") followed by the command "proof (induct)". |
|
379 In this way, Isabelle uses default settings for which induction should |
|
380 be performed. These default settings can be overridden by giving more |
|
381 information, like the variable over which a structural induction should |
|
382 proceed, or a specific induction principle such as well-founded inductions. |
|
383 |
|
384 After the induction is set up, the proof proceeds by cases. In Isar these |
|
385 cases can be given in any order, but must be started with "case" and the |
|
386 name of the case, and optionally some legible names for the variables |
|
387 referred to inside the case. |
|
388 |
|
389 The possible case-names can be found by looking inside the menu "Isabelle -> |
|
390 Show me -> cases". In each "case", we need to establish a statement introduced |
|
391 by "show". Once this has been done, the next case can be started using "next". |
|
392 When all cases are completed, the proof can be finished using "qed". |
|
393 |
|
394 This means a typical induction proof has the following pattern |
|
395 |
|
396 proof (induct) |
|
397 case \<dots> |
|
398 \<dots> |
|
399 show \<dots> |
|
400 next |
|
401 case \<dots> |
|
402 \<dots> |
|
403 show \<dots> |
|
404 \<dots> |
|
405 qed |
|
406 |
|
407 The four lemmas are by induction on the predicate machines. All proofs establish |
|
408 the same property, namely a transitivity rule for machines. The complete Isar |
|
409 proofs are given for the first three proofs. The point of these three proofs is |
|
410 that each proof increases the readability for humans. |
|
411 |
|
412 *} |
|
413 |
|
414 text {***************************************************************** |
|
415 |
|
416 2.) EXERCISE |
|
417 ------------ |
|
418 |
|
419 Remove the sorries in the proof below and fill in the correct |
|
420 justifications. |
|
421 *} |
|
422 |
|
423 lemma |
|
424 assumes a: "<e1,Es1> \<mapsto>* <e2,Es2>" |
|
425 and b: "<e2,Es2> \<mapsto>* <e3,Es3>" |
|
426 shows "<e1,Es1> \<mapsto>* <e3,Es3>" |
|
427 using a b |
|
428 proof(induct) |
|
429 case (ms1 e1 Es1) |
|
430 have c: "<e1,Es1> \<mapsto>* <e3,Es3>" by fact |
|
431 show "<e1,Es1> \<mapsto>* <e3,Es3>" sorry |
|
432 next |
|
433 case (ms2 e1 Es1 e2 Es2 e2' Es2') |
|
434 have ih: "<e2',Es2'> \<mapsto>* <e3,Es3> \<Longrightarrow> <e2,Es2> \<mapsto>* <e3,Es3>" by fact |
|
435 have d1: "<e2',Es2'> \<mapsto>* <e3,Es3>" by fact |
|
436 have d2: "<e1,Es1> \<mapsto> <e2,Es2>" by fact |
|
437 |
|
438 show "<e1,Es1> \<mapsto>* <e3,Es3>" sorry |
|
439 qed |
|
440 |
|
441 text {* |
|
442 Just like gotos in the Basic programming language, labels can reduce |
|
443 the readability of proofs. Therefore one can use in Isar the notation |
|
444 "then have" in order to feed a have-statement to the proof of |
|
445 the next have-statement. In the proof below this has been used |
|
446 in order to get rid of the labels c and d1. |
|
447 *} |
|
448 |
|
449 lemma |
|
450 assumes a: "<e1,Es1> \<mapsto>* <e2,Es2>" |
|
451 and b: "<e2,Es2> \<mapsto>* <e3,Es3>" |
|
452 shows "<e1,Es1> \<mapsto>* <e3,Es3>" |
|
453 using a b |
|
454 proof(induct) |
|
455 case (ms1 e1 Es1) |
|
456 show "<e1,Es1> \<mapsto>* <e3,Es3>" by fact |
|
457 next |
|
458 case (ms2 e1 Es1 e2 Es2 e2' Es2') |
|
459 have ih: "<e2',Es2'> \<mapsto>* <e3,Es3> \<Longrightarrow> <e2,Es2> \<mapsto>* <e3,Es3>" by fact |
|
460 have "<e2',Es2'> \<mapsto>* <e3,Es3>" by fact |
|
461 then have d3: "<e2,Es2> \<mapsto>* <e3,Es3>" using ih by simp |
|
462 have d2: "<e1,Es1> \<mapsto> <e2,Es2>" by fact |
|
463 show "<e1,Es1> \<mapsto>* <e3,Es3>" using d2 d3 by auto |
|
464 qed |
|
465 |
|
466 text {* |
|
467 The labels d2 and d3 cannot be got rid of in this way, because both |
|
468 facts are needed to prove the show-statement. We can still avoid the |
|
469 labels by feeding a sequence of facts into a proof using the chaining |
|
470 mechanism: |
|
471 |
|
472 have "statement1" \<dots> |
|
473 moreover |
|
474 have "statement2" \<dots> |
|
475 \<dots> |
|
476 moreover |
|
477 have "statementn" \<dots> |
|
478 ultimately have "statement" \<dots> |
|
479 |
|
480 In this chain, all "statementi" can be used in the proof of the final |
|
481 "statement". With this we can simplify our proof further to: |
|
482 *} |
|
483 |
|
484 lemma |
|
485 assumes a: "<e1,Es1> \<mapsto>* <e2,Es2>" |
|
486 and b: "<e2,Es2> \<mapsto>* <e3,Es3>" |
|
487 shows "<e1,Es1> \<mapsto>* <e3,Es3>" |
|
488 using a b |
|
489 proof(induct) |
|
490 case (ms1 e1 Es1) |
|
491 show "<e1,Es1> \<mapsto>* <e3,Es3>" by fact |
|
492 next |
|
493 case (ms2 e1 Es1 e2 Es2 e2' Es2') |
|
494 have ih: "<e2',Es2'> \<mapsto>* <e3,Es3> \<Longrightarrow> <e2,Es2> \<mapsto>* <e3,Es3>" by fact |
|
495 have "<e2',Es2'> \<mapsto>* <e3,Es3>" by fact |
|
496 then have "<e2,Es2> \<mapsto>* <e3,Es3>" using ih by simp |
|
497 moreover |
|
498 have "<e1,Es1> \<mapsto> <e2,Es2>" by fact |
|
499 ultimately show "<e1,Es1> \<mapsto>* <e3,Es3>" by auto |
|
500 qed |
|
501 |
|
502 text {* |
|
503 While automatic proof procedures in Isabelle are not able to prove statements |
|
504 like "P = NP" assuming usual definitions for P and NP, they can automatically |
|
505 discharge the lemma we just proved. For this we only have to set up the induction |
|
506 and auto will take care of the rest. This means we can write: |
|
507 *} |
|
508 |
|
509 lemma ms3[intro]: |
|
510 assumes a: "<e1,Es1> \<mapsto>* <e2,Es2>" |
|
511 and b: "<e2,Es2> \<mapsto>* <e3,Es3>" |
|
512 shows "<e1,Es1> \<mapsto>* <e3,Es3>" |
|
513 using a b by (induct) (auto) |
|
514 |
|
515 text {* |
|
516 The attribute [intro] indicates that this lemma should be from now on used in |
|
517 any proof obtained by "auto" or "blast". |
|
518 *} |
|
519 |
|
520 |
|
521 text {***************************************************************** |
|
522 |
|
523 A simple fact we need later on is that if t \<Down> t' then t' is a value. |
|
524 |
|
525 *} |
|
526 |
|
527 lemma eval_to_val: |
|
528 assumes a: "t \<Down> t'" |
|
529 shows "val t'" |
|
530 using a by (induct) (auto) |
|
531 |
|
532 |
|
533 text {***************************************************************** |
|
534 |
|
535 3.) EXERCISE |
|
536 ------------ |
|
537 |
|
538 Fill in the details in the proof below. The proof will establish the fact |
|
539 that if t \<Down> t' then <t,[]> \<mapsto>* <t',[]>. As can be seen, the proof is by |
|
540 induction on the definition of eval. If you want to know how the predicates |
|
541 machine and machines can be introduced, then use |
|
542 |
|
543 thm machine.intros[no_vars] |
|
544 thm machines.intros[no_vars] |
|
545 |
|
546 *} |
|
547 |
|
548 theorem |
|
549 assumes a: "t \<Down> t'" |
|
550 shows "<t,[]> \<mapsto>* <t',[]>" |
|
551 using a |
|
552 proof (induct) |
|
553 case (e_Lam x t) |
|
554 (* no assumptions *) |
|
555 show "<Lam [x].t,[]> \<mapsto>* <Lam [x].t,[]>" sorry |
|
556 next |
|
557 case (e_App t1 x t t2 v' v) |
|
558 (* all assumptions in this case *) |
|
559 have a1: "t1 \<Down> Lam [x].t" by fact |
|
560 have ih1: "<t1,[]> \<mapsto>* <Lam [x].t,[]>" by fact |
|
561 have a2: "t2 \<Down> v'" by fact |
|
562 have ih2: "<t2,[]> \<mapsto>* <v',[]>" by fact |
|
563 have a3: "t[x::=v'] \<Down> v" by fact |
|
564 have ih3: "<t[x::=v'],[]> \<mapsto>* <v,[]>" by fact |
|
565 (* your details *) |
|
566 show "<App t1 t2,[]> \<mapsto>* <v,[]>" sorry |
|
567 qed |
|
568 |
|
569 text {* |
|
570 Again the automatic tools in Isabelle can discharge automatically |
|
571 of the routine work in these proofs. We can write: *} |
|
572 |
|
573 theorem eval_implies_machines_ctx: |
|
574 assumes a: "t \<Down> t'" |
|
575 shows "<t,Es> \<mapsto>* <t',Es>" |
|
576 using a |
|
577 by (induct arbitrary: Es) |
|
578 (metis eval_to_val machine.intros ms1 ms2 ms3 v_Lam)+ |
|
579 |
|
580 corollary eval_implies_machines: |
|
581 assumes a: "t \<Down> t'" |
|
582 shows "<t,[]> \<mapsto>* <t',[]>" |
|
583 using a eval_implies_machines_ctx by simp |
|
584 |
|
585 text {***************************************************************** |
|
586 |
|
587 The Weakening Lemma |
|
588 ------------------- |
|
589 |
|
590 The proof of the weakening lemma is often said to be simple, |
|
591 routine or trivial. Below we will see how this lemma can be |
|
592 proved in Nominal Isabelle. First we define types, which |
|
593 we however do not define as datatypes, but as nominal datatypes. |
|
594 |
|
595 *} |
|
596 |
|
597 nominal_datatype ty = |
|
598 tVar "string" |
|
599 | tArr "ty" "ty" ("_ \<rightarrow> _") |
|
600 |
|
601 text {* |
|
602 Having defined them as nominal datatypes gives us additional |
|
603 definitions and theorems that come with types (see below). |
|
604 *} |
|
605 |
|
606 text {* |
|
607 We next define the type of typing contexts, a predicate that |
|
608 defines valid contexts (i.e. lists that contain only unique |
|
609 variables) and the typing judgement. |
|
610 |
|
611 *} |
|
612 |
|
613 types ty_ctx = "(name\<times>ty) list" |
|
614 |
|
615 inductive |
|
616 valid :: "ty_ctx \<Rightarrow> bool" |
|
617 where |
|
618 v1[intro]: "valid []" |
|
619 | v2[intro]: "\<lbrakk>valid \<Gamma>; x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((x,T)#\<Gamma>)" |
|
620 |
|
621 inductive |
|
622 typing :: "ty_ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _") |
|
623 where |
|
624 t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T" |
|
625 | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2" |
|
626 | t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T1)#\<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1 \<rightarrow> T2" |
|
627 |
|
628 |
|
629 text {* |
|
630 The predicate x\<sharp>\<Gamma>, read as x fresh for \<Gamma>, is defined by Nominal Isabelle. |
|
631 Freshness is defined for lambda-terms, products, lists etc. For example |
|
632 we have: |
|
633 *} |
|
634 |
|
635 lemma |
|
636 fixes x::"name" |
|
637 shows "x\<sharp>Lam [x].t" |
|
638 and "x\<sharp>(t1,t2) \<Longrightarrow> x\<sharp>App t1 t2" |
|
639 and "x\<sharp>(Var y) \<Longrightarrow> x\<sharp>y" |
|
640 and "\<lbrakk>x\<sharp>t1; x\<sharp>t2\<rbrakk> \<Longrightarrow> x\<sharp>(t1,t2)" |
|
641 and "\<lbrakk>x\<sharp>l1; x\<sharp>l2\<rbrakk> \<Longrightarrow> x\<sharp>(l1@l2)" |
|
642 and "x\<sharp>y \<Longrightarrow> x\<noteq>y" |
|
643 by (simp_all add: abs_fresh fresh_prod fresh_list_append fresh_atm) |
|
644 |
|
645 text {* We can also prove that every variable is fresh for a type *} |
|
646 |
|
647 lemma ty_fresh: |
|
648 fixes x::"name" |
|
649 and T::"ty" |
|
650 shows "x\<sharp>T" |
|
651 by (induct T rule: ty.induct) |
|
652 (simp_all add: fresh_string) |
|
653 |
|
654 text {* |
|
655 In order to state the weakening lemma in a convenient form, we overload |
|
656 the subset-notation and define the abbreviation below. Abbreviations behave |
|
657 like definitions, except that they are always automatically folded and |
|
658 unfolded. |
|
659 *} |
|
660 |
|
661 abbreviation |
|
662 "sub_ty_ctx" :: "ty_ctx \<Rightarrow> ty_ctx \<Rightarrow> bool" ("_ \<subseteq> _") |
|
663 where |
|
664 "\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x. x \<in> set \<Gamma>1 \<longrightarrow> x \<in> set \<Gamma>2" |
|
665 |
|
666 text {***************************************************************** |
|
667 |
|
668 4.) Exercise |
|
669 ------------ |
|
670 |
|
671 Fill in the details and give a proof of the weakening lemma. |
|
672 |
|
673 *} |
|
674 |
|
675 lemma |
|
676 fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list" |
|
677 assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
678 and b: "valid \<Gamma>2" |
|
679 and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
|
680 shows "\<Gamma>2 \<turnstile> t : T" |
|
681 using a b c |
|
682 proof (induct arbitrary: \<Gamma>2) |
|
683 case (t_Var \<Gamma>1 x T) |
|
684 have a1: "valid \<Gamma>1" by fact |
|
685 have a2: "(x,T) \<in> set \<Gamma>1" by fact |
|
686 have a3: "valid \<Gamma>2" by fact |
|
687 have a4: "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
|
688 |
|
689 show "\<Gamma>2 \<turnstile> Var x : T" sorry |
|
690 next |
|
691 case (t_Lam x \<Gamma>1 T1 t T2) |
|
692 have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x,T1)#\<Gamma>1 \<subseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact |
|
693 have a0: "x\<sharp>\<Gamma>1" by fact |
|
694 have a1: "valid \<Gamma>2" by fact |
|
695 have a2: "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
|
696 |
|
697 show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry |
|
698 qed (auto) |
|
699 |
|
700 |
|
701 text {* |
|
702 Despite the frequent claim that the weakening lemma is trivial, |
|
703 routine or obvious, the proof in the lambda-case does not go |
|
704 smoothly through. Painful variable renamings seem to be necessary. |
|
705 But the proof using renamings is horribly complicated. It is really |
|
706 interesting whether people who claim this proof is trivial, routine |
|
707 or obvious had this proof in mind. |
|
708 |
|
709 BTW: The following two commands help already with showing that validity |
|
710 and typing are invariant under variable (permutative) renamings. |
|
711 *} |
|
712 |
|
713 equivariance valid |
|
714 equivariance typing |
|
715 |
|
716 lemma not_to_be_tried_at_home_weakening: |
|
717 fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list" |
|
718 assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
719 and b: "valid \<Gamma>2" |
|
720 and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
|
721 shows "\<Gamma>2 \<turnstile> t : T" |
|
722 using a b c |
|
723 proof (induct arbitrary: \<Gamma>2) |
|
724 case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *) |
|
725 have fc0: "x\<sharp>\<Gamma>1" by fact |
|
726 have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x,T1)#\<Gamma>1 \<subseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact |
|
727 obtain c::"name" where fc1: "c\<sharp>(x,t,\<Gamma>1,\<Gamma>2)" (* we obtain a fresh name *) |
|
728 by (rule exists_fresh) (auto simp add: fs_name1) |
|
729 have "Lam [c].([(c,x)]\<bullet>t) = Lam [x].t" using fc1 (* we then alpha-rename the lambda-abstraction *) |
|
730 by (auto simp add: lam.inject alpha fresh_prod fresh_atm) |
|
731 moreover |
|
732 have "\<Gamma>2 \<turnstile> Lam [c].([(c,x)]\<bullet>t) : T1 \<rightarrow> T2" (* we can then alpha-rename our original goal *) |
|
733 proof - |
|
734 (* we establish (x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2) and valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2)) *) |
|
735 have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2)" |
|
736 proof - |
|
737 have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
|
738 then have "[(c,x)]\<bullet>((x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2))" using fc0 fc1 |
|
739 by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh) |
|
740 then show "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#([(c,x)]\<bullet>\<Gamma>2)" by (rule perm_boolE) |
|
741 qed |
|
742 moreover |
|
743 have "valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2))" |
|
744 proof - |
|
745 have "valid \<Gamma>2" by fact |
|
746 then show "valid ((x,T1)#([(c,x)]\<bullet>\<Gamma>2))" using fc1 |
|
747 by (auto intro!: v2 simp add: fresh_left calc_atm eqvts) |
|
748 qed |
|
749 (* these two facts give us by induction hypothesis the following *) |
|
750 ultimately have "(x,T1)#([(c,x)]\<bullet>\<Gamma>2) \<turnstile> t : T2" using ih by simp |
|
751 (* we now apply renamings to get to our goal *) |
|
752 then have "[(c,x)]\<bullet>((x,T1)#([(c,x)]\<bullet>\<Gamma>2) \<turnstile> t : T2)" by (rule perm_boolI) |
|
753 then have "(c,T1)#\<Gamma>2 \<turnstile> ([(c,x)]\<bullet>t) : T2" using fc1 |
|
754 by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh) |
|
755 then show "\<Gamma>2 \<turnstile> Lam [c].([(c,x)]\<bullet>t) : T1 \<rightarrow> T2" using fc1 by auto |
|
756 qed |
|
757 ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp |
|
758 qed (auto) (* var and app cases *) |
|
759 |
|
760 |
|
761 text {* |
|
762 The remedy to the complicated proof of the weakening proof |
|
763 shown above is to use a stronger induction principle that |
|
764 has the usual variable convention already build in. The |
|
765 following command derives this induction principle for us. |
|
766 (We shall explain what happens here later.) |
|
767 |
|
768 *} |
|
769 |
|
770 nominal_inductive typing |
|
771 by (simp_all add: abs_fresh ty_fresh) |
|
772 |
|
773 text {* Compare the two induction principles *} |
|
774 thm typing.induct[no_vars] |
|
775 thm typing.strong_induct[no_vars] |
|
776 |
|
777 text {* |
|
778 We can use the stronger induction principle by replacing |
|
779 the line |
|
780 |
|
781 proof (induct arbitrary: \<Gamma>2) |
|
782 |
|
783 with |
|
784 |
|
785 proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct) |
|
786 |
|
787 Try now the proof. |
|
788 |
|
789 *} |
|
790 |
|
791 |
|
792 lemma |
|
793 fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list" |
|
794 and t ::"lam" |
|
795 and \<tau> ::"ty" |
|
796 assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
797 and b: "valid \<Gamma>2" |
|
798 and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
|
799 shows "\<Gamma>2 \<turnstile> t : T" |
|
800 using a b c |
|
801 proof (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct) |
|
802 case (t_Var \<Gamma>1 x T) (* variable case *) |
|
803 have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
|
804 moreover |
|
805 have "valid \<Gamma>2" by fact |
|
806 moreover |
|
807 have "(x,T)\<in> set \<Gamma>1" by fact |
|
808 ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto |
|
809 next |
|
810 case (t_Lam x \<Gamma>1 T1 t T2) |
|
811 have vc: "x\<sharp>\<Gamma>2" by fact (* additional fact *) |
|
812 have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; (x,T1)#\<Gamma>1 \<subseteq> \<Gamma>3\<rbrakk> \<Longrightarrow> \<Gamma>3 \<turnstile> t : T2" by fact |
|
813 have a0: "x\<sharp>\<Gamma>1" by fact |
|
814 have a1: "valid \<Gamma>2" by fact |
|
815 have a2: "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
|
816 |
|
817 show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" sorry |
|
818 qed (auto) (* app case *) |
|
819 |
|
820 text {* |
|
821 Since we can use the stronger induction principle, the |
|
822 proof of the weakening lemma can actually be found |
|
823 automatically by Isabelle. Maybe the weakening lemma |
|
824 is actually trivial (in Nominal Isabelle ;o). |
|
825 *} |
|
826 |
|
827 lemma weakening: |
|
828 fixes \<Gamma>1 \<Gamma>2::"ty_ctx" |
|
829 assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
830 and b: "valid \<Gamma>2" |
|
831 and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
|
832 shows "\<Gamma>2 \<turnstile> t : T" |
|
833 using a b c |
|
834 by (nominal_induct avoiding: \<Gamma>2 rule: typing.strong_induct) |
|
835 (auto) |
|
836 |
|
837 |
|
838 text {***************************************************************** |
|
839 |
|
840 Function Definitions: Filling a Lambda-Term into a Context |
|
841 ---------------------------------------------------------- |
|
842 |
|
843 Many functions over datatypes can be defined by recursion on the |
|
844 structure. For this purpose, Isabelle provides "fun". To use it one needs |
|
845 to give a name for the function, its type, optionally some pretty-syntax |
|
846 and then some equations defining the function. Like in "inductive", |
|
847 "fun" expects that more than one such equation is separated by "|". |
|
848 |
|
849 *} |
|
850 |
|
851 fun |
|
852 filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>") |
|
853 where |
|
854 "\<box>\<lbrakk>t\<rbrakk> = t" |
|
855 | "(CAppL E t')\<lbrakk>t\<rbrakk> = App (E\<lbrakk>t\<rbrakk>) t'" |
|
856 | "(CAppR t' E)\<lbrakk>t\<rbrakk> = App t' (E\<lbrakk>t\<rbrakk>)" |
|
857 |
|
858 text {* |
|
859 After this definition Isabelle will be able to simplify |
|
860 statements like: *} |
|
861 |
|
862 lemma |
|
863 shows "(CAppL \<box> (Var x))\<lbrakk>Var y\<rbrakk> = App (Var y) (Var x)" |
|
864 by simp |
|
865 |
|
866 |
|
867 fun |
|
868 ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _" [101,100] 100) |
|
869 where |
|
870 "\<box> \<circ> E' = E'" |
|
871 | "(CAppL E t') \<circ> E' = CAppL (E \<circ> E') t'" |
|
872 | "(CAppR t' E) \<circ> E' = CAppR t' (E \<circ> E')" |
|
873 |
|
874 fun |
|
875 ctx_composes :: "ctxs \<Rightarrow> ctx" ("_\<down>" [110] 110) |
|
876 where |
|
877 "[]\<down> = \<box>" |
|
878 | "(E#Es)\<down> = (Es\<down>) \<circ> E" |
|
879 |
|
880 text {* |
|
881 Notice that we not just have given a pretty syntax for the functions, but |
|
882 also some precedences..The numbers inside the [\<dots>] stand for the precedences |
|
883 of the arguments; the one next to it the precedence of the whole term. |
|
884 |
|
885 This means we have to write (Es1 \<circ> Es2) \<circ> Es3 otherwise Es1 \<circ> Es2 \<circ> Es3 is |
|
886 interpreted as Es1 \<circ> (Es2 \<circ> Es3). |
|
887 *} |
|
888 |
|
889 text {****************************************************************** |
|
890 |
|
891 Structural Inductions over Contexts |
|
892 ------------------------------------ |
|
893 |
|
894 So far we have had a look at an induction over an inductive predicate. |
|
895 Another important induction principle is structural inductions for |
|
896 datatypes. To illustrate structural inductions we prove some facts |
|
897 about context composition, some of which we will need later on. |
|
898 |
|
899 *} |
|
900 |
|
901 text {****************************************************************** |
|
902 |
|
903 5.) EXERCISE |
|
904 ------------ |
|
905 |
|
906 Complete the proof and remove the sorries. |
|
907 |
|
908 *} |
|
909 |
|
910 lemma ctx_compose: |
|
911 shows "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" |
|
912 proof (induct E1) |
|
913 case Hole |
|
914 show "\<box> \<circ> E2\<lbrakk>t\<rbrakk> = \<box>\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry |
|
915 next |
|
916 case (CAppL E1 t') |
|
917 have ih: "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact |
|
918 show "((CAppL E1 t') \<circ> E2)\<lbrakk>t\<rbrakk> = (CAppL E1 t')\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry |
|
919 next |
|
920 case (CAppR t' E1) |
|
921 have ih: "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" by fact |
|
922 show "((CAppR t' E1) \<circ> E2)\<lbrakk>t\<rbrakk> = (CAppR t' E1)\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" sorry |
|
923 qed |
|
924 |
|
925 |
|
926 text {****************************************************************** |
|
927 |
|
928 6.) EXERCISE |
|
929 ------------ |
|
930 |
|
931 Prove associativity of \<circ> using the lemmas neut_hole and circ_assoc. |
|
932 |
|
933 *} |
|
934 |
|
935 lemma neut_hole: |
|
936 shows "E \<circ> \<box> = E" |
|
937 by (induct E) (simp_all) |
|
938 |
|
939 lemma circ_assoc: |
|
940 fixes E1 E2 E3::"ctx" |
|
941 shows "(E1 \<circ> E2) \<circ> E3 = E1 \<circ> (E2 \<circ> E3)" |
|
942 by (induct E1) (simp_all) |
|
943 |
|
944 lemma |
|
945 shows "(Es1@Es2)\<down> = (Es2\<down>) \<circ> (Es1\<down>)" |
|
946 proof (induct Es1) |
|
947 case Nil |
|
948 show "([]@Es2)\<down> = Es2\<down> \<circ> []\<down>" sorry |
|
949 next |
|
950 case (Cons E Es1) |
|
951 have ih: "(Es1@Es2)\<down> = Es2\<down> \<circ> Es1\<down>" by fact |
|
952 |
|
953 show "((E#Es1)@Es2)\<down> = Es2\<down> \<circ> (E#Es1)\<down>" sorry |
|
954 qed |
|
955 |
|
956 |
|
957 text {* |
|
958 The last proof involves several steps of equational reasoning. |
|
959 Isar provides some convenient means to express such equational |
|
960 reasoning in a much cleaner fashion using the "also have" |
|
961 construction. *} |
|
962 |
|
963 lemma |
|
964 shows "(Es1@Es2)\<down> = (Es2\<down>) \<circ> (Es1\<down>)" |
|
965 proof (induct Es1) |
|
966 case Nil |
|
967 show "([]@Es2)\<down> = Es2\<down> \<circ> []\<down>" using neut_hole by simp |
|
968 next |
|
969 case (Cons E Es1) |
|
970 have ih: "(Es1@Es2)\<down> = Es2\<down> \<circ> Es1\<down>" by fact |
|
971 have "((E#Es1)@Es2)\<down> = (Es1@Es2)\<down> \<circ> E" by simp |
|
972 also have "\<dots> = (Es2\<down> \<circ> Es1\<down>) \<circ> E" using ih by simp |
|
973 also have "\<dots> = Es2\<down> \<circ> (Es1\<down> \<circ> E)" using circ_assoc by simp |
|
974 also have "\<dots> = Es2\<down> \<circ> (E#Es1)\<down>" by simp |
|
975 finally show "((E#Es1)@Es2)\<down> = Es2\<down> \<circ> (E#Es1)\<down>" by simp |
|
976 qed |
|
977 |
|
978 |
|
979 text {****************************************************************** |
|
980 |
|
981 Formalising Barendregt's Proof of the Substitution Lemma |
|
982 -------------------------------------------------------- |
|
983 |
|
984 Barendregt's proof needs in the variable case a case distinction. |
|
985 One way to do this in Isar is to use blocks. A block is some sequent |
|
986 or reasoning steps enclosed in curly braces |
|
987 |
|
988 { \<dots> |
|
989 |
|
990 have "statement" |
|
991 } |
|
992 |
|
993 Such a block can contain local assumptions like |
|
994 |
|
995 { assume "A" |
|
996 assume "B" |
|
997 \<dots> |
|
998 have "C" by \<dots> |
|
999 } |
|
1000 |
|
1001 Where "C" is the last have-statement in this block. The behaviour |
|
1002 of such a block to the 'outside' is the implication |
|
1003 |
|
1004 \<lbrakk>A; B\<rbrakk> \<Longrightarrow> "C" |
|
1005 |
|
1006 Now if we want to prove a property "smth" using the case-distinctions |
|
1007 P\<^isub>1, P\<^isub>2 and P\<^isub>3 then we can use the following reasoning: |
|
1008 |
|
1009 { assume "P\<^isub>1" |
|
1010 \<dots> |
|
1011 have "smth" |
|
1012 } |
|
1013 moreover |
|
1014 { assume "P\<^isub>2" |
|
1015 \<dots> |
|
1016 have "smth" |
|
1017 } |
|
1018 moreover |
|
1019 { assume "P\<^isub>3" |
|
1020 \<dots> |
|
1021 have "smth" |
|
1022 } |
|
1023 ultimately have "smth" by blast |
|
1024 |
|
1025 The blocks establish the implications |
|
1026 |
|
1027 P\<^isub>1 \<Longrightarrow> smth |
|
1028 P\<^isub>2 \<Longrightarrow> smth |
|
1029 P\<^isub>3 \<Longrightarrow> smth |
|
1030 |
|
1031 If we know that P\<^isub>1, P\<^isub>2 and P\<^isub>3 cover all the cases, that is P\<^isub>1 \<or> P\<^isub>2 \<or> P\<^isub>3 is |
|
1032 true, then we have 'ultimately' established the property "smth" |
|
1033 |
|
1034 *} |
|
1035 |
|
1036 text {****************************************************************** |
|
1037 |
|
1038 7.) Exercise |
|
1039 ------------ |
|
1040 |
|
1041 Fill in the cases 1.2 and 1.3 and the equational reasoning |
|
1042 in the lambda-case. |
|
1043 *} |
|
1044 |
|
1045 thm forget[no_vars] |
|
1046 thm fresh_fact[no_vars] |
|
1047 |
|
1048 lemma |
|
1049 assumes a: "x\<noteq>y" |
|
1050 and b: "x\<sharp>L" |
|
1051 shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" |
|
1052 using a b |
|
1053 proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct) |
|
1054 case (Var z) |
|
1055 have a1: "x\<noteq>y" by fact |
|
1056 have a2: "x\<sharp>L" by fact |
|
1057 show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS") |
|
1058 proof - |
|
1059 { (*Case 1.1*) |
|
1060 assume c1: "z=x" |
|
1061 have "(1)": "?LHS = N[y::=L]" using c1 by simp |
|
1062 have "(2)": "?RHS = N[y::=L]" using c1 a1 by simp |
|
1063 have "?LHS = ?RHS" using "(1)" "(2)" by simp |
|
1064 } |
|
1065 moreover |
|
1066 { (*Case 1.2*) |
|
1067 assume c2: "z=y" "z\<noteq>x" |
|
1068 |
|
1069 have "?LHS = ?RHS" sorry |
|
1070 } |
|
1071 moreover |
|
1072 { (*Case 1.3*) |
|
1073 assume c3: "z\<noteq>x" "z\<noteq>y" |
|
1074 |
|
1075 have "?LHS = ?RHS" sorry |
|
1076 } |
|
1077 ultimately show "?LHS = ?RHS" by blast |
|
1078 qed |
|
1079 next |
|
1080 case (Lam z M1) (* case 2: lambdas *) |
|
1081 have ih: "\<lbrakk>x\<noteq>y; x\<sharp>L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact |
|
1082 have a1: "x\<noteq>y" by fact |
|
1083 have a2: "x\<sharp>L" by fact |
|
1084 have fs: "z\<sharp>x" "z\<sharp>y" "z\<sharp>N" "z\<sharp>L" by fact+ |
|
1085 then have b: "z\<sharp>N[y::=L]" by (simp add: fresh_fact) |
|
1086 show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") |
|
1087 proof - |
|
1088 have "?LHS = \<dots>" sorry |
|
1089 |
|
1090 also have "\<dots> = ?RHS" sorry |
|
1091 finally show "?LHS = ?RHS" by simp |
|
1092 qed |
|
1093 next |
|
1094 case (App M1 M2) (* case 3: applications *) |
|
1095 then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp |
|
1096 qed |
|
1097 |
|
1098 text {* |
|
1099 Again the strong induction principle enables Isabelle to find |
|
1100 the proof of the substitution lemma automatically. |
|
1101 *} |
|
1102 |
|
1103 lemma substitution_lemma_version: |
|
1104 assumes asm: "x\<noteq>y" "x\<sharp>L" |
|
1105 shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" |
|
1106 using asm |
|
1107 by (nominal_induct M avoiding: x y N L rule: lam.strong_induct) |
|
1108 (auto simp add: fresh_fact forget) |
|
1109 |
|
1110 text {****************************************************************** |
|
1111 |
|
1112 The CBV Reduction Relation (Small-Step Semantics) |
|
1113 ------------------------------------------------- |
|
1114 |
|
1115 In order to establish the property that the CK Machine |
|
1116 calculates a nomrmalform which corresponds to the |
|
1117 evaluation relation, we introduce the call-by-value |
|
1118 small-step semantics. |
|
1119 |
|
1120 *} |
|
1121 |
|
1122 inductive |
|
1123 cbv :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>cbv _") |
|
1124 where |
|
1125 cbv1: "\<lbrakk>val v; x\<sharp>v\<rbrakk> \<Longrightarrow> App (Lam [x].t) v \<longrightarrow>cbv t[x::=v]" |
|
1126 | cbv2[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> App t t2 \<longrightarrow>cbv App t' t2" |
|
1127 | cbv3[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> App t2 t \<longrightarrow>cbv App t2 t'" |
|
1128 |
|
1129 equivariance val |
|
1130 equivariance cbv |
|
1131 nominal_inductive cbv |
|
1132 by (simp_all add: abs_fresh fresh_fact) |
|
1133 |
|
1134 text {* |
|
1135 In order to satisfy the vc-condition we have to formulate |
|
1136 this relation with the additional freshness constraint |
|
1137 x\<sharp>v. Though this makes the definition vc-ompatible, it |
|
1138 makes the definition less useful. We can with some pain |
|
1139 show that the more restricted rule is equivalent to the |
|
1140 usual rule. *} |
|
1141 |
|
1142 thm subst_rename |
|
1143 |
|
1144 lemma better_cbv1[intro]: |
|
1145 assumes a: "val v" |
|
1146 shows "App (Lam [x].t) v \<longrightarrow>cbv t[x::=v]" |
|
1147 proof - |
|
1148 obtain y::"name" where fs: "y\<sharp>(x,t,v)" by (rule exists_fresh) (auto simp add: fs_name1) |
|
1149 have "App (Lam [x].t) v = App (Lam [y].([(y,x)]\<bullet>t)) v" using fs |
|
1150 by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) |
|
1151 also have "\<dots> \<longrightarrow>cbv ([(y,x)]\<bullet>t)[y::=v]" using fs a by (auto intro: cbv1) |
|
1152 also have "\<dots> = t[x::=v]" using fs by (simp add: subst_rename[symmetric]) |
|
1153 finally show "App (Lam [x].t) v \<longrightarrow>cbv t[x::=v]" by simp |
|
1154 qed |
|
1155 |
|
1156 text {* |
|
1157 The transitive closure of the cbv-reduction relation: *} |
|
1158 |
|
1159 inductive |
|
1160 "cbvs" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>cbv* _") |
|
1161 where |
|
1162 cbvs1[intro]: "e \<longrightarrow>cbv* e" |
|
1163 | cbvs2[intro]: "\<lbrakk>e1\<longrightarrow>cbv e2; e2 \<longrightarrow>cbv* e3\<rbrakk> \<Longrightarrow> e1 \<longrightarrow>cbv* e3" |
|
1164 |
|
1165 lemma cbvs3[intro]: |
|
1166 assumes a: "e1 \<longrightarrow>cbv* e2" "e2 \<longrightarrow>cbv* e3" |
|
1167 shows "e1 \<longrightarrow>cbv* e3" |
|
1168 using a by (induct) (auto) |
|
1169 |
|
1170 text {****************************************************************** |
|
1171 |
|
1172 8.) Exercise |
|
1173 ------------ |
|
1174 |
|
1175 If more simple exercises are needed, then complete the following proof. |
|
1176 |
|
1177 *} |
|
1178 |
|
1179 lemma cbv_in_ctx: |
|
1180 assumes a: "t \<longrightarrow>cbv t'" |
|
1181 shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" |
|
1182 using a |
|
1183 proof (induct E) |
|
1184 case Hole |
|
1185 have "t \<longrightarrow>cbv t'" by fact |
|
1186 then show "\<box>\<lbrakk>t\<rbrakk> \<longrightarrow>cbv \<box>\<lbrakk>t'\<rbrakk>" sorry |
|
1187 next |
|
1188 case (CAppL E s) |
|
1189 have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact |
|
1190 have a: "t \<longrightarrow>cbv t'" by fact |
|
1191 show "(CAppL E s)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppL E s)\<lbrakk>t'\<rbrakk>" sorry |
|
1192 next |
|
1193 case (CAppR s E) |
|
1194 have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact |
|
1195 have a: "t \<longrightarrow>cbv t'" by fact |
|
1196 show "(CAppR s E)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppR s E)\<lbrakk>t'\<rbrakk>" sorry |
|
1197 qed |
|
1198 |
|
1199 |
|
1200 text {****************************************************************** |
|
1201 |
|
1202 9.) Exercise |
|
1203 ------------ |
|
1204 |
|
1205 The point of the cbv-reduction was that we can easily relatively |
|
1206 establish the follwoing property: |
|
1207 |
|
1208 *} |
|
1209 |
|
1210 lemma machine_implies_cbvs_ctx: |
|
1211 assumes a: "<e,Es> \<mapsto> <e',Es'>" |
|
1212 shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>" |
|
1213 using a |
|
1214 proof (induct) |
|
1215 case (m1 t1 t2 Es) |
|
1216 |
|
1217 show "Es\<down>\<lbrakk>App t1 t2\<rbrakk> \<longrightarrow>cbv* ((CAppL \<box> t2)#Es)\<down>\<lbrakk>t1\<rbrakk>" sorry |
|
1218 next |
|
1219 case (m2 v t2 Es) |
|
1220 have "val v" by fact |
|
1221 |
|
1222 show "((CAppL \<box> t2)#Es)\<down>\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (CAppR v \<box> # Es)\<down>\<lbrakk>t2\<rbrakk>" sorry |
|
1223 next |
|
1224 case (m3 v x t Es) |
|
1225 have "val v" by fact |
|
1226 |
|
1227 show "(((CAppR (Lam [x].t) \<box>)#Es)\<down>)\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (Es\<down>)\<lbrakk>t[x::=v]\<rbrakk>" sorry |
|
1228 qed |
|
1229 |
|
1230 text {* |
|
1231 It is not difficult to extend the lemma above to |
|
1232 arbitrary reductions sequences of the CK machine. *} |
|
1233 |
|
1234 lemma machines_implies_cbvs_ctx: |
|
1235 assumes a: "<e,Es> \<mapsto>* <e',Es'>" |
|
1236 shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>" |
|
1237 using a |
|
1238 by (induct) (auto dest: machine_implies_cbvs_ctx) |
|
1239 |
|
1240 text {* |
|
1241 So whenever we let the CL machine start in an initial |
|
1242 state and it arrives at a final state, then there exists |
|
1243 a corresponding cbv-reduction sequence. *} |
|
1244 |
|
1245 corollary machines_implies_cbvs: |
|
1246 assumes a: "<e,[]> \<mapsto>* <e',[]>" |
|
1247 shows "e \<longrightarrow>cbv* e'" |
|
1248 using a by (auto dest: machines_implies_cbvs_ctx) |
|
1249 |
|
1250 text {* |
|
1251 We now want to relate the cbv-reduction to the evaluation |
|
1252 relation. For this we need two auxiliary lemmas. *} |
|
1253 |
|
1254 lemma eval_val: |
|
1255 assumes a: "val t" |
|
1256 shows "t \<Down> t" |
|
1257 using a by (induct) (auto) |
|
1258 |
|
1259 lemma e_App_elim: |
|
1260 assumes a: "App t1 t2 \<Down> v" |
|
1261 shows "\<exists>x t v'. t1 \<Down> Lam [x].t \<and> t2 \<Down> v' \<and> t[x::=v'] \<Down> v" |
|
1262 using a by (cases) (auto simp add: lam.inject) |
|
1263 |
|
1264 text {****************************************************************** |
|
1265 |
|
1266 10.) Exercise |
|
1267 ------------- |
|
1268 |
|
1269 Complete the first case in the proof below. |
|
1270 |
|
1271 *} |
|
1272 |
|
1273 lemma cbv_eval: |
|
1274 assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3" |
|
1275 shows "t1 \<Down> t3" |
|
1276 using a |
|
1277 proof(induct arbitrary: t3) |
|
1278 case (cbv1 v x t t3) |
|
1279 have a1: "val v" by fact |
|
1280 have a2: "t[x::=v] \<Down> t3" by fact |
|
1281 |
|
1282 show "App Lam [x].t v \<Down> t3" sorry |
|
1283 next |
|
1284 case (cbv2 t t' t2 t3) |
|
1285 have ih: "\<And>t3. t' \<Down> t3 \<Longrightarrow> t \<Down> t3" by fact |
|
1286 have "App t' t2 \<Down> t3" by fact |
|
1287 then obtain x t'' v' |
|
1288 where a1: "t' \<Down> Lam [x].t''" |
|
1289 and a2: "t2 \<Down> v'" |
|
1290 and a3: "t''[x::=v'] \<Down> t3" using e_App_elim by blast |
|
1291 have "t \<Down> Lam [x].t''" using ih a1 by auto |
|
1292 then show "App t t2 \<Down> t3" using a2 a3 by auto |
|
1293 qed (auto dest!: e_App_elim) |
|
1294 |
|
1295 |
|
1296 text {* |
|
1297 Next we extend the lemma above to arbitray initial |
|
1298 sequences of cbv-reductions. *} |
|
1299 |
|
1300 lemma cbvs_eval: |
|
1301 assumes a: "t1 \<longrightarrow>cbv* t2" "t2 \<Down> t3" |
|
1302 shows "t1 \<Down> t3" |
|
1303 using a by (induct) (auto intro: cbv_eval) |
|
1304 |
|
1305 text {* |
|
1306 Finally, we can show that if from a term t we reach a value |
|
1307 by a cbv-reduction sequence, then t evaluates to this value. *} |
|
1308 |
|
1309 lemma cbvs_implies_eval: |
|
1310 assumes a: "t \<longrightarrow>cbv* v" "val v" |
|
1311 shows "t \<Down> v" |
|
1312 using a |
|
1313 by (induct) (auto intro: eval_val cbvs_eval) |
|
1314 |
|
1315 text {* |
|
1316 All facts tied together give us the desired property about |
|
1317 K machines. *} |
|
1318 |
|
1319 theorem machines_implies_eval: |
|
1320 assumes a: "<t1,[]> \<mapsto>* <t2,[]>" |
|
1321 and b: "val t2" |
|
1322 shows "t1 \<Down> t2" |
|
1323 proof - |
|
1324 have "t1 \<longrightarrow>cbv* t2" using a by (simp add: machines_implies_cbvs) |
|
1325 then show "t1 \<Down> t2" using b by (simp add: cbvs_implies_eval) |
|
1326 qed |
|
1327 |
|
1328 text {****************************************************************** |
|
1329 |
|
1330 Formalising a Type-Soundness and Progress Lemma for CBV |
|
1331 ------------------------------------------------------- |
|
1332 |
|
1333 The central lemma for type-soundness is type-substitutity. In |
|
1334 our setting type-substitutivity is slightly painful to establish. |
|
1335 |
|
1336 *} |
|
1337 |
|
1338 lemma valid_elim: |
|
1339 assumes a: "valid ((x,T)#\<Gamma>)" |
|
1340 shows "x\<sharp>\<Gamma> \<and> valid \<Gamma>" |
|
1341 using a by (cases) (auto) |
|
1342 |
|
1343 lemma valid_insert: |
|
1344 assumes a: "valid (\<Delta>@[(x,T)]@\<Gamma>)" |
|
1345 shows "valid (\<Delta>@\<Gamma>)" |
|
1346 using a |
|
1347 by (induct \<Delta>) |
|
1348 (auto simp add: fresh_list_append fresh_list_cons dest!: valid_elim) |
|
1349 |
|
1350 lemma fresh_list: |
|
1351 shows "y\<sharp>xs = (\<forall>x\<in>set xs. y\<sharp>x)" |
|
1352 by (induct xs) (simp_all add: fresh_list_nil fresh_list_cons) |
|
1353 |
|
1354 lemma context_unique: |
|
1355 assumes a1: "valid \<Gamma>" |
|
1356 and a2: "(x,T) \<in> set \<Gamma>" |
|
1357 and a3: "(x,U) \<in> set \<Gamma>" |
|
1358 shows "T = U" |
|
1359 using a1 a2 a3 |
|
1360 by (induct) (auto simp add: fresh_list fresh_prod fresh_atm) |
|
1361 |
|
1362 lemma type_substitution_aux: |
|
1363 assumes a: "(\<Delta>@[(x,T')]@\<Gamma>) \<turnstile> e : T" |
|
1364 and b: "\<Gamma> \<turnstile> e' : T'" |
|
1365 shows "(\<Delta>@\<Gamma>) \<turnstile> e[x::=e'] : T" |
|
1366 using a b |
|
1367 proof (nominal_induct \<Gamma>'\<equiv>"\<Delta>@[(x,T')]@\<Gamma>" e T avoiding: x e' \<Delta> rule: typing.strong_induct) |
|
1368 case (t_Var \<Gamma>' y T x e' \<Delta>) |
|
1369 then have a1: "valid (\<Delta>@[(x,T')]@\<Gamma>)" |
|
1370 and a2: "(y,T) \<in> set (\<Delta>@[(x,T')]@\<Gamma>)" |
|
1371 and a3: "\<Gamma> \<turnstile> e' : T'" by simp_all |
|
1372 from a1 have a4: "valid (\<Delta>@\<Gamma>)" by (rule valid_insert) |
|
1373 { assume eq: "x=y" |
|
1374 from a1 a2 have "T=T'" using eq by (auto intro: context_unique) |
|
1375 with a3 have "\<Delta>@\<Gamma> \<turnstile> Var y[x::=e'] : T" using eq a4 by (auto intro: weakening) |
|
1376 } |
|
1377 moreover |
|
1378 { assume ineq: "x\<noteq>y" |
|
1379 from a2 have "(y,T) \<in> set (\<Delta>@\<Gamma>)" using ineq by simp |
|
1380 then have "\<Delta>@\<Gamma> \<turnstile> Var y[x::=e'] : T" using ineq a4 by auto |
|
1381 } |
|
1382 ultimately show "\<Delta>@\<Gamma> \<turnstile> Var y[x::=e'] : T" by blast |
|
1383 qed (force simp add: fresh_list_append fresh_list_cons)+ |
|
1384 |
|
1385 corollary type_substitution: |
|
1386 assumes a: "(x,T')#\<Gamma> \<turnstile> e : T" |
|
1387 and b: "\<Gamma> \<turnstile> e' : T'" |
|
1388 shows "\<Gamma> \<turnstile> e[x::=e'] : T" |
|
1389 using a b type_substitution_aux[where \<Delta>="[]"] |
|
1390 by (auto) |
|
1391 |
|
1392 lemma t_App_elim: |
|
1393 assumes a: "\<Gamma> \<turnstile> App t1 t2 : T" |
|
1394 shows "\<exists>T'. \<Gamma> \<turnstile> t1 : T' \<rightarrow> T \<and> \<Gamma> \<turnstile> t2 : T'" |
|
1395 using a |
|
1396 by (cases) (auto simp add: lam.inject) |
|
1397 |
|
1398 lemma t_Lam_elim: |
|
1399 assumes ty: "\<Gamma> \<turnstile> Lam [x].t : T" |
|
1400 and fc: "x\<sharp>\<Gamma>" |
|
1401 shows "\<exists>T1 T2. T = T1 \<rightarrow> T2 \<and> (x,T1)#\<Gamma> \<turnstile> t : T2" |
|
1402 using ty fc |
|
1403 by (cases rule: typing.strong_cases) |
|
1404 (auto simp add: alpha lam.inject abs_fresh ty_fresh) |
|
1405 |
|
1406 theorem cbv_type_preservation: |
|
1407 assumes a: "t \<longrightarrow>cbv t'" |
|
1408 and b: "\<Gamma> \<turnstile> t : T" |
|
1409 shows "\<Gamma> \<turnstile> t' : T" |
|
1410 using a b |
|
1411 by (nominal_induct avoiding: \<Gamma> T rule: cbv.strong_induct) |
|
1412 (auto dest!: t_Lam_elim t_App_elim simp add: type_substitution ty.inject) |
|
1413 |
|
1414 corollary cbvs_type_preservation: |
|
1415 assumes a: "t \<longrightarrow>cbv* t'" |
|
1416 and b: "\<Gamma> \<turnstile> t : T" |
|
1417 shows "\<Gamma> \<turnstile> t' : T" |
|
1418 using a b |
|
1419 by (induct) (auto intro: cbv_type_preservation) |
|
1420 |
|
1421 text {* |
|
1422 The Type-Preservation Property for the Machine and Evaluation Relation. *} |
|
1423 |
|
1424 theorem machine_type_preservation: |
|
1425 assumes a: "<t,[]> \<mapsto>* <t',[]>" |
|
1426 and b: "\<Gamma> \<turnstile> t : T" |
|
1427 shows "\<Gamma> \<turnstile> t' : T" |
|
1428 proof - |
|
1429 from a have "t \<longrightarrow>cbv* t'" by (simp add: machines_implies_cbvs) |
|
1430 then show "\<Gamma> \<turnstile> t' : T" using b by (simp add: cbvs_type_preservation) |
|
1431 qed |
|
1432 |
|
1433 theorem eval_type_preservation: |
|
1434 assumes a: "t \<Down> t'" |
|
1435 and b: "\<Gamma> \<turnstile> t : T" |
|
1436 shows "\<Gamma> \<turnstile> t' : T" |
|
1437 proof - |
|
1438 from a have "<t,[]> \<mapsto>* <t',[]>" by (simp add: eval_implies_machines) |
|
1439 then show "\<Gamma> \<turnstile> t' : T" using b by (simp add: machine_type_preservation) |
|
1440 qed |
|
1441 |
|
1442 text {* The Progress Property *} |
|
1443 |
|
1444 lemma canonical_tArr: |
|
1445 assumes a: "[] \<turnstile> t : T1 \<rightarrow> T2" |
|
1446 and b: "val t" |
|
1447 shows "\<exists>x t'. t = Lam [x].t'" |
|
1448 using b a by (induct) (auto) |
|
1449 |
|
1450 theorem progress: |
|
1451 assumes a: "[] \<turnstile> t : T" |
|
1452 shows "(\<exists>t'. t \<longrightarrow>cbv t') \<or> (val t)" |
|
1453 using a |
|
1454 by (induct \<Gamma>\<equiv>"[]::ty_ctx" t T) |
|
1455 (auto intro: cbv.intros dest!: canonical_tArr) |
|
1456 |
|
1457 |
|
1458 text {*********************************************************** |
|
1459 |
|
1460 Strong Induction Principle |
|
1461 -------------------------- |
|
1462 |
|
1463 A proof for the strong (structural) induction principle in the |
|
1464 lambda-calculus. |
|
1465 |
|
1466 *} |
|
1467 |
|
1468 |
|
1469 lemma lam_strong_induct: |
|
1470 fixes c::"'a::fs_name" |
|
1471 assumes h1: "\<And>x c. P c (Var x)" |
|
1472 and h2: "\<And>t1 t2 c. \<lbrakk>\<forall>d. P d t1; \<forall>d. P d t2\<rbrakk> \<Longrightarrow> P c (App t1 t2)" |
|
1473 and h3: "\<And>x t c. \<lbrakk>x\<sharp>c; \<forall>d. P d t\<rbrakk> \<Longrightarrow> P c (Lam [x].t)" |
|
1474 shows "P c t" |
|
1475 proof - |
|
1476 have "\<forall>(\<pi>::name prm) c. P c (\<pi>\<bullet>t)" |
|
1477 proof (induct t rule: lam.induct) |
|
1478 case (Lam x t) |
|
1479 have ih: "\<forall>(\<pi>::name prm) c. P c (\<pi>\<bullet>t)" by fact |
|
1480 { fix \<pi>::"name prm" and c::"'a::fs_name" |
|
1481 obtain y::"name" where fc: "y\<sharp>(\<pi>\<bullet>x,\<pi>\<bullet>t,c)" |
|
1482 by (rule exists_fresh) (auto simp add: fs_name1) |
|
1483 from ih have "\<forall>c. P c (([(y,\<pi>\<bullet>x)]@\<pi>)\<bullet>t)" by simp |
|
1484 then have "\<forall>c. P c ([(y,\<pi>\<bullet>x)]\<bullet>(\<pi>\<bullet>t))" by (auto simp only: pt_name2) |
|
1485 with h3 have "P c (Lam [y].[(y,\<pi>\<bullet>x)]\<bullet>(\<pi>\<bullet>t))" using fc by (simp add: fresh_prod) |
|
1486 moreover |
|
1487 have "Lam [y].[(y,\<pi>\<bullet>x)]\<bullet>(\<pi>\<bullet>t) = Lam [(\<pi>\<bullet>x)].(\<pi>\<bullet>t)" |
|
1488 using fc by (simp add: lam.inject alpha fresh_atm fresh_prod) |
|
1489 ultimately have "P c (Lam [(\<pi>\<bullet>x)].(\<pi>\<bullet>t))" by simp |
|
1490 } |
|
1491 then have "\<forall>(\<pi>::name prm) c. P c (Lam [(\<pi>\<bullet>x)].(\<pi>\<bullet>t))" by simp |
|
1492 then show "\<forall>(\<pi>::name prm) c. P c (\<pi>\<bullet>(Lam [x].t))" by simp |
|
1493 qed (auto intro: h1 h2) (* var and app case *) |
|
1494 then have "P c (([]::name prm)\<bullet>t)" by blast |
|
1495 then show "P c t" by simp |
|
1496 qed |
|
1497 |
|
1498 text {*********************************************************** |
|
1499 |
|
1500 --------- |
|
1501 SOLUTIONS |
|
1502 --------- |
|
1503 |
|
1504 *} |
|
1505 |
|
1506 text {************************************************************ |
|
1507 |
|
1508 1.) MINI EXERCISE |
|
1509 |
|
1510 The way we defined contexts does not allow us to |
|
1511 apply a Hole to a Hole. Therefore the following |
|
1512 will result in a typing error. *} |
|
1513 |
|
1514 (* term "CAppL \<box> \<box>" *) |
|
1515 |
|
1516 |
|
1517 text {************************************************************ |
|
1518 |
|
1519 2. EXERCISE |
|
1520 |
|
1521 A readable proof for this lemma is as follows: |
|
1522 |
|
1523 *} |
|
1524 |
|
1525 lemma |
|
1526 assumes a: "<e1,Es1> \<mapsto>* <e2,Es2>" |
|
1527 and b: "<e2,Es2> \<mapsto>* <e3,Es3>" |
|
1528 shows "<e1,Es1> \<mapsto>* <e3,Es3>" |
|
1529 using a b |
|
1530 proof(induct) |
|
1531 case (ms1 e1 Es1) |
|
1532 show "<e1,Es1> \<mapsto>* <e3,Es3>" by fact |
|
1533 next |
|
1534 case (ms2 e1 Es1 e2 Es2 e2' Es2') |
|
1535 have ih: "<e2',Es2'> \<mapsto>* <e3,Es3> \<Longrightarrow> <e2,Es2> \<mapsto>* <e3,Es3>" by fact |
|
1536 have "<e2',Es2'> \<mapsto>* <e3,Es3>" by fact |
|
1537 then have "<e2,Es2> \<mapsto>* <e3,Es3>" using ih by simp |
|
1538 moreover |
|
1539 have "<e1,Es1> \<mapsto> <e2,Es2>" by fact |
|
1540 ultimately show "<e1,Es1> \<mapsto>* <e3,Es3>" by auto |
|
1541 qed |
|
1542 |
|
1543 |
|
1544 text {************************************************************ |
|
1545 |
|
1546 3.) Exercise |
|
1547 |
|
1548 As one can quickly see in the second case, the theorem as stated |
|
1549 does not go through. We need to generalise the induction hypothesis |
|
1550 so that we show the lemma for all contexts Es. In Isar, variables |
|
1551 can be generalised by declaring "arbitrary: variable \<dots>" when the |
|
1552 induction is set up. |
|
1553 |
|
1554 *} |
|
1555 |
|
1556 theorem |
|
1557 assumes a: "t \<Down> t'" |
|
1558 shows "<t,Es> \<mapsto>* <t',Es>" |
|
1559 using a |
|
1560 proof (induct arbitrary: Es) (* here we generalise over Es *) |
|
1561 case (e_Lam x t Es) |
|
1562 show "<Lam [x].t,Es> \<mapsto>* <Lam [x].t,Es>" by auto |
|
1563 next |
|
1564 case (e_App t1 x t t2 v' v Es) |
|
1565 have ih1: "\<And>Es. <t1,Es> \<mapsto>* <Lam [x].t,Es>" by fact |
|
1566 have ih2: "\<And>Es. <t2,Es> \<mapsto>* <v',Es>" by fact |
|
1567 have ih3: "\<And>Es. <t[x::=v'],Es> \<mapsto>* <v,Es>" by fact |
|
1568 have "<App t1 t2,Es> \<mapsto>* <t1,CAppL \<box> t2#Es>" by auto |
|
1569 moreover |
|
1570 have "<t1,CAppL \<box> t2#Es> \<mapsto>* <Lam [x].t,CAppL \<box> t2#Es>" using ih1 by auto |
|
1571 moreover |
|
1572 have "<Lam [x].t,CAppL \<box> t2#Es> \<mapsto>* <t2,CAppR (Lam [x].t) \<box>#Es>" by auto |
|
1573 moreover |
|
1574 have "<t2,CAppR (Lam [x].t) \<box>#Es> \<mapsto>* <v',CAppR (Lam [x].t) \<box>#Es>" using ih2 by auto |
|
1575 moreover |
|
1576 have "t2 \<Down> v'" by fact |
|
1577 then have "val v'" using eval_to_val by auto |
|
1578 then have "<v',CAppR (Lam [x].t) \<box>#Es> \<mapsto>* <t[x::=v'],Es>" by auto |
|
1579 moreover |
|
1580 have "<t[x::=v'],Es> \<mapsto>* <v,Es>" using ih3 by auto |
|
1581 ultimately show "<App t1 t2,Es> \<mapsto>* <v,Es>" by blast |
|
1582 qed |
|
1583 |
|
1584 text {************************************************************ |
|
1585 |
|
1586 4.) Exercise |
|
1587 |
|
1588 A proof for the weakening lemma: |
|
1589 |
|
1590 *} |
|
1591 |
|
1592 lemma |
|
1593 fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list" |
|
1594 assumes a: "\<Gamma>1 \<turnstile> t : T" |
|
1595 and b: "valid \<Gamma>2" |
|
1596 and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
|
1597 shows "\<Gamma>2 \<turnstile> t : T" |
|
1598 using a b c |
|
1599 proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct) |
|
1600 case (t_Var \<Gamma>1 x T) (* variable case *) |
|
1601 have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
|
1602 moreover |
|
1603 have "valid \<Gamma>2" by fact |
|
1604 moreover |
|
1605 have "(x,T)\<in> set \<Gamma>1" by fact |
|
1606 ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto |
|
1607 next |
|
1608 case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *) |
|
1609 have vc: "x\<sharp>\<Gamma>2" by fact (* variable convention *) |
|
1610 have ih: "\<lbrakk>valid ((x,T1)#\<Gamma>2); (x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2\<rbrakk> \<Longrightarrow> (x,T1)#\<Gamma>2 \<turnstile> t:T2" by fact |
|
1611 have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact |
|
1612 then have "(x,T1)#\<Gamma>1 \<subseteq> (x,T1)#\<Gamma>2" by simp |
|
1613 moreover |
|
1614 have "valid \<Gamma>2" by fact |
|
1615 then have "valid ((x,T1)#\<Gamma>2)" using vc by (simp add: v2) |
|
1616 ultimately have "(x,T1)#\<Gamma>2 \<turnstile> t : T2" using ih by simp |
|
1617 with vc show "\<Gamma>2 \<turnstile> Lam [x].t : T1\<rightarrow>T2" by auto |
|
1618 qed (auto) (* app case *) |
|
1619 |
|
1620 text {************************************************************ |
|
1621 |
|
1622 5.) Exercise |
|
1623 |
|
1624 A proof for context omposition |
|
1625 |
|
1626 *} |
|
1627 |
|
1628 lemma |
|
1629 shows "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" |
|
1630 by (induct E1) (simp_all) |
|
1631 |
|
1632 text {****************************************************************** |
|
1633 |
|
1634 6.) EXERCISE |
|
1635 ------------ |
|
1636 |
|
1637 A proof for the assoiativity of \<circ>. |
|
1638 |
|
1639 *} |
|
1640 |
|
1641 lemma |
|
1642 shows "(Es1@Es2)\<down> = (Es2\<down>) \<circ> (Es1\<down>)" |
|
1643 proof (induct Es1) |
|
1644 case Nil |
|
1645 show "([]@Es2)\<down> = Es2\<down> \<circ> []\<down>" using neut_hole by simp |
|
1646 next |
|
1647 case (Cons E Es1) |
|
1648 have ih: "(Es1@Es2)\<down> = Es2\<down> \<circ> Es1\<down>" by fact |
|
1649 have "((E#Es1)@Es2)\<down> = (Es1@Es2)\<down> \<circ> E" by simp |
|
1650 also have "\<dots> = (Es2\<down> \<circ> Es1\<down>) \<circ> E" using ih by simp |
|
1651 also have "\<dots> = Es2\<down> \<circ> (Es1\<down> \<circ> E)" using circ_assoc by simp |
|
1652 also have "\<dots> = Es2\<down> \<circ> (E#Es1)\<down>" by simp |
|
1653 finally show "((E#Es1)@Es2)\<down> = Es2\<down> \<circ> (E#Es1)\<down>" by simp |
|
1654 qed |
|
1655 |
|
1656 text {****************************************************************** |
|
1657 |
|
1658 7.) EXERCISE |
|
1659 ------------ |
|
1660 |
|
1661 A proof for the substitution lemma. |
|
1662 |
|
1663 *} |
|
1664 |
|
1665 lemma |
|
1666 assumes a: "x\<noteq>y" |
|
1667 and b: "x\<sharp>L" |
|
1668 shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" |
|
1669 using a b |
|
1670 proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct) |
|
1671 case (Var z) (* case 1: Variables*) |
|
1672 show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS") |
|
1673 proof - |
|
1674 { (*Case 1.1*) |
|
1675 assume "z=x" |
|
1676 have "(1)": "?LHS = N[y::=L]" using `z=x` by simp |
|
1677 have "(2)": "?RHS = N[y::=L]" using `z=x` `x\<noteq>y` by simp |
|
1678 from "(1)" "(2)" have "?LHS = ?RHS" by simp |
|
1679 } |
|
1680 moreover |
|
1681 { (*Case 1.2*) |
|
1682 assume "z=y" and "z\<noteq>x" |
|
1683 have "(1)": "?LHS = L" using `z\<noteq>x` `z=y` by simp |
|
1684 have "(2)": "?RHS = L[x::=N[y::=L]]" using `z=y` by simp |
|
1685 have "(3)": "L[x::=N[y::=L]] = L" using `x\<sharp>L` by (simp add: forget) |
|
1686 from "(1)" "(2)" "(3)" have "?LHS = ?RHS" by simp |
|
1687 } |
|
1688 moreover |
|
1689 { (*Case 1.3*) |
|
1690 assume "z\<noteq>x" and "z\<noteq>y" |
|
1691 have "(1)": "?LHS = Var z" using `z\<noteq>x` `z\<noteq>y` by simp |
|
1692 have "(2)": "?RHS = Var z" using `z\<noteq>x` `z\<noteq>y` by simp |
|
1693 from "(1)" "(2)" have "?LHS = ?RHS" by simp |
|
1694 } |
|
1695 ultimately show "?LHS = ?RHS" by blast |
|
1696 qed |
|
1697 next |
|
1698 case (Lam z M1) (* case 2: lambdas *) |
|
1699 have ih: "\<lbrakk>x\<noteq>y; x\<sharp>L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact |
|
1700 have fs: "z\<sharp>x" "z\<sharp>y" "z\<sharp>N" "z\<sharp>L" by fact+ |
|
1701 then have "z\<sharp>N[y::=L]" by (simp add: fresh_fact) |
|
1702 show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") |
|
1703 proof - |
|
1704 have "?LHS = Lam [z].(M1[x::=N][y::=L])" using `z\<sharp>x` `z\<sharp>y` `z\<sharp>N` `z\<sharp>L` by simp |
|
1705 also from ih have "\<dots> = Lam [z].(M1[y::=L][x::=N[y::=L]])" using `x\<noteq>y` `x\<sharp>L` by simp |
|
1706 also have "\<dots> = (Lam [z].(M1[y::=L]))[x::=N[y::=L]]" using `z\<sharp>x` `z\<sharp>N[y::=L]` by simp |
|
1707 also have "\<dots> = ?RHS" using `z\<sharp>y` `z\<sharp>L` by simp |
|
1708 finally show "?LHS = ?RHS" by simp |
|
1709 qed |
|
1710 next |
|
1711 case (App M1 M2) (* case 3: applications *) |
|
1712 then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp |
|
1713 qed |
|
1714 |
|
1715 text {****************************************************************** |
|
1716 |
|
1717 8.) Exercise |
|
1718 ------------ |
|
1719 |
|
1720 Left out if not needed. |
|
1721 *} |
|
1722 |
|
1723 lemma |
|
1724 assumes a: "t \<longrightarrow>cbv t'" |
|
1725 shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" |
|
1726 using a |
|
1727 proof (induct E) |
|
1728 case Hole |
|
1729 have "t \<longrightarrow>cbv t'" by fact |
|
1730 then show "\<box>\<lbrakk>t\<rbrakk> \<longrightarrow>cbv \<box>\<lbrakk>t'\<rbrakk>" by simp |
|
1731 next |
|
1732 case (CAppL E s) |
|
1733 have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact |
|
1734 have a: "t \<longrightarrow>cbv t'" by fact |
|
1735 show "(CAppL E s)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppL E s)\<lbrakk>t'\<rbrakk>" using ih a by auto |
|
1736 next |
|
1737 case (CAppR s E) |
|
1738 have ih: "t \<longrightarrow>cbv t' \<Longrightarrow> E\<lbrakk>t\<rbrakk> \<longrightarrow>cbv E\<lbrakk>t'\<rbrakk>" by fact |
|
1739 have a: "t \<longrightarrow>cbv t'" by fact |
|
1740 show "(CAppR s E)\<lbrakk>t\<rbrakk> \<longrightarrow>cbv (CAppR s E)\<lbrakk>t'\<rbrakk>" using ih a by auto |
|
1741 qed |
|
1742 |
|
1743 text {****************************************************************** |
|
1744 |
|
1745 9.) Exercise |
|
1746 ------------ |
|
1747 |
|
1748 The point of the cbv-reduction was that we can easily relatively |
|
1749 establish the follwoing property: |
|
1750 |
|
1751 *} |
|
1752 |
|
1753 lemma |
|
1754 assumes a: "<e,Es> \<mapsto> <e',Es'>" |
|
1755 shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>" |
|
1756 using a |
|
1757 proof (induct) |
|
1758 case (m1 t1 t2 Es) |
|
1759 show "Es\<down>\<lbrakk>App t1 t2\<rbrakk> \<longrightarrow>cbv* ((CAppL \<box> t2)#Es)\<down>\<lbrakk>t1\<rbrakk>" by (auto simp add: ctx_compose) |
|
1760 next |
|
1761 case (m2 v t2 Es) |
|
1762 have "val v" by fact |
|
1763 then show "((CAppL \<box> t2)#Es)\<down>\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (CAppR v \<box> # Es)\<down>\<lbrakk>t2\<rbrakk>" |
|
1764 by (auto simp add: ctx_compose) |
|
1765 next |
|
1766 case (m3 v x t Es) |
|
1767 have "val v" by fact |
|
1768 then show "((CAppR (Lam [x].t) \<box>)#Es)\<down>\<lbrakk>v\<rbrakk> \<longrightarrow>cbv* (Es\<down>)\<lbrakk>t[x::=v]\<rbrakk>" |
|
1769 by (auto simp add: ctx_compose intro: cbv_in_ctx) |
|
1770 qed |
|
1771 |
|
1772 lemma |
|
1773 assumes a: "<e,Es> \<mapsto> <e',Es'>" |
|
1774 shows "(Es\<down>)\<lbrakk>e\<rbrakk> \<longrightarrow>cbv* (Es'\<down>)\<lbrakk>e'\<rbrakk>" |
|
1775 using a by (induct) (auto simp add: ctx_compose intro: cbv_in_ctx) |
|
1776 |
|
1777 text {****************************************************************** |
|
1778 |
|
1779 10.) Exercise |
|
1780 ------------- |
|
1781 |
|
1782 Complete the first case in the proof below. |
|
1783 |
|
1784 *} |
|
1785 |
|
1786 lemma |
|
1787 assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3" |
|
1788 shows "t1 \<Down> t3" |
|
1789 using a |
|
1790 proof(induct arbitrary: t3) |
|
1791 case (cbv1 v x t t3) |
|
1792 have a1: "val v" by fact |
|
1793 have a2: "t[x::=v] \<Down> t3" by fact |
|
1794 show "App Lam [x].t v \<Down> t3" using eval_val a1 a2 by auto |
|
1795 next |
|
1796 case (cbv2 t t' t2 t3) |
|
1797 have "t \<longrightarrow>cbv t'" by fact |
|
1798 have ih: "\<And>t3. t' \<Down> t3 \<Longrightarrow> t \<Down> t3" by fact |
|
1799 have "App t' t2 \<Down> t3" by fact |
|
1800 then obtain x t'' v' |
|
1801 where a1: "t' \<Down> Lam [x].t''" |
|
1802 and a2: "t2 \<Down> v'" |
|
1803 and a3: "t''[x::=v'] \<Down> t3" using e_App_elim by blast |
|
1804 have "t \<Down> Lam [x].t''" using ih a1 by auto |
|
1805 then show "App t t2 \<Down> t3" using a2 a3 by auto |
|
1806 qed (auto dest!: e_App_elim) |
|
1807 |
|
1808 lemma |
|
1809 assumes a: "t1 \<longrightarrow>cbv t2" "t2 \<Down> t3" |
|
1810 shows "t1 \<Down> t3" |
|
1811 using a |
|
1812 by (induct arbitrary: t3) |
|
1813 (auto elim!: eval_elim intro: eval_val) |
|
1814 |
|
1815 |
|
1816 end |
|
1817 |
|
1818 |
|
1819 |
|
1820 |
|
1821 |
|
1822 |
|
1823 |
|
1824 |
|
1825 |
|
1826 |