# HG changeset patch # User urbanc # Date 1162210071 -3600 # Node ID e69c0e82955a767ab1222c58d09a3cba877fca81 # Parent 51599a81b30804140029109aa820f8ed24e8ef52 new file for defining functions in the lambda-calculus diff -r 51599a81b308 -r e69c0e82955a src/HOL/Nominal/Examples/Lam_Funs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nominal/Examples/Lam_Funs.thy Mon Oct 30 13:07:51 2006 +0100 @@ -0,0 +1,244 @@ +(* $Id: *) + +theory Lam_Funs +imports Nominal +begin + +atom_decl name + +nominal_datatype lam = Var "name" + | App "lam" "lam" + | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + + +constdefs + depth_Var :: "name \ nat" + "depth_Var \ \(a::name). 1" + + depth_App :: "lam \ lam \ nat \ nat \ nat" + "depth_App \ \_ _ n1 n2. (max n1 n2)+1" + + depth_Lam :: "name \ lam \ nat \ nat" + "depth_Lam \ \_ _ n. n+1" + + depth_lam :: "lam \ nat" + "depth_lam t \ (lam_rec depth_Var depth_App depth_Lam) t" + +lemma fin_supp_depth_lam: + shows "finite ((supp depth_Var)::name set)" + and "finite ((supp depth_Lam)::name set)" + and "finite ((supp depth_App)::name set)" + by (finite_guess add: depth_Var_def depth_Lam_def depth_App_def perm_nat_def)+ + +lemma fcb_depth_lam: + fixes a::"name" + shows "a\(depth_Lam a t n)" + by (simp add: fresh_nat) + +lemma depth_lam: + shows "depth_lam (Var a) = 1" + and "depth_lam (App t1 t2) = (max (depth_lam t1) (depth_lam t2))+1" + and "depth_lam (Lam [a].t) = (depth_lam t)+1" +apply - +apply(unfold depth_lam_def) +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp_all add: fin_supp_depth_lam supp_unit) +apply(simp add: fcb_depth_lam) +apply(simp add: depth_Var_def) +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp_all add: fin_supp_depth_lam supp_unit) +apply(simp add: fcb_depth_lam) +apply(simp add: depth_App_def) +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp_all add: fin_supp_depth_lam supp_unit) +apply(simp add: fcb_depth_lam) +apply(simp add: depth_Var_def, fresh_guess add: perm_nat_def) +apply(simp add: depth_App_def, fresh_guess add: perm_nat_def) +apply(simp add: depth_Lam_def, fresh_guess add: perm_nat_def) +apply(simp add: depth_Lam_def) +done + +text {* capture-avoiding substitution *} +constdefs + subst_Var :: "name \lam \ name \ lam" + "subst_Var b t \ \a. (if a=b then t else (Var a))" + + subst_App :: "name \ lam \ lam \ lam \ lam \ lam \ lam" + "subst_App b t \ \_ _ r1 r2. App r1 r2" + + subst_Lam :: "name \ lam \ name \ lam \ lam \ lam" + "subst_Lam b t \ \a _ r. Lam [a].r" + +abbreviation + subst_syn :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 900) + "t'[b::=t] \ (lam_rec (subst_Var b t) (subst_App b t) (subst_Lam b t)) t'" + +(* FIXME: this lemma needs to be in Nominal.thy *) +lemma eq_eqvt: + fixes pi::"name prm" + and x::"'a::pt_name" + shows "pi\(x=y) = ((pi\x)=(pi\y))" + apply(simp add: perm_bool perm_bij) + done + +lemma fin_supp_subst: + shows "finite ((supp (subst_Var b t))::name set)" + and "finite ((supp (subst_Lam b t))::name set)" + and "finite ((supp (subst_App b t))::name set)" +apply - +apply(finite_guess add: subst_Var_def perm_if eq_eqvt fs_name1) +apply(finite_guess add: subst_Lam_def fs_name1) +apply(finite_guess add: subst_App_def fs_name1) +done + +lemma fcb_subst: + fixes a::"name" + shows "a\(subst_Lam b t) a t' r" + by (simp add: subst_Lam_def abs_fresh) + +lemma subst[simp]: + shows "(Var a)[b::=t] = (if a=b then t else (Var a))" + and "(App t1 t2)[b::=t] = App (t1[b::=t]) (t2[b::=t])" + and "a\(b,t) \ (Lam [a].t1)[b::=t] = Lam [a].(t1[b::=t])" + and "\a\b; a\t\ \ (Lam [a].t1)[b::=t] = Lam [a].(t1[b::=t])" + and "\a\b; a\t\ \ (Lam [a].t1)[b::=t] = Lam [a].(t1[b::=t])" +apply - +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp add: fs_name1 fin_supp_subst)+ +apply(simp add: fcb_subst) +apply(simp add: subst_Var_def) +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp add: fs_name1 fin_supp_subst)+ +apply(simp add: fcb_subst) +apply(simp add: subst_App_def) +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp add: fs_name1 fin_supp_subst)+ +apply(simp add: fcb_subst) +apply(fresh_guess add: subst_Var_def perm_if eq_eqvt fs_name1) +apply(fresh_guess add: subst_App_def fs_name1) +apply(fresh_guess add: subst_Lam_def fs_name1) +apply(simp add: subst_Lam_def) +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp add: fs_name1 fin_supp_subst)+ +apply(simp add: fcb_subst) +apply(fresh_guess add: subst_Var_def perm_if eq_eqvt fs_name1 fresh_atm) +apply(fresh_guess add: subst_App_def fs_name1 fresh_atm) +apply(fresh_guess add: subst_Lam_def fs_name1 fresh_atm) +apply(simp add: subst_Lam_def) +apply(rule trans) +apply(rule lam.recs[where P="\_. True" and z="()", simplified]) +apply(simp add: fs_name1 fin_supp_subst)+ +apply(simp add: fcb_subst) +apply(fresh_guess add: subst_Var_def perm_if eq_eqvt fs_name1) +apply(fresh_guess add: subst_App_def fs_name1) +apply(fresh_guess add: subst_Lam_def fs_name1) +apply(simp add: subst_Lam_def) +done + +lemma subst_eqvt[simp]: + fixes pi:: "name prm" + and t1:: "lam" + and t2:: "lam" + and a :: "name" + shows "pi\(t1[b::=t2]) = (pi\t1)[(pi\b)::=(pi\t2)]" +apply(nominal_induct t1 avoiding: b t2 rule: lam.induct) +apply(auto simp add: perm_bij fresh_prod fresh_atm fresh_bij) +done + +lemma subst_supp: + shows "supp(t1[a::=t2]) \ (((supp(t1)-{a})\supp(t2))::name set)" +apply(nominal_induct t1 avoiding: a t2 rule: lam.induct) +apply(auto simp add: lam.supp supp_atm fresh_prod abs_supp) +apply(blast)+ +done + +(* parallel substitution *) + +(* simultaneous substitution *) +consts + "domain" :: "(name\lam) list \ (name list)" +primrec + "domain [] = []" + "domain (x#\) = (fst x)#(domain \)" + +consts + "apply_sss" :: "(name\lam) list \ name \ lam" (" _ < _ >" [80,80] 80) +primrec +"(x#\) = (if (a = fst x) then (snd x) else \)" + +lemma apply_sss_eqvt: + fixes pi::"name prm" + assumes a: "a\set (domain \)" + shows "pi\(\) = (pi\\)a>" +using a +by (induct \) + (auto simp add: perm_bij) + +lemma domain_eqvt: + fixes pi::"name prm" + shows "((pi\a)\set (domain \)) = (a\set (domain ((rev pi)\\)))" +apply(induct \) +apply(auto) +apply(perm_simp)+ +done + +constdefs + psubst_Var :: "(name\lam) list \ name \ lam" + "psubst_Var \ \ \a. (if a\set (domain \) then \ else (Var a))" + + psubst_App :: "(name\lam) list \ lam \ lam \ lam \ lam \ lam" + "psubst_App \ \ \_ _ r1 r2. App r1 r2" + + psubst_Lam :: "(name\lam) list \ name \ lam \ lam \ lam" + "psubst_Lam \ \ \a _ r. Lam [a].r" + +abbreviation + psubst_lam :: "lam \ (name\lam) list \ lam" ("_[<_>]" [100,100] 900) + "t[<\>] \ (lam_rec (psubst_Var \) (psubst_App \) (psubst_Lam \)) t" + +lemma fin_supp_psubst: + shows "finite ((supp (psubst_Var \))::name set)" + and "finite ((supp (psubst_Lam \))::name set)" + and "finite ((supp (psubst_App \))::name set)" + apply - + apply(finite_guess add: fs_name1 psubst_Var_def domain_eqvt apply_sss_eqvt) + apply(finite_guess add: fs_name1 psubst_Lam_def) + apply(finite_guess add: fs_name1 psubst_App_def) +done + +lemma fcb_psubst_Lam: + shows "x\(psubst_Lam \) x t r" + by (simp add: psubst_Lam_def abs_fresh) + +lemma psubst[simp]: + shows "(Var a)[<\>] = (if a\set (domain \) then \ else (Var a))" + and "(App t1 t2)[<\>] = App (t1[<\>]) (t2[<\>])" + and "a\\ \ (Lam [a].t1)[<\>] = Lam [a].(t1[<\>])" + apply(rule trans) + apply(rule lam.recs[where P="\_. True" and z="()", simplified]) + apply(simp add: fin_supp_psubst fs_name1)+ + apply(simp add: fcb_psubst_Lam) + apply(simp add: psubst_Var_def) + apply(rule trans) + apply(rule lam.recs[where P="\_. True" and z="()", simplified]) + apply(simp add: fin_supp_psubst fs_name1)+ + apply(simp add: fcb_psubst_Lam) + apply(simp add: psubst_App_def) + apply(rule trans) + apply(rule lam.recs[where P="\_. True" and z="()", simplified]) + apply(simp add: fin_supp_psubst fs_name1)+ + apply(simp add: fcb_psubst_Lam) + apply(fresh_guess add: psubst_Var_def domain_eqvt apply_sss_eqvt fs_name1) + apply(fresh_guess add: psubst_App_def fs_name1) + apply(fresh_guess add: psubst_Lam_def fs_name1) + apply(simp add: psubst_Lam_def) +done + +end \ No newline at end of file diff -r 51599a81b308 -r e69c0e82955a src/HOL/Nominal/Examples/SN.thy --- a/src/HOL/Nominal/Examples/SN.thy Thu Oct 26 16:08:40 2006 +0200 +++ b/src/HOL/Nominal/Examples/SN.thy Mon Oct 30 13:07:51 2006 +0100 @@ -1,7 +1,7 @@ (* $Id$ *) theory SN -imports Lam_substs +imports Lam_Funs begin text {* Strong Normalisation proof from the Proofs and Types book *} @@ -343,7 +343,7 @@ thus "P x \ t \" by simp qed -constdefs +abbreviation "sub" :: "(name\ty) list \ (name\ty) list \ bool" (" _ \ _ " [80,80] 80) "\1 \ \2 \ \a \. (a,\)\set \1 \ (a,\)\set \2" @@ -354,11 +354,9 @@ shows "\2 \ t:\" using a b c apply(nominal_induct \1 t \ avoiding: \2 rule: typing_induct) -apply(auto simp add: sub_def) +apply(auto | atomize)+ (* FIXME: before using meta-connectives and the new induction *) (* method, this was completely automatic *) -apply(atomize) -apply(auto) done lemma in_ctxt: @@ -456,7 +454,7 @@ by (auto dest: valid_elim simp add: fresh_list_cons) from c12 have c14: "((c,\)#(a,\1)#\) \ s : \2" proof - - have c2: "((a,\1)#(c,\)#\) \ ((c,\)#(a,\1)#\)" by (force simp add: sub_def) + have c2: "((a,\1)#(c,\)#\) \ ((c,\)#(a,\1)#\)" by force have c3: "valid ((c,\)#(a,\1)#\)" by (rule v2, rule v2, auto simp add: fresh_list_cons fresh_prod ca cb cc f2' fresh_ty) from c12 c2 c3 show ?thesis by (force intro: weakening) @@ -464,7 +462,7 @@ assume c8: "\ \ t2 : \" have c81: "((a,\1)#\)\t2 :\" proof - - have c82: "\ \ ((a,\1)#\)" by (force simp add: sub_def) + have c82: "\ \ ((a,\1)#\)" by force have c83: "valid ((a,\1)#\)" using f1 ca by force with c8 c82 c83 show ?thesis by (force intro: weakening) qed @@ -857,6 +855,7 @@ (*A*) apply(simp add: fresh_list_cons fresh_prod) done + lemma all_RED: assumes a: "\\t:\" and b: "\a \. (a,\)\set(\) \ (a\set(domain \)\\\RED \)" diff -r 51599a81b308 -r e69c0e82955a src/HOL/Nominal/Examples/Weakening.thy --- a/src/HOL/Nominal/Examples/Weakening.thy Thu Oct 26 16:08:40 2006 +0200 +++ b/src/HOL/Nominal/Examples/Weakening.thy Mon Oct 30 13:07:51 2006 +0100 @@ -124,6 +124,72 @@ thus "P x \ t \" by simp qed +lemma typing_induct_test[consumes 1, case_names t_Var t_App t_Lam]: + fixes P :: "'a::fs_name\(name\ty) list \ lam \ ty \bool" + and \ :: "(name\ty) list" + and t :: "lam" + and \ :: "ty" + and x :: "'a::fs_name" + assumes a: "\ \ t : \" + and a1: "\\ a \ x. \valid \; (a,\) \ set \\ \ P x \ (Var a) \" + and a2: "\\ \ \ t1 t2 x. + \\ \ t1 : \\\; \z. P z \ t1 (\\\); \ \ t2 : \; \z. P z \ t2 \\ + \ P x \ (App t1 t2) \" + and a3: "\a \ \ \ t x. \a\x; a\\; ((a,\)#\) \ t : \; \z. P z ((a,\)#\) t \\ + \ P x \ (Lam [a].t) (\\\)" + shows "P x \ t \" +proof - + from a have "\(pi::name prm) x. P x (pi\\) (pi\t) \" + proof (induct) + case (t_Var \ a \) + have "valid \" by fact + then have "valid (pi\\)" by (rule eqvt_valid) + moreover + have "(a,\)\set \" by fact + then have "pi\(a,\)\pi\(set \)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + then have "(pi\a,\)\set (pi\\)" by (simp add: pt_list_set_pi[OF pt_name_inst]) + ultimately show "P x (pi\\) (pi\(Var a)) \" using a1 by simp + next + case (t_App \ t1 \ \ t2) + thus "P x (pi\\) (pi\(App t1 t2)) \" using a2 by (simp, blast intro: eqvt_typing) + next + case (t_Lam a \ \ t \ pi x) + have p1: "((a,\)#\)\t:\" by fact + have ih1: "\(pi::name prm) x. P x (pi\((a,\)#\)) (pi\t) \" by fact + have f: "a\\" by fact + then have f': "(pi\a)\(pi\\)" by (simp add: fresh_bij) + have "\c::name. c\(pi\a,pi\t,pi\\,x)" + by (rule exists_fresh, simp add: fs_name1) + then obtain c::"name" + where fs: "c\(pi\a)" "c\x" "c\(pi\t)" "c\(pi\\)" + by (force simp add: fresh_prod fresh_atm) + let ?pi'="[(pi\a,c)]@pi" + have eq: "((pi\a,c)#pi)\a = c" by (simp add: calc_atm) + have p1': "(?pi'\((a,\)#\))\(?pi'\t):\" using p1 by (simp only: eqvt_typing) + have ih1': "\x. P x (?pi'\((a,\)#\)) (?pi'\t) \" using ih1 by simp + have "P x (?pi'\\) (?pi'\(Lam [a].t)) (\\\)" using f f' fs p1' ih1' eq + apply - + apply(simp del: append_Cons) + apply(rule a3) + apply(simp_all add: fresh_left calc_atm pt_name2) + done + then have "P x ([(pi\a,c)]\(pi\\)) ([(pi\a,c)]\(Lam [(pi\a)].(pi\t))) (\\\)" + by (simp del: append_Cons add: pt_name2) + then show "P x (pi\\) (pi\(Lam [a].t)) (\ \ \)" using f f' fs + apply - + apply(subgoal_tac "c\Lam [(pi\a)].(pi\t)") + apply(subgoal_tac "(pi\a)\Lam [(pi\a)].(pi\t)") + apply(simp only: perm_fresh_fresh) + apply(simp) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + done + qed + hence "P x (([]::name prm)\\) (([]::name prm)\t) \" by blast + thus "P x \ t \" by simp +qed + + text {* definition of a subcontext *} abbreviation