# HG changeset patch # User urbanc # Date 1212334761 -7200 # Node ID 6fd85edc403d325e775b10205551d2e57e8390be # Parent 504b6f597233620ea1a9ffc80135e049063070d4 new example diff -r 504b6f597233 -r 6fd85edc403d src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sat May 31 00:34:04 2008 +0200 +++ b/src/HOL/IsaMakefile Sun Jun 01 17:39:21 2008 +0200 @@ -826,6 +826,7 @@ Nominal/Examples/SN.thy \ Nominal/Examples/SOS.thy \ Nominal/Examples/Support.thy \ + Nominal/Examples/Type_Preservation.thy \ Nominal/Examples/VC_Condition.thy \ Nominal/Examples/W.thy \ Nominal/Examples/Weakening.thy diff -r 504b6f597233 -r 6fd85edc403d src/HOL/Nominal/Examples/Type_Preservation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nominal/Examples/Type_Preservation.thy Sun Jun 01 17:39:21 2008 +0200 @@ -0,0 +1,219 @@ +theory Type_Preservation + imports Nominal +begin + +atom_decl name + +nominal_datatype lam = + Var "name" +| App "lam" "lam" +| Lam "\name\lam" ("Lam [_]._") + +text {* Capture-Avoiding Substitution *} + +consts + subst :: "lam \ name \ lam \ lam" ("_[_::=_]") + +nominal_primrec + "(Var x)[y::=s] = (if x=y then s else (Var x))" + "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" + "x\(y,s) \ (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" +apply(finite_guess)+ +apply(rule TrueI)+ +apply(simp add: abs_fresh) +apply(fresh_guess)+ +done + +lemma subst_eqvt[eqvt]: + fixes pi::"name prm" + shows "pi\(t1[x::=t2]) = (pi\t1)[(pi\x)::=(pi\t2)]" +by (nominal_induct t1 avoiding: x t2 rule: lam.strong_induct) + (auto simp add: perm_bij fresh_atm fresh_bij) + +lemma fresh_fact: + fixes z::"name" + shows "\z\s; (z=y \ z\t)\ \ z\t[y::=s]" +by (nominal_induct t avoiding: z y s rule: lam.strong_induct) + (auto simp add: abs_fresh fresh_prod fresh_atm) + +text {* Types *} + +nominal_datatype ty = + TVar "string" + | TArr "ty" "ty" ("_ \ _") + +lemma ty_fresh: + fixes x::"name" + and T::"ty" + shows "x\T" +by (induct T rule: ty.induct) + (auto simp add: fresh_string) + +text {* Typing Contexts *} + +types ctx = "(name\ty) list" + +abbreviation + "sub_ctx" :: "ctx \ ctx \ bool" ("_ \ _") +where + "\\<^isub>1 \ \\<^isub>2 \ \x. x \ set \\<^isub>1 \ x \ set \\<^isub>2" + +text {* Validity of typing contexts *} + +inductive + valid :: "(name\ty) list \ bool" +where + v1[intro]: "valid []" + | v2[intro]: "\valid \; x\\\\ valid ((x,T)#\)" + +equivariance valid + +lemma valid_elim[dest]: + 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_set: + 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_set fresh_prod fresh_atm) + +text {* typing relation *} + +inductive + typing :: "ctx \ lam \ ty \ bool" ("_ \ _ : _") +where + t_Var[intro]: "\valid \; (x,T)\set \\ \ \ \ Var x : T" +| t_App[intro]: "\\ \ t\<^isub>1 : T\<^isub>1\T\<^isub>2; \ \ t\<^isub>2 : T\<^isub>1\ \ \ \ App t\<^isub>1 t\<^isub>2 : T\<^isub>2" +| t_Lam[intro]: "\x\\; (x,T\<^isub>1)#\ \ t : T\<^isub>2\ \ \ \ Lam [x].t : T\<^isub>1\T\<^isub>2" + +equivariance typing +nominal_inductive typing + by (simp_all add: abs_fresh ty_fresh) + +lemma t_Lam_inversion[dest]: + assumes ty: "\ \ Lam [x].t : T" + and fc: "x\\" + shows "\T\<^isub>1 T\<^isub>2. T = T\<^isub>1 \ T\<^isub>2 \ (x,T\<^isub>1)#\ \ t : T\<^isub>2" +using ty fc +by (cases rule: typing.strong_cases) + (auto simp add: alpha lam.inject abs_fresh ty_fresh) + +lemma t_App_inversion[dest]: + assumes ty: "\ \ App t1 t2 : T" + shows "\T'. \ \ t1 : T' \ T \ \ \ t2 : T'" +using ty +by (cases) (auto simp add: lam.inject) + +lemma weakening: + fixes \1 \2::"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 \1 t T avoiding: \2 rule: typing.strong_induct) + (auto | atomize)+ + +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 +by (auto intro: type_substitution_aux[where \="[]",simplified]) + +text {* Beta reduction *} + +inductive + "beta" :: "lam\lam\bool" (" _ \\<^isub>\ _") +where + b1[intro]: "t1 \\<^isub>\ t2 \ App t1 s \\<^isub>\ App t2 s" +| b2[intro]: "s1 \\<^isub>\ s2 \ App t s1 \\<^isub>\ App t s2" +| b3[intro]: "t1 \\<^isub>\ t2 \ Lam [x].t1 \\<^isub>\ Lam [x].t2" +| b4[intro]: "x\s \ App (Lam [x].t) s \\<^isub>\ t[x::=s]" + +equivariance beta +nominal_inductive beta + by (auto simp add: abs_fresh fresh_fact) + + +theorem type_preservation: + assumes a: "t \\<^isub>\ t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +using a b +proof (nominal_induct avoiding: \ T rule: beta.strong_induct) + case (b1 t1 t2 s \ T) + have "\ \ App t1 s : T" by fact + then obtain T' where a1: "\ \ t1 : T' \ T" and a2: "\ \ s : T'" by auto + have ih: "\ \ t1 : T' \ T \ \ \ t2 : T' \ T" by fact + with a1 a2 show "\ \ App t2 s : T" by auto +next + case (b2 s1 s2 t \ T) + have "\ \ App t s1 : T" by fact + then obtain T' where a1: "\ \ t : T' \ T" and a2: "\ \ s1 : T'" by auto + have ih: "\ \ s1 : T' \ \ \ s2 : T'" by fact + with a1 a2 show "\ \ App t s2 : T" by auto +next + case (b3 t1 t2 x \ T) + have vc: "x\\" by fact + have "\ \ Lam [x].t1 : T" by fact + then obtain T1 T2 where a1: "(x,T1)#\ \ t1 : T2" and a2: "T = T1 \ T2" using vc by auto + have ih: "(x,T1)#\ \ t1 : T2 \ (x,T1)#\ \ t2 : T2" by fact + with a1 a2 show "\ \ Lam [x].t2 : T" using vc by auto +next + case (b4 x s t \ T) + have vc: "x\\" by fact + have "\ \ App (Lam [x].t) s : T" by fact + then obtain T' where a1: "\ \ Lam [x].t : T' \ T" and a2: "\ \ s : T'" by auto + from a1 obtain T1 T2 where a3: "(x,T')#\ \ t : T" using vc by (auto simp add: ty.inject) + from a3 a2 show "\ \ t[x::=s] : T" by (simp add: type_substitution) +qed + + +theorem type_preservation_automatic: + assumes a: "t \\<^isub>\ t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +using a b +by (nominal_induct avoiding: \ T rule: beta.strong_induct) + (auto dest!: t_Lam_inversion t_App_inversion simp add: ty.inject type_substitution) + +end