# HG changeset patch # User urbanc # Date 1164287509 -3600 # Node ID 45f9163d79e7d3c7ee0bc34729fb115b5555c3bb # Parent b1fdc05138123876fdc344cba56473f026c02bd1 tuned the proof of the strong induction principle diff -r b1fdc0513812 -r 45f9163d79e7 src/HOL/Nominal/Examples/Weakening.thy --- a/src/HOL/Nominal/Examples/Weakening.thy Thu Nov 23 13:32:19 2006 +0100 +++ b/src/HOL/Nominal/Examples/Weakening.thy Thu Nov 23 14:11:49 2006 +0100 @@ -98,94 +98,31 @@ thus "P x (pi\\) (pi\(App t1 t2)) \" using a2 by (simp, blast intro: eqvt_typing) next case (t_Lam a \ \ t \) - have k1: "a\\" by fact - have k2: "((a,\)#\)\t:\" by fact - have k3: "\(pi::name prm) (x::'a::fs_name). P x (pi \((a,\)#\)) (pi\t) \" by fact - obtain c::"name" where "c\(pi\a,pi\t,pi\\,x)" by (erule exists_fresh[OF fs_name1]) - then have f1: "c\(pi\a)" and f2: "c\x" and f3: "c\(pi\t)" and f4: "c\(pi\\)" - by (simp_all add: fresh_atm[symmetric]) - from k1 have k1a: "(pi\a)\(pi\\)" by (simp add: fresh_bij) - have l1: "(([(c,pi\a)]@pi)\\) = (pi\\)" using f4 k1a - by (simp only: pt_name2, rule perm_fresh_fresh) - have "\x. P x (([(c,pi\a)]@pi)\((a,\)#\)) (([(c,pi\a)]@pi)\t) \" using k3 by force - hence l2: "\x. P x ((c, \)#(pi\\)) (([(c,pi\a)]@pi)\t) \" using f1 l1 - by (force simp add: pt_name2 calc_atm) - have "(([(c,pi\a)]@pi)\((a,\)#\)) \ (([(c,pi\a)]@pi)\t) : \" using k2 by (rule eqvt_typing) - hence l3: "((c, \)#(pi\\)) \ (([(c,pi\a)]@pi)\t) : \" using l1 f1 - by (force simp add: pt_name2 calc_atm) - have l4: "P x (pi\\) (Lam [c].(([(c,pi\a)]@pi)\t)) (\ \ \)" using f2 f4 l2 l3 a3 by auto - have alpha: "(Lam [c].([(c,pi\a)]\(pi\t))) = (Lam [(pi\a)].(pi\t))" using f1 f3 - by (simp add: lam.inject alpha) - show "P x (pi\\) (pi\(Lam [a].t)) (\ \ \)" using l4 alpha by (simp only: pt_name2, simp) + obtain c::"name" where fs: "c\(pi\a,pi\t,pi\\,x)" by (rule exists_fresh[OF fs_name1]) + let ?sw="[(pi\a,c)]" + let ?pi'="?sw@pi" + have f1: "a\\" by fact + have f2: "(pi\a)\(pi\\)" using f1 by (simp add: fresh_bij) + have f3: "c\?pi'\\" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp) + have pr1: "((a,\)#\)\t:\" by fact + then have "(?pi'\((a,\)#\)) \ (?pi'\t) : \" by (rule eqvt_typing) + then have "((c,\)#(?pi'\\)) \ (?pi'\t) : \" by (simp add: calc_atm) + moreover + have ih1: "\x. P x (?pi'\((a,\)#\)) (?pi'\t) \" by fact + then have "\x. P x ((c,\)#(?pi'\\)) (?pi'\t) \" by (simp add: calc_atm) + ultimately have "P x (?pi'\\) (Lam [c].(?pi'\t)) (\ \ \)" using a3 f3 fs by simp + then have "P x (?sw\pi\\) (?sw\(Lam [(pi\a)].(pi\t))) (\ \ \)" + by (simp del: append_Cons add: calc_atm pt_name2) + moreover have "(?sw\(pi\\)) = (pi\\)" + by (rule perm_fresh_fresh) (simp_all add: fs f2) + moreover have "(?sw\(Lam [(pi\a)].(pi\t))) = Lam [(pi\a)].(pi\t)" + by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh) + ultimately show "P x (pi\\) (pi\(Lam [a].t)) (\ \ \)" by (simp only: , simp) qed hence "P x (([]::name prm)\\) (([]::name prm)\t) \" by blast 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) - obtain c::"name" where "c\(pi\a,pi\t,pi\\,x)" by (erule exists_fresh[OF fs_name1]) - then have fs: "c\(pi\a)" "c\x" "c\(pi\t)" "c\(pi\\)" - by (simp_all add: fresh_atm[symmetric]) - 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