diff -r e0391947b7da -r 7b2691911fbc Tutorial/Tutorial5.thy --- a/Tutorial/Tutorial5.thy Sat Jan 22 16:37:00 2011 -0600 +++ b/Tutorial/Tutorial5.thy Sat Jan 22 18:59:48 2011 -0600 @@ -1,9 +1,22 @@ + + theory Tutorial5 imports Tutorial4 begin +section {* Type-Preservation and Progress Lemma*} -section {* Type Preservation (fixme separate file) *} +text {* + The point of this tutorial is to prove the + type-preservation and progress lemma. Since + we now know that \, \cbv* and the machine + correspond to each other, we only need to + prove this property for one of them. We chose + \cbv. + + First we need to establish two elimination + properties and two auxiliary lemmas about contexts. +*} lemma valid_elim: @@ -30,6 +43,16 @@ using a1 a2 a3 by (induct) (auto simp add: fresh_list fresh_Pair fresh_at_base) + +section {* EXERCISE 16 *} + +text {* + Next we want to show the type substitution lemma. Unfortunately, + we have to prove a slightly more general version of it, where + the variable being substituted occurs somewhere inside the + context. +*} + lemma type_substitution_aux: assumes a: "\ @ [(x, T')] @ \ \ e : T" and b: "\ \ e' : T'" @@ -40,10 +63,11 @@ have a1: "valid (\ @ [(x, T')] @ \)" by fact have a2: "(y,T) \ set (\ @ [(x, T')] @ \)" by fact have a3: "\ \ e' : T'" by fact + 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) + + have "\ @ \ \ Var y[x ::= e'] : T" sorry } moreover { assume ineq: "x \ y" @@ -51,15 +75,46 @@ 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_append fresh_Cons)+ +next + case (t_Lam y T1 t T2 x e' \) + have a1: "atom y \ e'" by fact + have a2: "atom y \ \ @ [(x, T')] @ \" by fact + have a3: "\ \ e' : T'" by fact + have ih: "\ \ e' : T' \ ((y, T1) # \) @ \ \ t [x ::= e'] : T2" + using t_Lam(6)[of "(y, T1) # \"] by auto + + + show "\ @ \ \ (Lam [y]. t)[x ::= e'] : T1 \ T2" sorry +next + case (t_App t1 T1 T2 t2 x e' \) + have ih1: "\ \ e' : T' \ \ @ \ \ t1 [x ::= e'] : T1 \ T2" using t_App(2) by auto + have ih2: "\ \ e' : T' \ \ @ \ \ t2 [x ::= e'] : T1" using t_App(4) by auto + have a: "\ \ e' : T'" by fact + + show "\ @ \ \ App t1 t2 [x ::= e'] : T2" sorry +qed + +text {* + From this we can derive the usual version of the substitution + lemma. +*} 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 \="[]"] +using a b type_substitution_aux[of "[]"] by auto + +section {* Type Preservation *} + +text {* + Finally we are in a position to establish the type preservation + property. We just need the following two inversion rules for + particualr typing instances. +*} + lemma t_App_elim: assumes a: "\ \ App t1 t2 : T" obtains T' where "\ \ t1 : T' \ T" "\ \ t2 : T'" @@ -81,13 +136,34 @@ apply(auto simp add: flip_def swap_fresh_fresh ty_fresh) done + +section {* EXERCISE 17 *} + +text {* + Fill in the gaps in the t_Lam case. You will need + the type substitution lemma proved above. +*} + 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 elim!: t_Lam_elim t_App_elim simp add: type_substitution ty.eq_iff) +proof (nominal_induct avoiding: \ T rule: cbv.strong_induct) + case (cbv1 v x t \ T) + have fc: "atom x \ \" by fact + have "\ \ App (Lam [x]. t) v : T" by fact + then obtain T' where + *: "\ \ Lam [x]. t : T' \ T" and + **: "\ \ v : T'" by (rule t_App_elim) + have "(x, T') # \ \ t : T" using * fc by (rule t_Lam_elim) (simp add: ty.eq_iff) + + show "\ \ t [x ::= v] : T " sorry +qed (auto elim!: t_App_elim) + +text {* + We can easily extend this to sequences of cbv* reductions. +*} corollary cbvs_type_preservation: assumes a: "t \cbv* t'"