diff -r 08e4d3cbcf8c -r 894930834ca8 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Sun Apr 11 22:47:45 2010 +0200 +++ b/Nominal/Ex/Lambda.thy Sun Apr 11 22:48:49 2010 +0200 @@ -107,7 +107,12 @@ nominal_datatype ty = TVar string -| TFun ty ty ("_ \ _") +| TFun ty ty + +notation + TFun ("_ \ _") + +declare ty.perm[eqvt] inductive valid :: "(name \ ty) list \ bool" @@ -128,7 +133,7 @@ *} -lemma +lemma assumes a: "valid Gamma" shows "valid (p \ Gamma)" using a @@ -137,6 +142,31 @@ apply(tactic {* my_tac @{context} @{thms valid.intros} 1 *}) done +lemma + "(p \ valid) = valid" +oops + +lemma temp[eqvt_raw]: + "(p \ valid) \ valid" +sorry + +inductive + typing :: "(name\ty) list \ lam \ ty \ bool" ("_ \ _ : _" [60,60,60] 60) +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]: "\atom x \ \; (x, T1) # \ \ t : T2\ \ \ \ Lam x t : T1 \ T2" + +lemma + assumes a: "Gamma \ t : T" + shows "(p \ Gamma) \ (p \ t) : (p \ T)" +using a +apply(induct) +apply(tactic {* my_tac @{context} @{thms typing.intros} 1 *}) +apply(tactic {* my_tac @{context} @{thms typing.intros} 1 *}) +apply(tactic {* my_tac @{context} @{thms typing.intros} 1 *}) +done + declare permute_lam_raw.simps[eqvt] thm alpha_gen_real_eqvt[no_vars] @@ -173,6 +203,30 @@ apply(perm_simp permute_minus_cancel(2)) oops +thm alpha_lam_raw.intros[no_vars] + +inductive + alpha_lam_raw' +where + "name = namea \ alpha_lam_raw' (Var_raw name) (Var_raw namea)" +| "\alpha_lam_raw' lam_raw1 lam_raw1a; alpha_lam_raw' lam_raw2 lam_raw2a\ \ + alpha_lam_raw' (App_raw lam_raw1 lam_raw2) (App_raw lam_raw1a lam_raw2a)" +| "\pi. ({atom name}, lam_raw) \gen alpha_lam_raw fv_lam_raw pi ({atom namea}, lam_rawa) \ + alpha_lam_raw' (Lam_raw name lam_raw) (Lam_raw namea lam_rawa)" + +lemma + assumes a: "alpha_lam_raw' t1 t2" + shows "alpha_lam_raw' (p \ t1) (p \ t2)" +using a +apply(induct) +apply(tactic {* my_tac @{context} @{thms alpha_lam_raw'.intros} 1 *}) +apply(tactic {* my_tac @{context} @{thms alpha_lam_raw'.intros} 1 *}) +apply(perm_strict_simp) +apply(rule alpha_lam_raw'.intros) +apply(simp add: alphas) +apply(rule_tac p="- p" in permute_boolE) +apply(perm_simp permute_minus_cancel(2)) +oops section {* size function *}