diff -r 05e5d68c9627 -r f1be8028a4a9 Nominal/activities/Lambda.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/activities/Lambda.thy Wed Mar 30 17:27:34 2016 +0100 @@ -0,0 +1,87 @@ +theory Lambda + imports "Nominal" +begin + +atom_decl name + +section {* Alpha-Equated Lambda-Terms *} + +nominal_datatype lam = + Var "name" +| App "lam" "lam" +| Lam "\name\lam" ("Lam [_]._") + +section {* Capture-Avoiding Substitution *} + +consts subst :: "lam \ name \ lam \ lam" ("_[_::=_]") + +nominal_primrec + "(Var x)[y::=s] = (if x=y then s else (Var x))" + "(App t1 t2)[y::=s] = App (t1[y::=s]) (t2[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 forget: + assumes a: "x\L" + shows "L[x::=P] = L" + using a +by (nominal_induct L avoiding: x P rule: lam.strong_induct) + (auto simp add: abs_fresh fresh_atm) + +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) + +lemma subst_rename: + assumes a: "y\t" + shows "t[x::=s] = ([(y,x)]\t)[y::=s]" +using a +by (nominal_induct t avoiding: x y s rule: lam.strong_induct) + (auto simp add: calc_atm fresh_atm abs_fresh) + +text {* + The purpose of the two lemmas below is to work + around some quirks in Isabelle's handling of + meta_quantifiers and meta_implications. + *} + +lemma meta_impCE: + assumes major: "P ==> PROP Q" + and 1: "~ P ==> R" + and 2: "PROP Q ==> R" + shows R +proof (cases P) + assume P + then have "PROP Q" by (rule major) + then show R by (rule 2) +next + assume "~ P" + then show R by (rule 1) +qed + +declare meta_allE [elim] + and meta_impCE [elim!] + +end + + + + + + + + + +