diff -r f325eefe803e -r abb6c3ac2df2 Tutorial/Tutorial5.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Tutorial/Tutorial5.thy Fri Jan 21 22:23:44 2011 +0100 @@ -0,0 +1,142 @@ +theory Tutorial5 +imports Tutorial4 +begin + + +section {* Type Preservation (fixme separate file) *} + + +lemma valid_elim: + assumes a: "valid ((x, T) # \)" + shows "atom 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_append fresh_Cons dest!: valid_elim) + +lemma fresh_list: + shows "atom y \ xs = (\x \ set xs. atom y \ x)" +by (induct xs) (simp_all add: fresh_Nil fresh_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_Pair fresh_at_base) + +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' \) + 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) + } + 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_append fresh_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" + obtains T' where "\ \ t1 : T' \ T" "\ \ t2 : T'" +using a +by (cases) (auto simp add: lam.eq_iff lam.distinct) + +text {* we have not yet generated strong elimination rules *} +lemma t_Lam_elim: + assumes ty: "\ \ Lam [x].t : T" + and fc: "atom x \ \" + obtains T1 T2 where "T = T1 \ T2" "(x, T1) # \ \ t : T2" +using ty fc +apply(cases) +apply(auto simp add: lam.eq_iff lam.distinct ty.eq_iff) +apply(auto simp add: Abs1_eq_iff) +apply(rotate_tac 3) +apply(drule_tac p="(x \ xa)" in permute_boolI) +apply(perm_simp) +apply(auto simp add: flip_def swap_fresh_fresh ty_fresh) +done + +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) + +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 - + have "t \cbv* t'" using a machines_implies_cbvs by simp + then show "\ \ t' : T" using b cbvs_type_preservation by simp +qed + +theorem eval_type_preservation: + assumes a: "t \ t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +proof - + have " \* " using a eval_implies_machines by simp + then show "\ \ t' : T" using b machine_type_preservation by simp +qed + +text {* The Progress Property *} + +lemma canonical_tArr: + assumes a: "[] \ t : T1 \ T2" + and b: "val t" + obtains x t' where "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 elim: canonical_tArr) + +text {* + Done! Congratulations! +*} + +end +