diff -r 1d55a607c81b -r 887d8bd4eb99 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Fri Jun 03 18:33:11 2011 +0900 +++ b/Nominal/Ex/Lambda.thy Fri Jun 03 12:46:23 2011 +0100 @@ -12,6 +12,19 @@ lemma cheat: "P" sorry +thm lam.strong_exhaust + +lemma lam_strong_exhaust2: + "\\name. y = Var name \ P; + \lam1 lam2. y = App lam1 lam2 \ P; + \name lam. \{atom name} \* c; y = Lam [name]. lam\ \ P; + finite (supp c)\ + \ P" +sorry + +abbreviation + "FCB f \ \x t r. atom x \ f x t r" + lemma Abs1_eq_fdest: fixes x y :: "'a :: at_base" and S T :: "'b :: fs" @@ -40,39 +53,43 @@ using fresh_fun_eqvt_app[OF a b(1)] a b by (metis fresh_fun_app) +locale test = + fixes f1::"name \ ('a::pt)" + and f2::"lam \ lam \ 'a \ 'a \ ('a::pt)" + and f3::"name \ lam \ 'a \ ('a::pt)" + assumes fs: "finite (supp (f1, f2, f3))" + and eq: "eqvt f1" "eqvt f2" "eqvt f3" + and fcb: "\x t r. atom x \ f3 x t r" +begin + nominal_primrec - f :: "(name \ 'a\{pt}) \ (lam \ lam \ 'a \ 'a \ 'a) \ (name \ lam \ 'a \ 'a) \ lam \ 'a" + f :: "lam \ ('a::pt)" where - "f f1 f2 f3 (Var x) = f1 x" -| "f f1 f2 f3 (App t1 t2) = f2 t1 t2 (f f1 f2 f3 t1) (f f1 f2 f3 t2)" -| "(\a t r. atom a \ f3 a t r) \ (eqvt f1 \ eqvt f2 \ eqvt f3) \ atom x \ (f1,f2,f3) \ (f f1 f2 f3 (Lam [x].t)) = f3 x t (f f1 f2 f3 t)" + "f (Var x) = f1 x" +| "f (App t1 t2) = f2 t1 t2 (f t1) (f t2)" +| "atom x \ (f1, f2, f3) \ f (Lam [x].t) = f3 x t (f t)" apply (simp add: eqvt_def f_graph_def) - apply (rule, perm_simp, rule) - apply(case_tac x) - apply(rule_tac y="d" and c="z" in lam.strong_exhaust) + apply (perm_simp) + apply(simp add: eq[simplified eqvt_def]) + apply(rule_tac y="x" and c="(f1,f2,f3)" in lam_strong_exhaust2) apply(auto simp add: fresh_star_def) - apply(blast) - defer + apply(rule fs) apply(simp add: fresh_Pair_elim) apply(erule Abs1_eq_fdest) apply simp_all - apply (rule_tac f="f3a" in fresh_fun_eqvt_app3) - apply assumption - apply (simp add: fresh_at_base) - apply assumption - apply (erule fresh_eqvt_at) - apply (simp add: supp_Pair supp_fun_eqvt finite_supp) - apply (simp add: fresh_Pair) - apply (subgoal_tac "\p y r. p \ (f3a x y r) = f3a (p \ x) (p \ y) (p \ r)") + apply(simp add: fcb) + apply (subgoal_tac "\p y r. p \ (f3 x y r) = f3 (p \ x) (p \ y) (p \ r)") apply (simp add: eqvt_at_def eqvt_def) - apply (simp add: permute_fun_app_eq) - apply (simp add: eqvt_def) - sorry (*this could be defined *) + apply (simp add: permute_fun_app_eq eq[simplified eqvt_def]) + sorry termination sorry -thm f.simps +end +thm test.f.simps + +thm test_def