# HG changeset patch # User urbanc # Date 1137001591 -3600 # Node ID dde117622dace696123b023d1f06cf8d8ef9eb46 # Parent 9968dc816cda3a778aa24b0346d4b167c4e1e320 updated to new induction principle diff -r 9968dc816cda -r dde117622dac src/HOL/Nominal/Examples/Class.thy --- a/src/HOL/Nominal/Examples/Class.thy Wed Jan 11 18:39:19 2006 +0100 +++ b/src/HOL/Nominal/Examples/Class.thy Wed Jan 11 18:46:31 2006 +0100 @@ -1,5 +1,5 @@ -theory class +theory Class imports "../nominal" begin @@ -12,153 +12,7 @@ | ImpL "\coname\trm" "\name\trm" "name"("ImpL [_]._ [_]._ _" [100,100,100,100,100] 100) | Cut "\coname\trm" "\name\trm" ("Cut [_]._ [_]._" [100,100,100,100] 100) -lemma trm_induct_aux: - fixes P :: "trm \ 'a \ bool" - and f1 :: "'a \ name set" - and f2 :: "'a \ coname set" - assumes fs1: "\x. finite (f1 x)" - and fs2: "\x. finite (f2 x)" - and h1: "\k x a. P (Ax x a) k" - and h2: "\k x a t b. x\f1 k \ a\f2 k \ (\l. P t l) \ P (ImpR [x].[a].t b) k" - and h3: "\k a t1 x t2 y. a\f2 k \ x\f1 k \ (\l. P t1 l) \ (\l. P t2 l) - \ P (ImpL [a].t1 [x].t2 y) k" - and h4: "\k a t1 x t2. a\f2 k \ x\f1 k \ (\l. P t1 l) \ (\l. P t2 l) - \ P (Cut [a].t1 [x].t2) k" - shows "\(pi1::name prm) (pi2::coname prm) k. P (pi1\(pi2\t)) k" -proof (induct rule: trm.induct_weak) - case (goal1 a) - show ?case using h1 by simp -next - case (goal2 x a t b) - assume i1: "\(pi1::name prm)(pi2::coname prm) k. P (pi1\(pi2\t)) k" - show ?case - proof (intro strip, simp add: abs_perm perm_dj) - fix pi1::"name prm" and pi2::"coname prm" and k::"'a" - have "\u::name. u\(f1 k,pi1\x,pi1\(pi2\t))" - by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1 - at_fin_set_supp[OF at_name_inst, OF fs1] fs1) - then obtain u::"name" - where f1: "u\(pi1\x)" and f2: "u\(f1 k)" and f3: "u\(pi1\(pi2\t))" - by (auto simp add: fresh_prod at_fresh[OF at_name_inst]) - have "\c::coname. c\(f2 k,pi2\a,pi1\(pi2\t))" - by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1 - at_fin_set_supp[OF at_coname_inst, OF fs2] fs2) - then obtain c::"coname" - where e1: "c\(pi2\a)" and e2: "c\(f2 k)" and e3: "c\(pi1\(pi2\t))" - by (auto simp add: fresh_prod at_fresh[OF at_coname_inst]) - have g: "ImpR [u].[c].([(u,pi1\x)]\(pi1\([(c,pi2\a)]\(pi2\t)))) (pi2\b) - =ImpR [(pi1\x)].[(pi2\a)].(pi1\(pi2\t)) (pi2\b)" using f1 f3 e1 e3 - apply (auto simp add: ImpR_inject alpha abs_fresh abs_perm perm_dj, - simp add: dj_cp[OF cp_name_coname_inst, OF dj_coname_name]) - apply(simp add: pt_fresh_left_ineq[OF pt_name_inst, OF pt_name_inst, - OF at_name_inst, OF cp_name_coname_inst] perm_dj) - done - from i1 have "\k. P (([(u,pi1\x)]@pi1)\(([(c,pi2\a)]@pi2)\t)) k" by force - hence i1b: "\k. P ([(u,pi1\x)]\(pi1\([(c,pi2\a)]\(pi2\t)))) k" - by (simp add: pt_name2[symmetric] pt_coname2[symmetric]) - with h2 f2 e2 have "P (ImpR [u].[c].([(u,pi1\x)]\(pi1\([(c,pi2\a)]\(pi2\t)))) (pi2\b)) k" - by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1] - at_fin_set_supp[OF at_coname_inst, OF fs2]) - with g show "P (ImpR [(pi1\x)].[(pi2\a)].(pi1\(pi2\t)) (pi2\b)) k" by simp - qed -next - case (goal3 a t1 x t2 y) - assume i1: "\(pi1::name prm)(pi2::coname prm) k. P (pi1\(pi2\t1)) k" - and i2: "\(pi1::name prm)(pi2::coname prm) k. P (pi1\(pi2\t2)) k" - show ?case - proof (intro strip, simp add: abs_perm) - fix pi1::"name prm" and pi2::"coname prm" and k::"'a" - have "\u::name. u\(f1 k,pi1\x,pi1\(pi2\t2))" - by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1 - at_fin_set_supp[OF at_name_inst, OF fs1] fs1) - then obtain u::"name" - where f1: "u\(pi1\x)" and f2: "u\(f1 k)" and f3: "u\(pi1\(pi2\t2))" - by (auto simp add: fresh_prod at_fresh[OF at_name_inst]) - have "\c::coname. c\(f2 k,pi2\a,pi1\(pi2\t1))" - by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1 - at_fin_set_supp[OF at_coname_inst, OF fs2] fs2) - then obtain c::"coname" - where e1: "c\(pi2\a)" and e2: "c\(f2 k)" and e3: "c\(pi1\(pi2\t1))" - by (auto simp add: fresh_prod at_fresh[OF at_coname_inst]) - have g: "ImpL [c].([(c,pi2\a)]\(pi1\(pi2\t1))) [u].([(u,pi1\x)]\(pi1\(pi2\t2))) (pi1\y) - =ImpL [(pi2\a)].(pi1\(pi2\t1)) [(pi1\x)].(pi1\(pi2\t2)) (pi1\y)" using f1 f3 e1 e3 - by (simp add: ImpL_inject alpha abs_fresh abs_perm perm_dj) - from i2 have "\k. P (([(u,pi1\x)]@pi1)\(pi2\t2)) k" by force - hence i2b: "\k. P ([(u,pi1\x)]\(pi1\(pi2\t2))) k" - by (simp add: pt_name2[symmetric]) - from i1 have "\k. P (pi1\(([(c,pi2\a)]@pi2)\t1)) k" by force - hence i1b: "\k. P ([(c,pi2\a)]\(pi1\(pi2\t1))) k" - by (simp add: pt_coname2[symmetric] dj_cp[OF cp_name_coname_inst, OF dj_coname_name]) - from h3 f2 e2 i1b i2b - have "P (ImpL [c].([(c,pi2\a)]\(pi1\(pi2\t1))) [u].([(u,pi1\x)]\(pi1\(pi2\t2))) (pi1\y)) k" - by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1] - at_fin_set_supp[OF at_coname_inst, OF fs2]) - with g show "P (ImpL [(pi2\a)].(pi1\(pi2\t1)) [(pi1\x)].(pi1\(pi2\t2)) (pi1\y)) k" by simp - qed -next - case (goal4 a t1 x t2) - assume i1: "\(pi1::name prm)(pi2::coname prm) k. P (pi1\(pi2\t1)) k" - and i2: "\(pi1::name prm)(pi2::coname prm) k. P (pi1\(pi2\t2)) k" - show ?case - proof (intro strip, simp add: abs_perm) - fix pi1::"name prm" and pi2::"coname prm" and k::"'a" - have "\u::name. u\(f1 k,pi1\x,pi1\(pi2\t2))" - by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1 - at_fin_set_supp[OF at_name_inst, OF fs1] fs1) - then obtain u::"name" - where f1: "u\(pi1\x)" and f2: "u\(f1 k)" and f3: "u\(pi1\(pi2\t2))" - by (auto simp add: fresh_prod at_fresh[OF at_name_inst]) - have "\c::coname. c\(f2 k,pi2\a,pi1\(pi2\t1))" - by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1 - at_fin_set_supp[OF at_coname_inst, OF fs2] fs2) - then obtain c::"coname" - where e1: "c\(pi2\a)" and e2: "c\(f2 k)" and e3: "c\(pi1\(pi2\t1))" - by (auto simp add: fresh_prod at_fresh[OF at_coname_inst]) - have g: "Cut [c].([(c,pi2\a)]\(pi1\(pi2\t1))) [u].([(u,pi1\x)]\(pi1\(pi2\t2))) - =Cut [(pi2\a)].(pi1\(pi2\t1)) [(pi1\x)].(pi1\(pi2\t2))" using f1 f3 e1 e3 - by (simp add: Cut_inject alpha abs_fresh abs_perm perm_dj) - from i2 have "\k. P (([(u,pi1\x)]@pi1)\(pi2\t2)) k" by force - hence i2b: "\k. P ([(u,pi1\x)]\(pi1\(pi2\t2))) k" - by (simp add: pt_name2[symmetric]) - from i1 have "\k. P (pi1\(([(c,pi2\a)]@pi2)\t1)) k" by force - hence i1b: "\k. P ([(c,pi2\a)]\(pi1\(pi2\t1))) k" - by (simp add: pt_coname2[symmetric] dj_cp[OF cp_name_coname_inst, OF dj_coname_name]) - from h3 f2 e2 i1b i2b - have "P (Cut [c].([(c,pi2\a)]\(pi1\(pi2\t1))) [u].([(u,pi1\x)]\(pi1\(pi2\t2)))) k" - by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1] - at_fin_set_supp[OF at_coname_inst, OF fs2]) - with g show "P (Cut [(pi2\a)].(pi1\(pi2\t1)) [(pi1\x)].(pi1\(pi2\t2))) k" by simp - qed -qed +thm trm.induct +thm trm.inducts +thm trm.induct' -lemma trm_induct'[case_names Ax ImpR ImpL Cut]: - fixes P :: "trm \ 'a \ bool" - and f1 :: "'a \ name set" - and f2 :: "'a \ coname set" - assumes fs1: "\x. finite (f1 x)" - and fs2: "\x. finite (f2 x)" - and h1: "\k x a. P (Ax x a) k" - and h2: "\k x a t b. x\f1 k \ a\f2 k \ (\l. P t l) \ P (ImpR [x].[a].t b) k" - and h3: "\k a t1 x t2 y. a\f2 k \ x\f1 k \ (\l. P t1 l) \ (\l. P t2 l) - \ P (ImpL [a].t1 [x].t2 y) k" - and h4: "\k a t1 x t2. a\f2 k \ x\f1 k \ (\l. P t1 l) \ (\l. P t2 l) - \ P (Cut [a].t1 [x].t2) k" - shows "P t k" -proof - - have "\(pi1::name prm)(pi2::coname prm) k. P (pi1\(pi2\t)) k" - using fs1 fs2 h1 h2 h3 h4 by (rule trm_induct_aux, auto) - hence "P (([]::name prm)\(([]::coname prm)\t)) k" by blast - thus "P t k" by simp -qed - -lemma trm_induct[case_names Ax ImpR ImpL Cut]: - fixes P :: "trm \ ('a::{fs_name,fs_coname}) \ bool" - assumes h1: "\k x a. P (Ax x a) k" - and h2: "\k x a t b. x\k \ a\k \ (\l. P t l) \ P (ImpR [x].[a].t b) k" - and h3: "\k a t1 x t2 y. a\k \ x\k \ (\l. P t1 l) \ (\l. P t2 l) - \ P (ImpL [a].t1 [x].t2 y) k" - and h4: "\k a t1 x t2. a\k \ x\k \ (\l. P t1 l) \ (\l. P t2 l) - \ P (Cut [a].t1 [x].t2) k" - shows "P t k" -by (rule trm_induct'[of "\x. ((supp x)::name set)" "\x. ((supp x)::coname set)" "P"], - simp_all add: fs_name1 fs_coname1 fresh_def[symmetric], auto intro: h1 h2 h3 h4)