# HG changeset patch # User urbanc # Date 1133288798 -3600 # Node ID cd217d16c90d7fee88640aa37edc41f8a5c4ac60 # Parent f8a49f09a20215e6ff88f7259351546501cd0a5e changed the order of the induction variable and the context in lam_induct (test to see whether it simplifies the code for nominal_induct). diff -r f8a49f09a202 -r cd217d16c90d src/HOL/Nominal/Examples/Lam_substs.thy --- a/src/HOL/Nominal/Examples/Lam_substs.thy Tue Nov 29 18:09:12 2005 +0100 +++ b/src/HOL/Nominal/Examples/Lam_substs.thy Tue Nov 29 19:26:38 2005 +0100 @@ -17,18 +17,18 @@ section {* Strong induction principles for lam *} lemma lam_induct_aux: - fixes P :: "lam \ 'a \ bool" + fixes P :: "'a \ lam \ bool" and f :: "'a \ 'a" assumes fs: "\x. finite ((supp (f x))::name set)" - and h1: "\x a. P (Var a) x" - and h2: "\x t1 t2. (\z. P t1 z) \ (\z. P t2 z) \ P (App t1 t2) x" - and h3: "\x a t. a\f x \ (\z. P t z) \ P (Lam [a].t) x" - shows "\(pi::name prm) x. P (pi\t) x" + and h1: "\x a. P x (Var a)" + and h2: "\x t1 t2. (\z. P z t1) \ (\z. P z t2) \ P x (App t1 t2)" + and h3: "\x a t. a\f x \ (\z. P z t) \ P x (Lam [a].t)" + shows "\(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 (pi\t) x" by fact - show "\(pi::name prm) x. P (pi\(Lam [a].t)) x" - proof (intro strip, simp add: abs_perm) + 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 add: abs_perm) 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) @@ -38,40 +38,42 @@ 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 (([(c,pi\a)]@pi)\t) x" by force - hence "\x. P ([(c,pi\a)]\(pi\t)) x" by (simp add: pt_name2[symmetric]) - hence "P (Lam [c].([(c,pi\a)]\(pi\t))) x" using h3 f2 + 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 (Lam [(pi\a)].(pi\t)) x" by simp + ultimately show "P x (Lam [(pi\a)].(pi\t))" by simp qed qed (auto intro: h1 h2) lemma lam_induct'[case_names Fin Var App Lam]: - fixes P :: "lam \ 'a \ bool" + fixes P :: "'a \ lam \ bool" and x :: "'a" and t :: "lam" and f :: "'a \ 'a" assumes fs: "\x. finite ((supp (f x))::name set)" - and h1: "\x a. P (Var a) x" - and h2: "\x t1 t2. (\z. P t1 z)\(\z. P t2 z)\P (App t1 t2) x" - and h3: "\x a t. a\f x \ (\z. P t z)\ P (Lam [a].t) x" - shows "P t x" + and h1: "\x a. P x (Var a)" + and h2: "\x t1 t2. (\z. P z t1)\(\z. P z t2)\P x (App t1 t2)" + and h3: "\x a t. a\f x \ (\z. P z t)\ P x (Lam [a].t)" + shows "P x t" proof - - from fs h1 h2 h3 have "\(pi::name prm) x. P (pi\t) x" by (rule lam_induct_aux, auto) - hence "P (([]::name prm)\t) x" by blast - thus "P t x" by simp + from fs h1 h2 h3 have "\(pi::name prm) x. P x (pi\t)" by (rule lam_induct_aux, auto) + 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 :: "lam \ ('a::fs_name) \ bool" + fixes P :: "('a::fs_name) \ lam \ bool" and x :: "'a::fs_name" and t :: "lam" - assumes h1: "\x a. P (Var a) x" - and h2: "\x t1 t2. (\z. P t1 z)\(\z. P t2 z)\P (App t1 t2) x" - and h3: "\x a t. a\x \ (\z. P t z)\ P (Lam [a].t) x" - shows "P t x" -by (rule lam_induct'[of "\x. x" "P"], - simp_all add: fs_name1 fresh_def[symmetric], auto intro: h1 h2 h3) + assumes h1: "\x a. P x (Var a)" + and h2: "\x t1 t2. (\z. P z t1)\(\z. P z t2)\P x (App t1 t2)" + and h3: "\x a t. 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)" @@ -925,7 +927,7 @@ 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(nominal_induct t rule: lam_induct')*) -apply (rule lam_induct'[of "\_. (f1,f2,f3)" "\t _. fun t = rfun f1 f2 f3 t"]) +apply (rule lam_induct'[of "\_. (f1,f2,f3)" "\_ t. fun t = rfun f1 f2 f3 t"]) (* finite support *) apply(rule f) (* VAR *) @@ -1039,7 +1041,7 @@ 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 "\_. (f1,f2,f3)" "\_ t. fst (rfun_gen' f1 f2 f3 t) = t"]) apply(simp add: f) apply(unfold rfun_gen'_def) apply(simp only: rfun_Var[OF fsupp'[OF f],OF fcb'[OF f, OF c]]) @@ -1370,6 +1372,7 @@ and t2:: "lam" and a :: "name" shows "pi\(t1[a::=t2]) = (pi\t1)[(pi\a)::=(pi\t2)]" +thm lam_induct apply(nominal_induct t1 rule: lam_induct) apply(auto) apply(auto simp add: perm_bij fresh_prod fresh_atm)