diff -r 317a6f0ef8b9 -r 2ff0ae57431d src/HOL/Nominal/Examples/Lam_substs.thy --- a/src/HOL/Nominal/Examples/Lam_substs.thy Wed Jan 11 18:21:23 2006 +0100 +++ b/src/HOL/Nominal/Examples/Lam_substs.thy Wed Jan 11 18:38:32 2006 +0100 @@ -14,56 +14,6 @@ | App "lam" "lam" | Lam "\name\lam" ("Lam [_]._" [100,100] 100) -section {* Strong induction principles for lam *} - -lemma lam_induct'[case_names Fin Var App Lam]: - fixes P :: "'a \ lam \ bool" - and f :: "'a \ 'a" - assumes fs: "\x. finite ((supp (f x))::name set)" - and h1: "\a x. P x (Var a)" - and h2: "\t1 t2 x. (\z. P z t1) \ (\z. P z t2) \ P x (App t1 t2)" - and h3: "\a t x. a\f x \ (\z. P z t) \ P x (Lam [a].t)" - shows "P x t" -proof - - have "\(pi::name prm) x. P x (pi\t)" - proof (induct rule: lam.induct_weak) - case (Lam a t) - have ih: "\(pi::name prm) x. P x (pi\t)" by fact - show "\(pi::name prm) x. P x (pi\(Lam [a].t))" - proof (simp) - fix x::"'a" and pi::"name prm" - have "\c::name. c\(f x,pi\a,pi\t)" - by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1 fs) - then obtain c::"name" - where f1: "c\(pi\a)" and f2: "c\(f x)" and f3: "c\(pi\t)" - by (force simp add: fresh_prod fresh_atm) - have "Lam [c].([(c,pi\a)]\(pi\t)) = Lam [(pi\a)].(pi\t)" using f1 f3 - by (simp add: lam.inject alpha) - moreover - from ih have "\x. P x (([(c,pi\a)]@pi)\t)" by force - hence "\x. P x ([(c,pi\a)]\(pi\t))" by (simp add: pt_name2[symmetric]) - hence "P x (Lam [c].([(c,pi\a)]\(pi\t)))" using h3 f2 - by (auto simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs]) - ultimately show "P x (Lam [(pi\a)].(pi\t))" by simp - qed - qed (auto intro: h1 h2) - hence "P x (([]::name prm)\t)" by blast - thus "P x t" by simp -qed - -lemma lam_induct[case_names Var App Lam]: - fixes P :: "('a::fs_name) \ lam \ bool" - and x :: "'a::fs_name" - and t :: "lam" - assumes h1: "\a x. P x (Var a)" - and h2: "\t1 t2 x. (\z. P z t1)\(\z. P z t2)\P x (App t1 t2)" - and h3: "\a t x. a\x \ (\z. P z t)\ P x (Lam [a].t)" - shows "P x t" -apply(rule lam_induct'[of "\x. x" "P"]) -apply(simp add: fs_name1) -apply(auto intro: h1 h2 h3) -done - types 'a f1_ty = "name\('a::pt_name)" 'a f2_ty = "'a\'a\('a::pt_name)" 'a f3_ty = "name\'a\('a::pt_name)" @@ -915,7 +865,8 @@ and a2: "\t1 t2. fun (App t1 t2) = f2 (fun t1) (fun t2)" and a3: "\a t. a\(f1,f2,f3) \ fun (Lam [a].t) = f3 a (fun t)" shows "fun t = rfun f1 f2 f3 t" -apply (rule lam_induct'[of "\_. (f1,f2,f3)" "\_ t. fun t = rfun f1 f2 f3 t"]) +apply (rule lam.induct'[of "\_. ((supp (f1,f2,f3))::name set)" "\_ t. fun t = rfun f1 f2 f3 t"]) +apply(fold fresh_def) (* finite support *) apply(rule f) (* VAR *) @@ -1029,7 +980,8 @@ assumes f: "finite ((supp (f1,f2,f3))::name set)" and c: "(\(a::name). a\f3 \ (\(y::'a::pt_name) t. a\f3 t a y))" shows "fst (rfun_gen' f1 f2 f3 t) = t" -apply(rule lam_induct'[of "\_. (f1,f2,f3)" "\_ t. fst (rfun_gen' f1 f2 f3 t) = t"]) +apply(rule lam.induct'[of "\_. ((supp (f1,f2,f3))::name set)" "\_ t. fst (rfun_gen' f1 f2 f3 t) = t"]) +apply(fold fresh_def) apply(simp add: f) apply(unfold rfun_gen'_def) apply(simp only: rfun_Var[OF fsupp'[OF f],OF fcb'[OF f, OF c]]) @@ -1362,15 +1314,16 @@ 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(nominal_induct t1 avoiding: b t2 rule: lam.induct) apply(auto simp add: perm_bij fresh_prod fresh_atm) -apply(subgoal_tac "(pi\a)\(pi\b,pi\t2)")(*A*) -apply(simp) +apply(subgoal_tac "(pi\name)\(pi\b,pi\t2)")(*A*) +apply(simp) +(*A*) apply(simp add: perm_bij fresh_prod fresh_atm pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) done lemma subst_supp: "supp(t1[a::=t2])\(((supp(t1)-{a})\supp(t2))::name set)" -apply(nominal_induct t1 avoiding: a t2 rule: lam_induct) +apply(nominal_induct t1 avoiding: a t2 rule: lam.induct) apply(auto simp add: lam.supp supp_atm fresh_prod abs_supp) apply(blast) apply(blast)