diff -r 05e5d68c9627 -r f1be8028a4a9 Nominal/activities/CK_Machine.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/activities/CK_Machine.thy Wed Mar 30 17:27:34 2016 +0100 @@ -0,0 +1,1826 @@ +(***************************************************************** + + Nominal Isabelle Tutorial + ------------------------- + + 11th August 2008, Sydney + + This file contains most of the material that will be covered in the + tutorial. The file can be "stepped through"; though it contains much + commented code (purple text). + +*) + +(***************************************************************** + + Proof General + ------------- + + Proof General is the front end for Isabelle. + + To run Nominal Isabelle proof-scripts you must have HOL-Nominal enabled in + the menu Isabelle -> Logics. You also need to enable X-Symbols in the menu + Proof-General -> Options (make sure to save this option once enabled). + + Proof General "paints" blue the part of the proof-script that has already + been processed by Isabelle. You can advance or retract the "frontier" of + this processed part using the "Next" and "Undo" buttons in the + menubar. "Goto" will process everything up to the current cursor position, + or retract if the cursor is inside the blue part. The key-combination + control-c control-return is a short-cut for the "Goto"-operation. + + Proof General gives feedback inside the "Response" and "Goals" buffers. + The response buffer will contain warning messages (in yellow) and + error messages (in red). Warning messages can generally be ignored. Error + messages must be dealt with in order to advance the proof script. The goal + buffer shows which goals need to be proved. The sole idea of interactive + theorem proving is to get the message "No subgoals." ;o) + +*) + +(***************************************************************** + + X-Symbols + --------- + + X-symbols provide a nice way to input non-ascii characters. For example: + + \, \, \, \, \, \, \, \, \, \, \ + + They need to be input via the combination \ + where name-of-x-symbol coincides with the usual latex name of + that symbol. + + However, there are some handy short-cuts for frequently used X-symbols. + For example + + [| \ \ + |] \ \ + ==> \ \ + => \ \ + --> \ \ + /\ \ \ + \/ \ \ + |-> \ \ + =_ \ \ + +*) + +(***************************************************************** + + Theories + -------- + + Every Isabelle proof-script needs to have a name and must import + some pre-existing theory. For Nominal Isabelle proof-scripts this will + normally be the theory Nominal, but we use here the theory Lambda.thy, + which extends Nominal with a definition for lambda-terms and capture- + avoiding substitution. + + BTW, the Nominal theory builds directly on Isabelle/HOL and extends it + only with some definitions and some reasoning infrastructure. It does not + add any new axiom to Isabelle/HOL. So you can trust what you are doing. ;o) + +*) + +theory CK_Machine + imports "Lambda" +begin + +text {***************************************************************** + + Types + ----- + + Isabelle is based, roughly, on the theory of simple types including some + polymorphism. It also includes some overloading, which means that sometimes + explicit type annotations need to be given. + + - Base types include: nat, bool, string, lam + + - Type formers include: 'a list, ('a\'b), 'c set + + - Type variables are written like in ML with an apostrophe: 'a, 'b, \ + + Types known to Isabelle can be queried using: + +*} + +typ nat +typ bool +typ lam (* the type for alpha-equated lambda-terms *) +typ "('a \ 'b)" (* product type *) +typ "'c set" (* set type *) +typ "nat \ bool" (* the type for functions from nat to bool *) + +(* These give errors: *) +(*typ boolean *) +(*typ set*) + + +text {***************************************************************** + + Terms + ----- + + Every term in Isabelle needs to be well-typed (however they can have polymorphic + type). Whether a term is accepted can be queried using: + +*} + +term c (* a variable of polymorphic type *) +term "1::nat" (* the constant 1 in natural numbers *) +term 1 (* the constant 1 with polymorphic type *) +term "{1, 2, 3::nat}" (* the set containing natural numbers 1, 2 and 3 *) +term "[1, 2, 3]" (* the list containing the polymorphic numbers 1, 2 and 3 *) +term "Lam [x].(Var x)" (* a lambda-term *) +term "App t1 t2" (* another lambda-term *) + +text {* + Isabelle provides some useful colour feedback about what are constants (in black), + free variables (in blue) and bound variables (in green). *} + +term "True" (* a constant that is defined in HOL *) +term "true" (* not recognised as a constant, therefore it is interpreted as a variable *) +term "\x. P x" (* x is bound, P is free *) + +text {* Every formula in Isabelle needs to have type bool *} + +term "True" +term "True \ False" +term "{1,2,3} = {3,2,1}" +term "\x. P x" +term "A \ B" + +text {* + When working with Isabelle, one is confronted with an object logic (that is HOL) + and Isabelle's meta-logic (called Pure). Sometimes one has to pay attention + to this fact. *} + +term "A \ B" (* versus *) +term "A \ B" + +term "\x. P x" (* versus *) +term "\x. P x" + +term "A \ B \ C" (* is synonymous with *) +term "\A; B\ \ C" + + +text {***************************************************************** + + Inductive Definitions: The Evaluation Judgement and the Value Predicate + ----------------------------------------------------------------------- + + Inductive definitions start with the keyword "inductive" and a predicate name. + One also needs to provide a type for the predicate. Clauses of the inductive + predicate are introduced by "where" and more than two clauses need to be + separated by "|". Optionally one can give a name to each clause and indicate + that it should be added to the hints database ("[intro]"). A typical clause + has some premises and a conclusion. This is written in Isabelle as: + + "premise \ conclusion" + "\premise1; premise2; \ \ \ conclusion" + + If no premise is present, then one just writes + + "conclusion" + + Below we define the evaluation judgement for lambda-terms. This definition + introduces the predicate named "eval". After giving its type, we declare + the usual pretty syntax _ \ _. In this declaration _ stands for an argument + of eval. + +*} + +inductive + eval :: "lam \ lam \ bool" ("_ \ _") +where + e_Lam: "Lam [x].t \ Lam [x].t" +| e_App: "\t1 \ Lam [x].t; t2 \ v'; t[x::=v'] \ v\ \ App t1 t2 \ v" + +declare eval.intros[intro] + +text {* + Values are also defined using inductive. In our case values + are just lambda-abstractions. *} + +inductive + val :: "lam \ bool" +where + v_Lam[intro]: "val (Lam [x].t)" + + +text {***************************************************************** + + Theorems + -------- + + A central concept in Isabelle is that of theorems. Isabelle's theorem + database can be queried using + +*} + +thm e_Lam +thm e_App +thm conjI +thm conjunct1 + +text {* + Notice that theorems usually contain schematic variables (e.g. ?P, ?Q, \). + These schematic variables can be substituted with any term (of the right type + of course). It is sometimes useful to suppress the "?" from the schematic + variables. This can be achieved by using the attribute "[no_vars]". *} + +thm e_Lam[no_vars] +thm e_App[no_vars] +thm conjI[no_vars] +thm conjunct1[no_vars] + + +text {* + When defining the predicate eval, Isabelle provides us automatically with + the following theorems that state how evaluation judgments can be introduced + and what constitutes an induction over the predicate eval. *} + +thm eval.intros[no_vars] +thm eval.induct[no_vars] + +text {***************************************************************** + + Lemma / Theorem / Corollary Statements + -------------------------------------- + + Whether to use lemma, theorem or corollary makes no semantic difference + in Isabelle. A lemma starts with "lemma" and consists of a statement + ("shows \") and optionally a lemma name, some type-information for + variables ("fixes \") and some assumptions ("assumes \"). Lemmas + also need to have a proof, but ignore this 'detail' for the moment. + +*} + +lemma alpha_equ: + shows "Lam [x].Var x = Lam [y].Var y" +by (simp add: lam.inject alpha swap_simps fresh_atm) + +lemma Lam_freshness: + assumes a: "x\y" + shows "y\Lam [x].t \ y\t" +using a by (simp add: abs_fresh) + +lemma neutral_element: + fixes x::"nat" + shows "x + 0 = x" +by simp + + +text {***************************************************************** + + Datatypes: Evaluation Contexts + ------------------------------ + + Datatypes can be defined in Isabelle as follows: we have to provide the name + of the datatype and list its type-constructors. Each type-constructor needs + to have the information about the types of its arguments, and optionally + can also contain some information about pretty syntax. For example, we like + to write "\" for holes. + +*} + +datatype ctx = + Hole ("\") + | CAppL "ctx" "lam" + | CAppR "lam" "ctx" + +text {* Now Isabelle knows about: *} + +typ ctx +term "\" +term "CAppL" +term "CAppL \ (Var x)" + +text {* + + 1.) MINI EXERCISE + ----------------- + + Try and see what happens if you apply a Hole to a Hole? + +*} + + +text {***************************************************************** + + The CK Machine + -------------- + + The CK machine is also defined using an inductive predicate with four + arguments. The idea behind this abstract machine is to transform, or reduce, + a configuration consisting of a lambda-term and a framestack (a list of + contexts), to a new configuration. + + We use the type abbreviation "ctxs" for the type for framestacks. + + The pretty syntax for the CK machine is <_,_> \ <_,_>. + +*} + +types ctxs = "ctx list" + +inductive + machine :: "lam \ ctxs \lam \ ctxs \ bool" ("<_,_> \ <_,_>") +where + m1[intro]: " \ t2)#Es>" +| m2[intro]: "val v \ t2)#Es> \ )#Es>" +| m3[intro]: "val v \ )#Es> \ " + + +text {* + Since the machine defined above only performs a single reduction, + we need to define the transitive closure of this machine. *} + +inductive + machines :: "lam \ ctxs \ lam \ ctxs \ bool" ("<_,_> \* <_,_>") +where + ms1[intro]: " \* " +| ms2[intro]: "\ \ ; \* \ \ \* " + + +text {***************************************************************** + + Isar Proofs + ----------- + + Isar is a language for writing down proofs that can be understood by humans + and by Isabelle. An Isar proof can be thought of as a sequence of 'stepping stones' + that start with the assumptions and lead to the goal to be established. Every such + stepping stone is introduced by "have" followed by the statement of the stepping + stone. An exception is the goal to be proved, which need to be introduced with "show". + + Since proofs usually do not proceed in a linear fashion, a label can be given + to each stepping stone and then used later to refer to this stepping stone + ("using"). + + Each stepping stone (or have-statement) needs to have a justification. The + simplest justification is "sorry" which admits any stepping stone, even false + ones (this is good during the development of proofs). Assumption can be + "justified" using "by fact". Derived facts can be justified using + + - by simp (* simplification *) + - by auto (* proof search and simplification *) + - by blast (* only proof search *) + + If facts or lemmas are needed in order to justify a have-statement, then + one can feed these facts into the proof by using "using label \" or + "using theorem-name \". More than one label at the time is allowed. + + Induction proofs in Isar are set up by indicating over which predicate(s) + the induction proceeds ("using a b") followed by the command "proof (induct)". + In this way, Isabelle uses default settings for which induction should + be performed. These default settings can be overridden by giving more + information, like the variable over which a structural induction should + proceed, or a specific induction principle such as well-founded inductions. + + After the induction is set up, the proof proceeds by cases. In Isar these + cases can be given in any order, but must be started with "case" and the + name of the case, and optionally some legible names for the variables + referred to inside the case. + + The possible case-names can be found by looking inside the menu "Isabelle -> + Show me -> cases". In each "case", we need to establish a statement introduced + by "show". Once this has been done, the next case can be started using "next". + When all cases are completed, the proof can be finished using "qed". + + This means a typical induction proof has the following pattern + + proof (induct) + case \ + \ + show \ + next + case \ + \ + show \ + \ + qed + + The four lemmas are by induction on the predicate machines. All proofs establish + the same property, namely a transitivity rule for machines. The complete Isar + proofs are given for the first three proofs. The point of these three proofs is + that each proof increases the readability for humans. + +*} + +text {***************************************************************** + + 2.) EXERCISE + ------------ + + Remove the sorries in the proof below and fill in the correct + justifications. +*} + +lemma + assumes a: " \* " + and b: " \* " + shows " \* " +using a b +proof(induct) + case (ms1 e1 Es1) + have c: " \* " by fact + show " \* " sorry +next + case (ms2 e1 Es1 e2 Es2 e2' Es2') + have ih: " \* \ \* " by fact + have d1: " \* " by fact + have d2: " \ " by fact + + show " \* " sorry +qed + +text {* + Just like gotos in the Basic programming language, labels can reduce + the readability of proofs. Therefore one can use in Isar the notation + "then have" in order to feed a have-statement to the proof of + the next have-statement. In the proof below this has been used + in order to get rid of the labels c and d1. +*} + +lemma + assumes a: " \* " + and b: " \* " + shows " \* " +using a b +proof(induct) + case (ms1 e1 Es1) + show " \* " by fact +next + case (ms2 e1 Es1 e2 Es2 e2' Es2') + have ih: " \* \ \* " by fact + have " \* " by fact + then have d3: " \* " using ih by simp + have d2: " \ " by fact + show " \* " using d2 d3 by auto +qed + +text {* + The labels d2 and d3 cannot be got rid of in this way, because both + facts are needed to prove the show-statement. We can still avoid the + labels by feeding a sequence of facts into a proof using the chaining + mechanism: + + have "statement1" \ + moreover + have "statement2" \ + \ + moreover + have "statementn" \ + ultimately have "statement" \ + + In this chain, all "statementi" can be used in the proof of the final + "statement". With this we can simplify our proof further to: +*} + +lemma + assumes a: " \* " + and b: " \* " + shows " \* " +using a b +proof(induct) + case (ms1 e1 Es1) + show " \* " by fact +next + case (ms2 e1 Es1 e2 Es2 e2' Es2') + have ih: " \* \ \* " by fact + have " \* " by fact + then have " \* " using ih by simp + moreover + have " \ " by fact + ultimately show " \* " by auto +qed + +text {* + While automatic proof procedures in Isabelle are not able to prove statements + like "P = NP" assuming usual definitions for P and NP, they can automatically + discharge the lemma we just proved. For this we only have to set up the induction + and auto will take care of the rest. This means we can write: +*} + +lemma ms3[intro]: + assumes a: " \* " + and b: " \* " + shows " \* " +using a b by (induct) (auto) + +text {* + The attribute [intro] indicates that this lemma should be from now on used in + any proof obtained by "auto" or "blast". +*} + + +text {***************************************************************** + + A simple fact we need later on is that if t \ t' then t' is a value. + +*} + +lemma eval_to_val: + assumes a: "t \ t'" + shows "val t'" +using a by (induct) (auto) + + +text {***************************************************************** + + 3.) EXERCISE + ------------ + + Fill in the details in the proof below. The proof will establish the fact + that if t \ t' then \* . As can be seen, the proof is by + induction on the definition of eval. If you want to know how the predicates + machine and machines can be introduced, then use + + thm machine.intros[no_vars] + thm machines.intros[no_vars] + +*} + +theorem + assumes a: "t \ t'" + shows " \* " +using a +proof (induct) + case (e_Lam x t) + (* no assumptions *) + show " \* " sorry +next + case (e_App t1 x t t2 v' v) + (* all assumptions in this case *) + have a1: "t1 \ Lam [x].t" by fact + have ih1: " \* " by fact + have a2: "t2 \ v'" by fact + have ih2: " \* " by fact + have a3: "t[x::=v'] \ v" by fact + have ih3: " \* " by fact + (* your details *) + show " \* " sorry +qed + +text {* + Again the automatic tools in Isabelle can discharge automatically + of the routine work in these proofs. We can write: *} + +theorem eval_implies_machines_ctx: + assumes a: "t \ t'" + shows " \* " +using a +by (induct arbitrary: Es) + (metis eval_to_val machine.intros ms1 ms2 ms3 v_Lam)+ + +corollary eval_implies_machines: + assumes a: "t \ t'" + shows " \* " +using a eval_implies_machines_ctx by simp + +text {***************************************************************** + + The Weakening Lemma + ------------------- + + The proof of the weakening lemma is often said to be simple, + routine or trivial. Below we will see how this lemma can be + proved in Nominal Isabelle. First we define types, which + we however do not define as datatypes, but as nominal datatypes. + +*} + +nominal_datatype ty = + tVar "string" +| tArr "ty" "ty" ("_ \ _") + +text {* + Having defined them as nominal datatypes gives us additional + definitions and theorems that come with types (see below). + *} + +text {* + We next define the type of typing contexts, a predicate that + defines valid contexts (i.e. lists that contain only unique + variables) and the typing judgement. + +*} + +types ty_ctx = "(name\ty) list" + +inductive + valid :: "ty_ctx \ bool" +where + v1[intro]: "valid []" +| v2[intro]: "\valid \; x\\\\ valid ((x,T)#\)" + +inductive + typing :: "ty_ctx \ lam \ ty \ bool" ("_ \ _ : _") +where + t_Var[intro]: "\valid \; (x,T)\set \\ \ \ \ Var x : T" +| t_App[intro]: "\\ \ t1 : T1\T2; \ \ t2 : T1\ \ \ \ App t1 t2 : T2" +| t_Lam[intro]: "\x\\; (x,T1)#\ \ t : T2\ \ \ \ Lam [x].t : T1 \ T2" + + +text {* + The predicate x\\, read as x fresh for \, is defined by Nominal Isabelle. + Freshness is defined for lambda-terms, products, lists etc. For example + we have: + *} + +lemma + fixes x::"name" + shows "x\Lam [x].t" + and "x\(t1,t2) \ x\App t1 t2" + and "x\(Var y) \ x\y" + and "\x\t1; x\t2\ \ x\(t1,t2)" + and "\x\l1; x\l2\ \ x\(l1@l2)" + and "x\y \ x\y" +by (simp_all add: abs_fresh fresh_prod fresh_list_append fresh_atm) + +text {* We can also prove that every variable is fresh for a type *} + +lemma ty_fresh: + fixes x::"name" + and T::"ty" + shows "x\T" +by (induct T rule: ty.induct) + (simp_all add: fresh_string) + +text {* + In order to state the weakening lemma in a convenient form, we overload + the subset-notation and define the abbreviation below. Abbreviations behave + like definitions, except that they are always automatically folded and + unfolded. +*} + +abbreviation + "sub_ty_ctx" :: "ty_ctx \ ty_ctx \ bool" ("_ \ _") +where + "\1 \ \2 \ \x. x \ set \1 \ x \ set \2" + +text {***************************************************************** + + 4.) Exercise + ------------ + + Fill in the details and give a proof of the weakening lemma. + +*} + +lemma + fixes \1 \2::"(name\ty) list" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +proof (induct arbitrary: \2) + case (t_Var \1 x T) + have a1: "valid \1" by fact + have a2: "(x,T) \ set \1" by fact + have a3: "valid \2" by fact + have a4: "\1 \ \2" by fact + + show "\2 \ Var x : T" sorry +next + case (t_Lam x \1 T1 t T2) + have ih: "\\3. \valid \3; (x,T1)#\1 \ \3\ \ \3 \ t : T2" by fact + have a0: "x\\1" by fact + have a1: "valid \2" by fact + have a2: "\1 \ \2" by fact + + show "\2 \ Lam [x].t : T1 \ T2" sorry +qed (auto) + + +text {* + Despite the frequent claim that the weakening lemma is trivial, + routine or obvious, the proof in the lambda-case does not go + smoothly through. Painful variable renamings seem to be necessary. + But the proof using renamings is horribly complicated. It is really + interesting whether people who claim this proof is trivial, routine + or obvious had this proof in mind. + + BTW: The following two commands help already with showing that validity + and typing are invariant under variable (permutative) renamings. +*} + +equivariance valid +equivariance typing + +lemma not_to_be_tried_at_home_weakening: + fixes \1 \2::"(name\ty) list" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +proof (induct arbitrary: \2) + case (t_Lam x \1 T1 t T2) (* lambda case *) + have fc0: "x\\1" by fact + have ih: "\\3. \valid \3; (x,T1)#\1 \ \3\ \ \3 \ t : T2" by fact + obtain c::"name" where fc1: "c\(x,t,\1,\2)" (* we obtain a fresh name *) + by (rule exists_fresh) (auto simp add: fs_name1) + have "Lam [c].([(c,x)]\t) = Lam [x].t" using fc1 (* we then alpha-rename the lambda-abstraction *) + by (auto simp add: lam.inject alpha fresh_prod fresh_atm) + moreover + have "\2 \ Lam [c].([(c,x)]\t) : T1 \ T2" (* we can then alpha-rename our original goal *) + proof - + (* we establish (x,T1)#\1 \ (x,T1)#([(c,x)]\\2) and valid ((x,T1)#([(c,x)]\\2)) *) + have "(x,T1)#\1 \ (x,T1)#([(c,x)]\\2)" + proof - + have "\1 \ \2" by fact + then have "[(c,x)]\((x,T1)#\1 \ (x,T1)#([(c,x)]\\2))" using fc0 fc1 + by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh) + then show "(x,T1)#\1 \ (x,T1)#([(c,x)]\\2)" by (rule perm_boolE) + qed + moreover + have "valid ((x,T1)#([(c,x)]\\2))" + proof - + have "valid \2" by fact + then show "valid ((x,T1)#([(c,x)]\\2))" using fc1 + by (auto intro!: v2 simp add: fresh_left calc_atm eqvts) + qed + (* these two facts give us by induction hypothesis the following *) + ultimately have "(x,T1)#([(c,x)]\\2) \ t : T2" using ih by simp + (* we now apply renamings to get to our goal *) + then have "[(c,x)]\((x,T1)#([(c,x)]\\2) \ t : T2)" by (rule perm_boolI) + then have "(c,T1)#\2 \ ([(c,x)]\t) : T2" using fc1 + by (perm_simp add: eqvts calc_atm perm_fresh_fresh ty_fresh) + then show "\2 \ Lam [c].([(c,x)]\t) : T1 \ T2" using fc1 by auto + qed + ultimately show "\2 \ Lam [x].t : T1 \ T2" by simp +qed (auto) (* var and app cases *) + + +text {* + The remedy to the complicated proof of the weakening proof + shown above is to use a stronger induction principle that + has the usual variable convention already build in. The + following command derives this induction principle for us. + (We shall explain what happens here later.) + +*} + +nominal_inductive typing + by (simp_all add: abs_fresh ty_fresh) + +text {* Compare the two induction principles *} +thm typing.induct[no_vars] +thm typing.strong_induct[no_vars] + +text {* + We can use the stronger induction principle by replacing + the line + + proof (induct arbitrary: \2) + + with + + proof (nominal_induct avoiding: \2 rule: typing.strong_induct) + + Try now the proof. + +*} + + +lemma + fixes \1 \2::"(name\ty) list" + and t ::"lam" + and \ ::"ty" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +proof (nominal_induct avoiding: \2 rule: typing.strong_induct) + case (t_Var \1 x T) (* variable case *) + have "\1 \ \2" by fact + moreover + have "valid \2" by fact + moreover + have "(x,T)\ set \1" by fact + ultimately show "\2 \ Var x : T" by auto +next + case (t_Lam x \1 T1 t T2) + have vc: "x\\2" by fact (* additional fact *) + have ih: "\\3. \valid \3; (x,T1)#\1 \ \3\ \ \3 \ t : T2" by fact + have a0: "x\\1" by fact + have a1: "valid \2" by fact + have a2: "\1 \ \2" by fact + + show "\2 \ Lam [x].t : T1 \ T2" sorry +qed (auto) (* app case *) + +text {* + Since we can use the stronger induction principle, the + proof of the weakening lemma can actually be found + automatically by Isabelle. Maybe the weakening lemma + is actually trivial (in Nominal Isabelle ;o). +*} + +lemma weakening: + fixes \1 \2::"ty_ctx" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +by (nominal_induct avoiding: \2 rule: typing.strong_induct) + (auto) + + +text {***************************************************************** + + Function Definitions: Filling a Lambda-Term into a Context + ---------------------------------------------------------- + + Many functions over datatypes can be defined by recursion on the + structure. For this purpose, Isabelle provides "fun". To use it one needs + to give a name for the function, its type, optionally some pretty-syntax + and then some equations defining the function. Like in "inductive", + "fun" expects that more than one such equation is separated by "|". + +*} + +fun + filling :: "ctx \ lam \ lam" ("_\_\") +where + "\\t\ = t" +| "(CAppL E t')\t\ = App (E\t\) t'" +| "(CAppR t' E)\t\ = App t' (E\t\)" + +text {* + After this definition Isabelle will be able to simplify + statements like: *} + +lemma + shows "(CAppL \ (Var x))\Var y\ = App (Var y) (Var x)" + by simp + + +fun + ctx_compose :: "ctx \ ctx \ ctx" ("_ \ _" [101,100] 100) +where + "\ \ E' = E'" +| "(CAppL E t') \ E' = CAppL (E \ E') t'" +| "(CAppR t' E) \ E' = CAppR t' (E \ E')" + +fun + ctx_composes :: "ctxs \ ctx" ("_\" [110] 110) +where + "[]\ = \" + | "(E#Es)\ = (Es\) \ E" + +text {* + Notice that we not just have given a pretty syntax for the functions, but + also some precedences..The numbers inside the [\] stand for the precedences + of the arguments; the one next to it the precedence of the whole term. + + This means we have to write (Es1 \ Es2) \ Es3 otherwise Es1 \ Es2 \ Es3 is + interpreted as Es1 \ (Es2 \ Es3). +*} + +text {****************************************************************** + + Structural Inductions over Contexts + ------------------------------------ + + So far we have had a look at an induction over an inductive predicate. + Another important induction principle is structural inductions for + datatypes. To illustrate structural inductions we prove some facts + about context composition, some of which we will need later on. + +*} + +text {****************************************************************** + + 5.) EXERCISE + ------------ + + Complete the proof and remove the sorries. + +*} + +lemma ctx_compose: + shows "(E1 \ E2)\t\ = E1\E2\t\\" +proof (induct E1) + case Hole + show "\ \ E2\t\ = \\E2\t\\" sorry +next + case (CAppL E1 t') + have ih: "(E1 \ E2)\t\ = E1\E2\t\\" by fact + show "((CAppL E1 t') \ E2)\t\ = (CAppL E1 t')\E2\t\\" sorry +next + case (CAppR t' E1) + have ih: "(E1 \ E2)\t\ = E1\E2\t\\" by fact + show "((CAppR t' E1) \ E2)\t\ = (CAppR t' E1)\E2\t\\" sorry +qed + + +text {****************************************************************** + + 6.) EXERCISE + ------------ + + Prove associativity of \ using the lemmas neut_hole and circ_assoc. + +*} + +lemma neut_hole: + shows "E \ \ = E" +by (induct E) (simp_all) + +lemma circ_assoc: + fixes E1 E2 E3::"ctx" + shows "(E1 \ E2) \ E3 = E1 \ (E2 \ E3)" +by (induct E1) (simp_all) + +lemma + shows "(Es1@Es2)\ = (Es2\) \ (Es1\)" +proof (induct Es1) + case Nil + show "([]@Es2)\ = Es2\ \ []\" sorry +next + case (Cons E Es1) + have ih: "(Es1@Es2)\ = Es2\ \ Es1\" by fact + + show "((E#Es1)@Es2)\ = Es2\ \ (E#Es1)\" sorry +qed + + +text {* + The last proof involves several steps of equational reasoning. + Isar provides some convenient means to express such equational + reasoning in a much cleaner fashion using the "also have" + construction. *} + +lemma + shows "(Es1@Es2)\ = (Es2\) \ (Es1\)" +proof (induct Es1) + case Nil + show "([]@Es2)\ = Es2\ \ []\" using neut_hole by simp +next + case (Cons E Es1) + have ih: "(Es1@Es2)\ = Es2\ \ Es1\" by fact + have "((E#Es1)@Es2)\ = (Es1@Es2)\ \ E" by simp + also have "\ = (Es2\ \ Es1\) \ E" using ih by simp + also have "\ = Es2\ \ (Es1\ \ E)" using circ_assoc by simp + also have "\ = Es2\ \ (E#Es1)\" by simp + finally show "((E#Es1)@Es2)\ = Es2\ \ (E#Es1)\" by simp +qed + + +text {****************************************************************** + + Formalising Barendregt's Proof of the Substitution Lemma + -------------------------------------------------------- + + Barendregt's proof needs in the variable case a case distinction. + One way to do this in Isar is to use blocks. A block is some sequent + or reasoning steps enclosed in curly braces + + { \ + + have "statement" + } + + Such a block can contain local assumptions like + + { assume "A" + assume "B" + \ + have "C" by \ + } + + Where "C" is the last have-statement in this block. The behaviour + of such a block to the 'outside' is the implication + + \A; B\ \ "C" + + Now if we want to prove a property "smth" using the case-distinctions + P\<^isub>1, P\<^isub>2 and P\<^isub>3 then we can use the following reasoning: + + { assume "P\<^isub>1" + \ + have "smth" + } + moreover + { assume "P\<^isub>2" + \ + have "smth" + } + moreover + { assume "P\<^isub>3" + \ + have "smth" + } + ultimately have "smth" by blast + + The blocks establish the implications + + P\<^isub>1 \ smth + P\<^isub>2 \ smth + P\<^isub>3 \ smth + + If we know that P\<^isub>1, P\<^isub>2 and P\<^isub>3 cover all the cases, that is P\<^isub>1 \ P\<^isub>2 \ P\<^isub>3 is + true, then we have 'ultimately' established the property "smth" + +*} + +text {****************************************************************** + + 7.) Exercise + ------------ + + Fill in the cases 1.2 and 1.3 and the equational reasoning + in the lambda-case. +*} + +thm forget[no_vars] +thm fresh_fact[no_vars] + +lemma + assumes a: "x\y" + and b: "x\L" + shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" +using a b +proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct) + case (Var z) + have a1: "x\y" by fact + have a2: "x\L" by fact + show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS") + proof - + { (*Case 1.1*) + assume c1: "z=x" + have "(1)": "?LHS = N[y::=L]" using c1 by simp + have "(2)": "?RHS = N[y::=L]" using c1 a1 by simp + have "?LHS = ?RHS" using "(1)" "(2)" by simp + } + moreover + { (*Case 1.2*) + assume c2: "z=y" "z\x" + + have "?LHS = ?RHS" sorry + } + moreover + { (*Case 1.3*) + assume c3: "z\x" "z\y" + + have "?LHS = ?RHS" sorry + } + ultimately show "?LHS = ?RHS" by blast + qed +next + case (Lam z M1) (* case 2: lambdas *) + have ih: "\x\y; x\L\ \ M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact + have a1: "x\y" by fact + have a2: "x\L" by fact + have fs: "z\x" "z\y" "z\N" "z\L" by fact+ + then have b: "z\N[y::=L]" by (simp add: fresh_fact) + show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") + proof - + have "?LHS = \" sorry + + also have "\ = ?RHS" sorry + finally show "?LHS = ?RHS" by simp + qed +next + case (App M1 M2) (* case 3: applications *) + then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp +qed + +text {* + Again the strong induction principle enables Isabelle to find + the proof of the substitution lemma automatically. +*} + +lemma substitution_lemma_version: + assumes asm: "x\y" "x\L" + shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" + using asm +by (nominal_induct M avoiding: x y N L rule: lam.strong_induct) + (auto simp add: fresh_fact forget) + +text {****************************************************************** + + The CBV Reduction Relation (Small-Step Semantics) + ------------------------------------------------- + + In order to establish the property that the CK Machine + calculates a nomrmalform which corresponds to the + evaluation relation, we introduce the call-by-value + small-step semantics. + +*} + +inductive + cbv :: "lam\lam\bool" ("_ \cbv _") +where + cbv1: "\val v; x\v\ \ App (Lam [x].t) v \cbv t[x::=v]" +| cbv2[intro]: "t \cbv t' \ App t t2 \cbv App t' t2" +| cbv3[intro]: "t \cbv t' \ App t2 t \cbv App t2 t'" + +equivariance val +equivariance cbv +nominal_inductive cbv + by (simp_all add: abs_fresh fresh_fact) + +text {* + In order to satisfy the vc-condition we have to formulate + this relation with the additional freshness constraint + x\v. Though this makes the definition vc-ompatible, it + makes the definition less useful. We can with some pain + show that the more restricted rule is equivalent to the + usual rule. *} + +thm subst_rename + +lemma better_cbv1[intro]: + assumes a: "val v" + shows "App (Lam [x].t) v \cbv t[x::=v]" +proof - + obtain y::"name" where fs: "y\(x,t,v)" by (rule exists_fresh) (auto simp add: fs_name1) + have "App (Lam [x].t) v = App (Lam [y].([(y,x)]\t)) v" using fs + by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) + also have "\ \cbv ([(y,x)]\t)[y::=v]" using fs a by (auto intro: cbv1) + also have "\ = t[x::=v]" using fs by (simp add: subst_rename[symmetric]) + finally show "App (Lam [x].t) v \cbv t[x::=v]" by simp +qed + +text {* + The transitive closure of the cbv-reduction relation: *} + +inductive + "cbvs" :: "lam\lam\bool" (" _ \cbv* _") +where + cbvs1[intro]: "e \cbv* e" +| cbvs2[intro]: "\e1\cbv e2; e2 \cbv* e3\ \ e1 \cbv* e3" + +lemma cbvs3[intro]: + assumes a: "e1 \cbv* e2" "e2 \cbv* e3" + shows "e1 \cbv* e3" +using a by (induct) (auto) + +text {****************************************************************** + + 8.) Exercise + ------------ + + If more simple exercises are needed, then complete the following proof. + +*} + +lemma cbv_in_ctx: + assumes a: "t \cbv t'" + shows "E\t\ \cbv E\t'\" +using a +proof (induct E) + case Hole + have "t \cbv t'" by fact + then show "\\t\ \cbv \\t'\" sorry +next + case (CAppL E s) + have ih: "t \cbv t' \ E\t\ \cbv E\t'\" by fact + have a: "t \cbv t'" by fact + show "(CAppL E s)\t\ \cbv (CAppL E s)\t'\" sorry +next + case (CAppR s E) + have ih: "t \cbv t' \ E\t\ \cbv E\t'\" by fact + have a: "t \cbv t'" by fact + show "(CAppR s E)\t\ \cbv (CAppR s E)\t'\" sorry +qed + + +text {****************************************************************** + + 9.) Exercise + ------------ + + The point of the cbv-reduction was that we can easily relatively + establish the follwoing property: + +*} + +lemma machine_implies_cbvs_ctx: + assumes a: " \ " + shows "(Es\)\e\ \cbv* (Es'\)\e'\" +using a +proof (induct) + case (m1 t1 t2 Es) + + show "Es\\App t1 t2\ \cbv* ((CAppL \ t2)#Es)\\t1\" sorry +next + case (m2 v t2 Es) + have "val v" by fact + + show "((CAppL \ t2)#Es)\\v\ \cbv* (CAppR v \ # Es)\\t2\" sorry +next + case (m3 v x t Es) + have "val v" by fact + + show "(((CAppR (Lam [x].t) \)#Es)\)\v\ \cbv* (Es\)\t[x::=v]\" sorry +qed + +text {* + It is not difficult to extend the lemma above to + arbitrary reductions sequences of the CK machine. *} + +lemma machines_implies_cbvs_ctx: + assumes a: " \* " + shows "(Es\)\e\ \cbv* (Es'\)\e'\" +using a +by (induct) (auto dest: machine_implies_cbvs_ctx) + +text {* + So whenever we let the CL machine start in an initial + state and it arrives at a final state, then there exists + a corresponding cbv-reduction sequence. *} + +corollary machines_implies_cbvs: + assumes a: " \* " + shows "e \cbv* e'" +using a by (auto dest: machines_implies_cbvs_ctx) + +text {* + We now want to relate the cbv-reduction to the evaluation + relation. For this we need two auxiliary lemmas. *} + +lemma eval_val: + assumes a: "val t" + shows "t \ t" +using a by (induct) (auto) + +lemma e_App_elim: + assumes a: "App t1 t2 \ v" + shows "\x t v'. t1 \ Lam [x].t \ t2 \ v' \ t[x::=v'] \ v" +using a by (cases) (auto simp add: lam.inject) + +text {****************************************************************** + + 10.) Exercise + ------------- + + Complete the first case in the proof below. + +*} + +lemma cbv_eval: + assumes a: "t1 \cbv t2" "t2 \ t3" + shows "t1 \ t3" +using a +proof(induct arbitrary: t3) + case (cbv1 v x t t3) + have a1: "val v" by fact + have a2: "t[x::=v] \ t3" by fact + + show "App Lam [x].t v \ t3" sorry +next + case (cbv2 t t' t2 t3) + have ih: "\t3. t' \ t3 \ t \ t3" by fact + have "App t' t2 \ t3" by fact + then obtain x t'' v' + where a1: "t' \ Lam [x].t''" + and a2: "t2 \ v'" + and a3: "t''[x::=v'] \ t3" using e_App_elim by blast + have "t \ Lam [x].t''" using ih a1 by auto + then show "App t t2 \ t3" using a2 a3 by auto +qed (auto dest!: e_App_elim) + + +text {* + Next we extend the lemma above to arbitray initial + sequences of cbv-reductions. *} + +lemma cbvs_eval: + assumes a: "t1 \cbv* t2" "t2 \ t3" + shows "t1 \ t3" +using a by (induct) (auto intro: cbv_eval) + +text {* + Finally, we can show that if from a term t we reach a value + by a cbv-reduction sequence, then t evaluates to this value. *} + +lemma cbvs_implies_eval: + assumes a: "t \cbv* v" "val v" + shows "t \ v" +using a +by (induct) (auto intro: eval_val cbvs_eval) + +text {* + All facts tied together give us the desired property about + K machines. *} + +theorem machines_implies_eval: + assumes a: " \* " + and b: "val t2" + shows "t1 \ t2" +proof - + have "t1 \cbv* t2" using a by (simp add: machines_implies_cbvs) + then show "t1 \ t2" using b by (simp add: cbvs_implies_eval) +qed + +text {****************************************************************** + + Formalising a Type-Soundness and Progress Lemma for CBV + ------------------------------------------------------- + + The central lemma for type-soundness is type-substitutity. In + our setting type-substitutivity is slightly painful to establish. + +*} + +lemma valid_elim: + assumes a: "valid ((x,T)#\)" + shows "x\\ \ valid \" +using a by (cases) (auto) + +lemma valid_insert: + assumes a: "valid (\@[(x,T)]@\)" + shows "valid (\@\)" +using a +by (induct \) + (auto simp add: fresh_list_append fresh_list_cons dest!: valid_elim) + +lemma fresh_list: + shows "y\xs = (\x\set xs. y\x)" +by (induct xs) (simp_all add: fresh_list_nil fresh_list_cons) + +lemma context_unique: + assumes a1: "valid \" + and a2: "(x,T) \ set \" + and a3: "(x,U) \ set \" + shows "T = U" +using a1 a2 a3 +by (induct) (auto simp add: fresh_list fresh_prod fresh_atm) + +lemma type_substitution_aux: + assumes a: "(\@[(x,T')]@\) \ e : T" + and b: "\ \ e' : T'" + shows "(\@\) \ e[x::=e'] : T" +using a b +proof (nominal_induct \'\"\@[(x,T')]@\" e T avoiding: x e' \ rule: typing.strong_induct) + case (t_Var \' y T x e' \) + then have a1: "valid (\@[(x,T')]@\)" + and a2: "(y,T) \ set (\@[(x,T')]@\)" + and a3: "\ \ e' : T'" by simp_all + from a1 have a4: "valid (\@\)" by (rule valid_insert) + { assume eq: "x=y" + from a1 a2 have "T=T'" using eq by (auto intro: context_unique) + with a3 have "\@\ \ Var y[x::=e'] : T" using eq a4 by (auto intro: weakening) + } + moreover + { assume ineq: "x\y" + from a2 have "(y,T) \ set (\@\)" using ineq by simp + then have "\@\ \ Var y[x::=e'] : T" using ineq a4 by auto + } + ultimately show "\@\ \ Var y[x::=e'] : T" by blast +qed (force simp add: fresh_list_append fresh_list_cons)+ + +corollary type_substitution: + assumes a: "(x,T')#\ \ e : T" + and b: "\ \ e' : T'" + shows "\ \ e[x::=e'] : T" +using a b type_substitution_aux[where \="[]"] +by (auto) + +lemma t_App_elim: + assumes a: "\ \ App t1 t2 : T" + shows "\T'. \ \ t1 : T' \ T \ \ \ t2 : T'" +using a +by (cases) (auto simp add: lam.inject) + +lemma t_Lam_elim: + assumes ty: "\ \ Lam [x].t : T" + and fc: "x\\" + shows "\T1 T2. T = T1 \ T2 \ (x,T1)#\ \ t : T2" +using ty fc +by (cases rule: typing.strong_cases) + (auto simp add: alpha lam.inject abs_fresh ty_fresh) + +theorem cbv_type_preservation: + assumes a: "t \cbv t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +using a b +by (nominal_induct avoiding: \ T rule: cbv.strong_induct) + (auto dest!: t_Lam_elim t_App_elim simp add: type_substitution ty.inject) + +corollary cbvs_type_preservation: + assumes a: "t \cbv* t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +using a b +by (induct) (auto intro: cbv_type_preservation) + +text {* + The Type-Preservation Property for the Machine and Evaluation Relation. *} + +theorem machine_type_preservation: + assumes a: " \* " + and b: "\ \ t : T" + shows "\ \ t' : T" +proof - + from a have "t \cbv* t'" by (simp add: machines_implies_cbvs) + then show "\ \ t' : T" using b by (simp add: cbvs_type_preservation) +qed + +theorem eval_type_preservation: + assumes a: "t \ t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +proof - + from a have " \* " by (simp add: eval_implies_machines) + then show "\ \ t' : T" using b by (simp add: machine_type_preservation) +qed + +text {* The Progress Property *} + +lemma canonical_tArr: + assumes a: "[] \ t : T1 \ T2" + and b: "val t" + shows "\x t'. t = Lam [x].t'" +using b a by (induct) (auto) + +theorem progress: + assumes a: "[] \ t : T" + shows "(\t'. t \cbv t') \ (val t)" +using a +by (induct \\"[]::ty_ctx" t T) + (auto intro: cbv.intros dest!: canonical_tArr) + + +text {*********************************************************** + + Strong Induction Principle + -------------------------- + + A proof for the strong (structural) induction principle in the + lambda-calculus. + +*} + + +lemma lam_strong_induct: + fixes c::"'a::fs_name" + assumes h1: "\x c. P c (Var x)" + and h2: "\t1 t2 c. \\d. P d t1; \d. P d t2\ \ P c (App t1 t2)" + and h3: "\x t c. \x\c; \d. P d t\ \ P c (Lam [x].t)" + shows "P c t" +proof - + have "\(\::name prm) c. P c (\\t)" + proof (induct t rule: lam.induct) + case (Lam x t) + have ih: "\(\::name prm) c. P c (\\t)" by fact + { fix \::"name prm" and c::"'a::fs_name" + obtain y::"name" where fc: "y\(\\x,\\t,c)" + by (rule exists_fresh) (auto simp add: fs_name1) + from ih have "\c. P c (([(y,\\x)]@\)\t)" by simp + then have "\c. P c ([(y,\\x)]\(\\t))" by (auto simp only: pt_name2) + with h3 have "P c (Lam [y].[(y,\\x)]\(\\t))" using fc by (simp add: fresh_prod) + moreover + have "Lam [y].[(y,\\x)]\(\\t) = Lam [(\\x)].(\\t)" + using fc by (simp add: lam.inject alpha fresh_atm fresh_prod) + ultimately have "P c (Lam [(\\x)].(\\t))" by simp + } + then have "\(\::name prm) c. P c (Lam [(\\x)].(\\t))" by simp + then show "\(\::name prm) c. P c (\\(Lam [x].t))" by simp + qed (auto intro: h1 h2) (* var and app case *) + then have "P c (([]::name prm)\t)" by blast + then show "P c t" by simp +qed + +text {*********************************************************** + + --------- + SOLUTIONS + --------- + +*} + +text {************************************************************ + + 1.) MINI EXERCISE + + The way we defined contexts does not allow us to + apply a Hole to a Hole. Therefore the following + will result in a typing error. *} + +(* term "CAppL \ \" *) + + +text {************************************************************ + + 2. EXERCISE + + A readable proof for this lemma is as follows: + +*} + +lemma + assumes a: " \* " + and b: " \* " + shows " \* " +using a b +proof(induct) + case (ms1 e1 Es1) + show " \* " by fact +next + case (ms2 e1 Es1 e2 Es2 e2' Es2') + have ih: " \* \ \* " by fact + have " \* " by fact + then have " \* " using ih by simp + moreover + have " \ " by fact + ultimately show " \* " by auto +qed + + +text {************************************************************ + + 3.) Exercise + + As one can quickly see in the second case, the theorem as stated + does not go through. We need to generalise the induction hypothesis + so that we show the lemma for all contexts Es. In Isar, variables + can be generalised by declaring "arbitrary: variable \" when the + induction is set up. + +*} + +theorem + assumes a: "t \ t'" + shows " \* " +using a +proof (induct arbitrary: Es) (* here we generalise over Es *) + case (e_Lam x t Es) + show " \* " by auto +next + case (e_App t1 x t t2 v' v Es) + have ih1: "\Es. \* " by fact + have ih2: "\Es. \* " by fact + have ih3: "\Es. \* " by fact + have " \* t2#Es>" by auto + moreover + have " t2#Es> \* t2#Es>" using ih1 by auto + moreover + have " t2#Es> \* #Es>" by auto + moreover + have "#Es> \* #Es>" using ih2 by auto + moreover + have "t2 \ v'" by fact + then have "val v'" using eval_to_val by auto + then have "#Es> \* " by auto + moreover + have " \* " using ih3 by auto + ultimately show " \* " by blast +qed + +text {************************************************************ + + 4.) Exercise + + A proof for the weakening lemma: + +*} + +lemma + fixes \1 \2::"(name\ty) list" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +proof (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) + case (t_Var \1 x T) (* variable case *) + have "\1 \ \2" by fact + moreover + have "valid \2" by fact + moreover + have "(x,T)\ set \1" by fact + ultimately show "\2 \ Var x : T" by auto +next + case (t_Lam x \1 T1 t T2) (* lambda case *) + have vc: "x\\2" by fact (* variable convention *) + have ih: "\valid ((x,T1)#\2); (x,T1)#\1 \ (x,T1)#\2\ \ (x,T1)#\2 \ t:T2" by fact + have "\1 \ \2" by fact + then have "(x,T1)#\1 \ (x,T1)#\2" by simp + moreover + have "valid \2" by fact + then have "valid ((x,T1)#\2)" using vc by (simp add: v2) + ultimately have "(x,T1)#\2 \ t : T2" using ih by simp + with vc show "\2 \ Lam [x].t : T1\T2" by auto +qed (auto) (* app case *) + +text {************************************************************ + + 5.) Exercise + + A proof for context omposition + +*} + +lemma + shows "(E1 \ E2)\t\ = E1\E2\t\\" +by (induct E1) (simp_all) + +text {****************************************************************** + + 6.) EXERCISE + ------------ + + A proof for the assoiativity of \. + +*} + +lemma + shows "(Es1@Es2)\ = (Es2\) \ (Es1\)" +proof (induct Es1) + case Nil + show "([]@Es2)\ = Es2\ \ []\" using neut_hole by simp +next + case (Cons E Es1) + have ih: "(Es1@Es2)\ = Es2\ \ Es1\" by fact + have "((E#Es1)@Es2)\ = (Es1@Es2)\ \ E" by simp + also have "\ = (Es2\ \ Es1\) \ E" using ih by simp + also have "\ = Es2\ \ (Es1\ \ E)" using circ_assoc by simp + also have "\ = Es2\ \ (E#Es1)\" by simp + finally show "((E#Es1)@Es2)\ = Es2\ \ (E#Es1)\" by simp +qed + +text {****************************************************************** + + 7.) EXERCISE + ------------ + + A proof for the substitution lemma. + +*} + +lemma + assumes a: "x\y" + and b: "x\L" + shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" +using a b +proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct) + case (Var z) (* case 1: Variables*) + show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS") + proof - + { (*Case 1.1*) + assume "z=x" + have "(1)": "?LHS = N[y::=L]" using `z=x` by simp + have "(2)": "?RHS = N[y::=L]" using `z=x` `x\y` by simp + from "(1)" "(2)" have "?LHS = ?RHS" by simp + } + moreover + { (*Case 1.2*) + assume "z=y" and "z\x" + have "(1)": "?LHS = L" using `z\x` `z=y` by simp + have "(2)": "?RHS = L[x::=N[y::=L]]" using `z=y` by simp + have "(3)": "L[x::=N[y::=L]] = L" using `x\L` by (simp add: forget) + from "(1)" "(2)" "(3)" have "?LHS = ?RHS" by simp + } + moreover + { (*Case 1.3*) + assume "z\x" and "z\y" + have "(1)": "?LHS = Var z" using `z\x` `z\y` by simp + have "(2)": "?RHS = Var z" using `z\x` `z\y` by simp + from "(1)" "(2)" have "?LHS = ?RHS" by simp + } + ultimately show "?LHS = ?RHS" by blast + qed +next + case (Lam z M1) (* case 2: lambdas *) + have ih: "\x\y; x\L\ \ M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact + have fs: "z\x" "z\y" "z\N" "z\L" by fact+ + then have "z\N[y::=L]" by (simp add: fresh_fact) + show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") + proof - + have "?LHS = Lam [z].(M1[x::=N][y::=L])" using `z\x` `z\y` `z\N` `z\L` by simp + also from ih have "\ = Lam [z].(M1[y::=L][x::=N[y::=L]])" using `x\y` `x\L` by simp + also have "\ = (Lam [z].(M1[y::=L]))[x::=N[y::=L]]" using `z\x` `z\N[y::=L]` by simp + also have "\ = ?RHS" using `z\y` `z\L` by simp + finally show "?LHS = ?RHS" by simp + qed +next + case (App M1 M2) (* case 3: applications *) + then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp +qed + +text {****************************************************************** + + 8.) Exercise + ------------ + + Left out if not needed. +*} + +lemma + assumes a: "t \cbv t'" + shows "E\t\ \cbv E\t'\" +using a +proof (induct E) + case Hole + have "t \cbv t'" by fact + then show "\\t\ \cbv \\t'\" by simp +next + case (CAppL E s) + have ih: "t \cbv t' \ E\t\ \cbv E\t'\" by fact + have a: "t \cbv t'" by fact + show "(CAppL E s)\t\ \cbv (CAppL E s)\t'\" using ih a by auto +next + case (CAppR s E) + have ih: "t \cbv t' \ E\t\ \cbv E\t'\" by fact + have a: "t \cbv t'" by fact + show "(CAppR s E)\t\ \cbv (CAppR s E)\t'\" using ih a by auto +qed + +text {****************************************************************** + + 9.) Exercise + ------------ + + The point of the cbv-reduction was that we can easily relatively + establish the follwoing property: + +*} + +lemma + assumes a: " \ " + shows "(Es\)\e\ \cbv* (Es'\)\e'\" +using a +proof (induct) + case (m1 t1 t2 Es) + show "Es\\App t1 t2\ \cbv* ((CAppL \ t2)#Es)\\t1\" by (auto simp add: ctx_compose) +next + case (m2 v t2 Es) + have "val v" by fact + then show "((CAppL \ t2)#Es)\\v\ \cbv* (CAppR v \ # Es)\\t2\" + by (auto simp add: ctx_compose) +next + case (m3 v x t Es) + have "val v" by fact + then show "((CAppR (Lam [x].t) \)#Es)\\v\ \cbv* (Es\)\t[x::=v]\" + by (auto simp add: ctx_compose intro: cbv_in_ctx) +qed + +lemma + assumes a: " \ " + shows "(Es\)\e\ \cbv* (Es'\)\e'\" +using a by (induct) (auto simp add: ctx_compose intro: cbv_in_ctx) + +text {****************************************************************** + + 10.) Exercise + ------------- + + Complete the first case in the proof below. + +*} + +lemma + assumes a: "t1 \cbv t2" "t2 \ t3" + shows "t1 \ t3" +using a +proof(induct arbitrary: t3) + case (cbv1 v x t t3) + have a1: "val v" by fact + have a2: "t[x::=v] \ t3" by fact + show "App Lam [x].t v \ t3" using eval_val a1 a2 by auto +next + case (cbv2 t t' t2 t3) + have "t \cbv t'" by fact + have ih: "\t3. t' \ t3 \ t \ t3" by fact + have "App t' t2 \ t3" by fact + then obtain x t'' v' + where a1: "t' \ Lam [x].t''" + and a2: "t2 \ v'" + and a3: "t''[x::=v'] \ t3" using e_App_elim by blast + have "t \ Lam [x].t''" using ih a1 by auto + then show "App t t2 \ t3" using a2 a3 by auto +qed (auto dest!: e_App_elim) + +lemma + assumes a: "t1 \cbv t2" "t2 \ t3" + shows "t1 \ t3" +using a +by (induct arbitrary: t3) + (auto elim!: eval_elim intro: eval_val) + + +end + + + + + + + + + +