diff -r 78623a0f294b -r 8f5420681039 Nominal/Ex/Foo2.thy --- a/Nominal/Ex/Foo2.thy Sat Nov 27 23:00:16 2010 +0000 +++ b/Nominal/Ex/Foo2.thy Sun Nov 28 16:37:34 2010 +0000 @@ -61,284 +61,23 @@ apply(simp add: atom_eqvt) done - -lemma Let1_rename: - assumes "supp ([bn assn1]lst. trm1) \* p" "supp ([bn assn2]lst. trm2) \* p" - shows "Let1 assn1 trm1 assn2 trm2 = Let1 (permute_bn p assn1) (p \ trm1) (permute_bn p assn2) (p \ trm2)" -using assms -apply - -apply(drule supp_perm_eq[symmetric]) -apply(drule supp_perm_eq[symmetric]) -apply(simp only: permute_Abs) -apply(simp only: tt1) -apply(simp only: foo.eq_iff) -apply(simp add: uu1) -done - -lemma Let2_rename: - assumes "(supp ([[atom x, atom y]]lst. t1)) \* p" and "(supp ([[atom y]]lst. t2)) \* p" - shows "Let2 x y t1 t2 = Let2 (p \ x) (p \ y) (p \ t1) (p \ t2)" - using assms - apply - - apply(drule supp_perm_eq[symmetric]) - apply(drule supp_perm_eq[symmetric]) - apply(simp only: foo.eq_iff) - apply(simp only: eqvts) - apply simp - done - -lemma Let2_rename2: - assumes "(supp ([[atom x, atom y]]lst. t1)) \* p" and "(atom y) \ p" - shows "Let2 x y t1 t2 = Let2 (p \ x) y (p \ t1) t2" - using assms - apply - - apply(drule supp_perm_eq[symmetric]) - apply(simp only: foo.eq_iff) - apply(simp only: eqvts) - apply simp - by (metis assms(2) atom_eqvt fresh_perm) - -lemma Let2_rename3: - assumes "(supp ([[atom x, atom y]]lst. t1)) \* p" - and "(supp ([[atom y]]lst. t2)) \* p" - and "(atom x) \ p" - shows "Let2 x y t1 t2 = Let2 x (p \ y) (p \ t1) (p \ t2)" - using assms - apply - - apply(drule supp_perm_eq[symmetric]) - apply(drule supp_perm_eq[symmetric]) - apply(simp only: foo.eq_iff) - apply(simp only: eqvts) - apply simp - by (metis assms(2) atom_eqvt fresh_perm) - -lemma strong_exhaust1_pre: - fixes c::"'a::fs" - assumes "\name. y = Var name \ P" - and "\trm1 trm2. y = App trm1 trm2 \ P" - and "\name trm. \{atom name} \* c; y = Lam name trm\ \ P" - and "\assn1 trm1 assn2 trm2. - \((set (bn assn1)) \ (set (bn assn2))) \* c; y = Let1 assn1 trm1 assn2 trm2\ \ P" - and "\x1 x2 trm1 trm2. \{atom x1} \* c; y = Let2 x1 x2 trm1 trm2\ \ P" - shows "P" -apply(rule_tac y="y" in foo.exhaust(1)) -apply(rule assms(1)) -apply(assumption) -apply(rule assms(2)) -apply(assumption) -apply(subgoal_tac "\q. (q \ {atom name}) \* c \ supp (Lam name trm) \* q") -apply(erule exE) -apply(erule conjE) -apply(rule assms(3)) -apply(perm_simp) -apply(assumption) -apply(simp) -apply(drule supp_perm_eq[symmetric]) -apply(perm_simp) -apply(simp) -apply(rule at_set_avoiding2) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: foo.fresh fresh_star_def) -apply(subgoal_tac "\q. (q \ (set (bn assg1))) \* c \ supp ([bn assg1]lst. trm1) \* q") -apply(subgoal_tac "\q. (q \ (set (bn assg2))) \* c \ supp ([bn assg2]lst. trm2) \* q") -apply(erule exE)+ -apply(erule conjE)+ -apply(rule assms(4)) -apply(simp add: set_eqvt union_eqvt) -apply(simp add: tt1) -apply(simp add: fresh_star_union) -apply(rule conjI) -apply(assumption) -apply(rotate_tac 3) -apply(assumption) -apply(simp add: foo.eq_iff) -apply(drule supp_perm_eq[symmetric])+ -apply(simp add: tt1 uu1) -apply(auto)[1] -apply(rule at_set_avoiding2) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: Abs_fresh_star) -apply(rule at_set_avoiding2) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: Abs_fresh_star) -apply(case_tac "name1 = name2") -apply(subgoal_tac - "\q. (q \ {atom name1, atom name2}) \* c \ (supp (([[atom name1, atom name2]]lst. trm1), ([[atom name2]]lst. trm2))) \* q") -apply(erule exE)+ -apply(erule conjE)+ -apply(perm_simp) -apply(rule assms(5)) - -apply (simp add: fresh_star_def eqvts) -apply (simp only: supp_Pair) -apply (simp only: fresh_star_Un_elim) -apply (subst Let2_rename) -apply assumption -apply assumption -apply (rule refl) -apply(rule at_set_avoiding2) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply clarify -apply (simp add: fresh_star_def) -apply (simp add: fresh_def supp_Pair supp_Abs) -apply(subgoal_tac - "\q. (q \ {atom name1}) \* c \ (supp ((([[atom name1, atom name2]]lst. trm1)), (atom name2))) \* q") -prefer 2 -apply(rule at_set_avoiding2) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply (simp add: fresh_star_def) -apply (simp add: fresh_def supp_Pair supp_Abs supp_atom) -apply(erule exE)+ -apply(erule conjE)+ -apply(perm_simp) -apply(rule assms(5)) -apply assumption -apply clarify -apply (rule_tac x="name1" and y="name2" and ?t1.0="trm1" and ?t2.0="trm2" in Let2_rename2) -apply (simp_all add: fresh_star_Un_elim supp_Pair supp_Abs) -apply (simp add: fresh_star_def supp_atom) -done - -lemma strong_exhaust1: - fixes c::"'a::fs" - assumes "\name. y = Var name \ P" - and "\trm1 trm2. y = App trm1 trm2 \ P" - and "\name trm. \{atom name} \* c; y = Lam name trm\ \ P" - and "\assn1 trm1 assn2 trm2. - \((set (bn assn1)) \ (set (bn assn2))) \* c; y = Let1 assn1 trm1 assn2 trm2\ \ P" - and "\x1 x2 trm1 trm2. \{atom x1, atom x2} \* c; y = Let2 x1 x2 trm1 trm2\ \ P" - shows "P" - apply (rule strong_exhaust1_pre) - apply (erule assms) - apply (erule assms) - apply (erule assms) apply assumption - apply (erule assms) apply assumption -apply(case_tac "x1 = x2") -apply(subgoal_tac - "\q. (q \ {atom x1, atom x2}) \* c \ (supp (([[atom x1, atom x2]]lst. trm1), ([[atom x2]]lst. trm2))) \* q") -apply(erule exE)+ -apply(erule conjE)+ -apply(perm_simp) -apply(rule assms(5)) -apply assumption -apply simp -apply (rule Let2_rename) -apply (simp only: supp_Pair) -apply (simp only: fresh_star_Un_elim) -apply (simp only: supp_Pair) -apply (simp only: fresh_star_Un_elim) -apply(rule at_set_avoiding2) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply clarify -apply (simp add: fresh_star_def) -apply (simp add: fresh_def supp_Pair supp_Abs) - - apply(subgoal_tac - "\q. (q \ {atom x2}) \* c \ supp (([[atom x2]]lst. trm2), ([[atom x1, atom x2]]lst. trm1), (atom x1)) \* q") - apply(erule exE)+ - apply(erule conjE)+ - apply(rule assms(5)) -apply(perm_simp) -apply(simp (no_asm) add: fresh_star_insert) -apply(rule conjI) -apply (simp add: fresh_star_def) -apply(rotate_tac 2) -apply(simp add: fresh_star_def) -apply(simp) -apply (rule Let2_rename3) -apply (simp add: supp_Pair fresh_star_union) -apply (simp add: supp_Pair fresh_star_union) -apply (simp add: supp_Pair fresh_star_union) -apply clarify -apply (simp add: fresh_star_def supp_atom) -apply(rule at_set_avoiding2) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: finite_supp) -apply(simp add: fresh_star_def) -apply (simp add: fresh_def supp_Pair supp_Abs supp_atom) -done - -lemma strong_induct: - fixes c :: "'a :: fs" - and assg :: assg and trm :: trm - assumes a0: "\name c. P1 c (Var name)" - and a1: "\trm1 trm2 c. \\d. P1 d trm1; \d. P1 d trm2\ \ P1 c (App trm1 trm2)" - and a2: "\name trm c. (\d. P1 d trm) \ P1 c (Lam name trm)" - and a3: "\assg1 trm1 assg2 trm2 c. \((set (bn assg1)) \ (set (bn assg2))) \* c; \d. P2 c assg1; \d. P1 d trm1; \d. P2 d assg2; \d. P1 d trm2\ \ P1 c (Let1 assg1 trm1 assg2 trm2)" - and a4: "\name1 name2 trm1 trm2 c. \{atom name1, atom name2} \* c; \d. P1 d trm1; \d. P1 d trm2\ \ P1 c (Let2 name1 name2 trm1 trm2)" - and a5: "\c. P2 c As_Nil" - and a6: "\name1 name2 trm assg c. \\d. P1 d trm; \d. P2 d assg\ \ P2 c (As name1 name2 trm assg)" - shows "P1 c trm" "P2 c assg" - using assms - apply(induction_schema) - apply(rule_tac y="trm" and c="c" in strong_exhaust1) -apply(simp_all)[5] -apply(rule_tac y="assg" in foo.exhaust(2)) -apply(simp_all)[2] -apply(relation "measure (sum_case (size o snd) (\y. size (snd y)))") -apply(simp_all add: foo.size) -done - - text {* tests by cu *} -thm at_set_avoiding2 supp_perm_eq at_set_avoiding - -lemma abs_rename_set: - fixes x::"'a::fs" - assumes "b' \ x" "sort_of b = sort_of b'" - shows "\y. [{b}]set. x = [{b'}]set. y" -apply(rule_tac x="(b \ b') \ x" in exI) -apply(subst Abs_swap1[where a="b" and b="b'"]) -apply(simp) -using assms -apply(simp add: fresh_def) -apply(perm_simp) -using assms -apply(simp) -done - -lemma abs_rename_set1: - fixes x::"'a::fs" - assumes "(p \ b) \ x" - shows "\y. [{b}]set. x = [{p \ b}]set. y" -apply(rule_tac x="(b \ (p \ b)) \ x" in exI) -apply(subst Abs_swap1[where a="b" and b="p \ b"]) -apply(simp) -using assms -apply(simp add: fresh_def) -apply(perm_simp) -apply(simp) -done - lemma yy: - assumes "finite bs" - and "bs \ p \ bs = {}" + assumes a: "p \ bs \ bs = {}" + and b: "finite bs" shows "\q. q \ bs = p \ bs \ supp q \ bs \ (p \ bs)" -using assms +using b a apply(induct) apply(simp add: permute_set_eq) apply(rule_tac x="0" in exI) apply(simp add: supp_perm) apply(perm_simp) apply(drule meta_mp) -apply(auto)[1] +apply(auto simp add: fresh_star_def fresh_finite_insert)[1] apply(erule exE) apply(simp) apply(case_tac "q \ x = p \ x") @@ -353,36 +92,35 @@ apply(simp) apply(rule conjI) apply(simp add: supp_swap) -defer +apply(simp add: supp_perm) +apply(auto)[1] apply(auto)[1] apply(erule conjE)+ apply(drule sym) -apply(auto)[1] apply(simp add: mem_permute_iff) apply(simp add: mem_permute_iff) -apply(auto)[1] -apply(simp add: supp_perm) -apply(auto)[1] done -lemma zz: - fixes bs::"atom list" - assumes "set bs \ (set (p \ bs)) = {}" - shows "\q. q \ bs = p \ bs \ supp q \ (set bs) \ (set (p \ bs))" -sorry +lemma yy1: + assumes "(p \ bs) \* bs" "finite bs" + shows "p \ bs \ bs = {}" +using assms +apply(auto simp add: fresh_star_def) +apply(simp add: fresh_def supp_finite_atom_set) +done -lemma abs_rename_set2: +lemma abs_rename_set: fixes x::"'a::fs" - assumes "(p \ bs) \* x" "bs \ p \ bs ={}" "finite bs" + assumes "(p \ bs) \* x" "(p \ bs) \* bs" "finite bs" shows "\y. [bs]set. x = [p \ bs]set. y" using yy assms apply - +apply(drule_tac x="p" in meta_spec) apply(drule_tac x="bs" in meta_spec) -apply(drule_tac x="p" in meta_spec) -apply(auto) +apply(auto simp add: yy1) apply(rule_tac x="q \ x" in exI) apply(subgoal_tac "(q \ ([bs]set. x)) = [bs]set. x") -apply(simp add: permute_Abs) +apply(simp) apply(rule supp_perm_eq) apply(rule fresh_star_supp_conv) apply(simp add: fresh_star_def) @@ -390,29 +128,73 @@ apply(auto) done +lemma zz0: + assumes "p \ bs = q \ bs" + and a: "a \ set bs" + shows "p \ a = q \ a" +using assms +by (induct bs) (auto) + +lemma zz: + fixes bs::"atom list" + assumes "set bs \ (p \ (set bs)) = {}" + shows "\q. q \ bs = p \ bs \ supp q \ (set bs) \ (p \ (set bs))" +using assms +apply(induct bs) +apply(simp add: permute_set_eq) +apply(rule_tac x="0" in exI) +apply(simp add: supp_perm) +apply(drule meta_mp) +apply(perm_simp) +apply(auto)[1] +apply(erule exE) +apply(simp) +apply(case_tac "a \ set bs") +apply(rule_tac x="q" in exI) +apply(perm_simp) +apply(auto dest: zz0)[1] +apply(subgoal_tac "q \ a \ p \ bs") +apply(subgoal_tac "p \ a \ p \ bs") +apply(rule_tac x="((q \ a) \ (p \ a)) + q" in exI) +apply(simp add: swap_fresh_fresh) +apply(rule subset_trans) +apply(rule supp_plus_perm) +apply(simp) +apply(rule conjI) +apply(simp add: supp_swap) +apply(simp add: supp_perm) +apply(perm_simp) +apply(auto simp add: fresh_def supp_of_atom_list)[1] +apply(perm_simp) +apply(auto)[1] +apply(simp add: fresh_permute_iff) +apply(simp add: fresh_def supp_of_atom_list) +apply(erule conjE)+ +apply(drule sym) +apply(drule sym) +apply(simp) +apply(simp add: fresh_permute_iff) +apply(auto simp add: fresh_def supp_of_atom_list)[1] +done + +lemma zz1: + assumes "(p \ (set bs)) \* bs" + shows "(set bs) \ (p \ (set bs)) = {}" +using assms +apply(auto simp add: fresh_star_def) +apply(simp add: fresh_def supp_of_atom_list) +done + lemma abs_rename_list: fixes x::"'a::fs" - assumes "b' \ x" "sort_of b = sort_of b'" - shows "\y. [[b]]lst. x = [[b']]lst. y" -apply(rule_tac x="(b \ b') \ x" in exI) -apply(subst Abs_swap2[where a="b" and b="b'"]) -apply(simp) -using assms -apply(simp add: fresh_def) -apply(perm_simp) -using assms -apply(simp) -done - -lemma abs_rename_list2: - fixes x::"'a::fs" - assumes "(set (p \ bs)) \* x" "(set bs) \ (set (p \ bs)) ={}" + assumes "(p \ (set bs)) \* x" "(p \ (set bs)) \* bs" shows "\y. [bs]lst. x = [p \ bs]lst. y" using zz assms apply - apply(drule_tac x="bs" in meta_spec) apply(drule_tac x="p" in meta_spec) -apply(auto simp add: set_eqvt) +apply(drule_tac zz1) +apply(auto) apply(rule_tac x="q \ x" in exI) apply(subgoal_tac "(q \ ([bs]lst. x)) = [bs]lst. x") apply(simp) @@ -423,72 +205,172 @@ apply(auto) done +lemma fresh_star_list: + shows "as \* (xs @ ys) \ as \* xs \ as \* ys" + and "as \* (x # xs) \ as \* x \ as \* xs" + and "as \* []" +by (auto simp add: fresh_star_def fresh_Nil fresh_Cons fresh_append) -lemma test3: + +lemma test6: fixes c::"'a::fs" - assumes a: "y = Let2 x1 x2 trm1 trm2" - and b: "\x1 x2 trm1 trm2. {atom x1, atom x2} \* c \ y = Let2 x1 x2 trm1 trm2 \ P" + assumes "\name. y = Var name \ P" + and "\trm1 trm2. y = App trm1 trm2 \ P" + and "\name trm. {atom name} \* c \ y = Lam name trm \ P" + and "\a1 trm1 a2 trm2. ((set (bn a1)) \ (set (bn a2))) \* c \ y = Let1 a1 trm1 a2 trm2 \ P" + and "\x1 x2 trm1 trm2. {atom x1, atom x2} \* c \ y = Let2 x1 x2 trm1 trm2 \ P" shows "P" -using b[simplified a] -apply - -apply(subgoal_tac "\q::perm. (q \ {atom x1, atom x2}) \* (c, x1, x2, trm1, trm2)") +apply(rule_tac y="y" in foo.exhaust(1)) +apply(rule assms(1)) +apply(rule exI)+ +apply(assumption) +apply(rule assms(2)) +apply(rule exI)+ +apply(assumption) +apply(subgoal_tac "\p. (p \ {atom name}) \* (c, [atom name], trm)") apply(erule exE) +apply(rule assms(3)) +apply(insert abs_rename_list)[1] +apply(drule_tac x="p" in meta_spec) +apply(drule_tac x="[atom name]" in meta_spec) +apply(drule_tac x="trm" in meta_spec) +apply(simp only: fresh_star_prod set.simps) +apply(drule meta_mp) +apply(rule TrueI) +apply(drule meta_mp) +apply(rule TrueI) +apply(erule exE) +apply(rule exI)+ apply(perm_simp) -apply(simp only: foo.eq_iff) -apply(drule_tac x="q \ x1" in spec) -apply(drule_tac x="q \ x2" in spec) -apply(simp) -apply(erule mp) +apply(erule conjE)+ apply(rule conjI) -apply(simp add: fresh_star_prod) +apply(assumption) +apply(simp add: foo.eq_iff) +defer +apply(subgoal_tac "\p. (p \ ((set (bn assg1)) \ (set (bn assg2)))) \* (c, bn assg1, bn assg2, trm1, trm2)") +apply(erule exE) +apply(rule assms(4)) +apply(simp only:) +apply(simp only: foo.eq_iff) +apply(insert abs_rename_list)[1] +apply(drule_tac x="p" in meta_spec) +apply(drule_tac x="bn assg1" in meta_spec) +apply(drule_tac x="trm1" in meta_spec) +apply(insert abs_rename_list)[1] +apply(drule_tac x="p" in meta_spec) +apply(drule_tac x="bn assg2" in meta_spec) +apply(drule_tac x="trm2" in meta_spec) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps fresh_star_union) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps fresh_star_union) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps fresh_star_union) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps fresh_star_union) +apply(erule exE)+ +apply(rule exI)+ +apply(perm_simp add: tt1) +apply(rule conjI) +apply(simp add: fresh_star_prod fresh_star_union) +apply(erule conjE)+ +apply(rule conjI) +apply(assumption) +apply(rotate_tac 4) +apply(assumption) +apply(rule conjI) +apply(assumption) apply(rule conjI) -apply(simp add: atom_eqvt[symmetric]) -apply(subst (2) permute_list.simps(1)[where p="q", symmetric]) -apply(subst permute_list.simps(2)[symmetric]) -apply(subst permute_list.simps(2)[symmetric]) -apply(rule abs_rename_list2) -apply(simp add: atom_eqvt fresh_star_prod) -apply(simp add: atom_eqvt fresh_star_prod fresh_star_def fresh_Pair fresh_def supp_Pair supp_at_base) -apply(simp add: atom_eqvt[symmetric]) -apply(subst (2) permute_list.simps(1)[where p="q", symmetric]) -apply(subst permute_list.simps(2)[symmetric]) -apply(rule abs_rename_list2) -apply(simp add: atom_eqvt fresh_star_prod) -apply(simp add: atom_eqvt fresh_star_prod fresh_star_def fresh_Pair fresh_def supp_Pair supp_at_base) -apply(simp add: atom_eqvt[symmetric]) -apply(simp add: atom_eqvt fresh_star_prod fresh_star_def fresh_Pair fresh_def supp_Pair supp_at_base) +apply(rule uu1) +apply(rule conjI) +apply(assumption) +apply(rule uu1) +defer +apply(subgoal_tac "\p. (p \ ({atom name1} \ {atom name2})) \* (c, atom name1, atom name2, trm1, trm2)") +apply(erule exE) +apply(rule assms(5)) +apply(simp only:) +apply(simp only: foo.eq_iff) +apply(insert abs_rename_list)[1] +apply(drule_tac x="p" in meta_spec) +apply(drule_tac x="[atom name1] @ [atom name2]" in meta_spec) +apply(drule_tac x="trm1" in meta_spec) +apply(insert abs_rename_list)[1] +apply(drule_tac x="p" in meta_spec) +apply(drule_tac x="[atom name2]" in meta_spec) +apply(drule_tac x="trm2" in meta_spec) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps set_append fresh_star_union, simp) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps fresh_star_union) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps fresh_star_union set_append fresh_star_list, simp) +apply(drule meta_mp) +apply(simp only: union_eqvt fresh_star_prod set.simps fresh_star_union set_append fresh_star_list, simp) +apply(erule exE)+ +apply(rule exI)+ +apply(perm_simp add: tt1) +apply(rule conjI) +prefer 2 +apply(rule conjI) +apply(assumption) +apply(assumption) +apply(simp add: fresh_star_prod) +apply(simp add: fresh_star_def) sorry -lemma test4: - fixes c::"'a::fs" - assumes a: "y = Let2 x1 x2 trm1 trm2" - and b: "\x1 x2 trm1 trm2. \{atom x1, atom x2} \* c; y = Let2 x1 x2 trm1 trm2\ \ P" - shows "P" -apply(rule test3) -apply(rule a) -using b -apply(auto) -done + lemma test5: fixes c::"'a::fs" assumes "\name. y = Var name \ P" and "\trm1 trm2. y = App trm1 trm2 \ P" and "\name trm. \{atom name} \* c; y = Lam name trm\ \ P" - and "\assn1 trm1 assn2 trm2. - \((set (bn assn1)) \ (set (bn assn2))) \* c; y = Let1 assn1 trm1 assn2 trm2\ \ P" + and "\a1 trm1 a2 trm2. \((set (bn a1)) \ (set (bn a2))) \* c; y = Let1 a1 trm1 a2 trm2\ \ P" and "\x1 x2 trm1 trm2. \{atom x1, atom x2} \* c; y = Let2 x1 x2 trm1 trm2\ \ P" shows "P" -apply(rule_tac y="y" in foo.exhaust(1)) -apply (erule assms) -apply (erule assms) -prefer 3 -apply(erule test4[where c="c"]) -apply (rule assms(5)) -apply assumption -apply(simp) -oops +apply(rule_tac y="y" in test6) +apply(erule exE)+ +apply(rule assms(1)) +apply(assumption) +apply(erule exE)+ +apply(rule assms(2)) +apply(assumption) +apply(erule exE)+ +apply(rule assms(3)) +apply(auto)[2] +apply(erule exE)+ +apply(rule assms(4)) +apply(auto)[2] +apply(erule exE)+ +apply(rule assms(5)) +apply(auto)[2] +done + +lemma strong_induct: + fixes c :: "'a :: fs" + and assg :: assg and trm :: trm + assumes a0: "\name c. P1 c (Var name)" + and a1: "\trm1 trm2 c. \\d. P1 d trm1; \d. P1 d trm2\ \ P1 c (App trm1 trm2)" + and a2: "\name trm c. (\d. P1 d trm) \ P1 c (Lam name trm)" + and a3: "\a1 t1 a2 t2 c. + \((set (bn a1)) \ (set (bn a2))) \* c; \d. P2 c a1; \d. P1 d t1; \d. P2 d a2; \d. P1 d t2\ + \ P1 c (Let1 a1 t1 a2 t2)" + and a4: "\n1 n2 t1 t2 c. + \{atom n1, atom n2} \* c; \d. P1 d t1; \d. P1 d t2\ \ P1 c (Let2 n1 n2 t1 t2)" + and a5: "\c. P2 c As_Nil" + and a6: "\name1 name2 trm assg c. \\d. P1 d trm; \d. P2 d assg\ \ P2 c (As name1 name2 trm assg)" + shows "P1 c trm" "P2 c assg" + using assms + apply(induction_schema) + apply(rule_tac y="trm" and c="c" in test5) + apply(simp_all)[5] + apply(rule_tac y="assg" in foo.exhaust(2)) + apply(simp_all)[2] + apply(relation "measure (sum_case (size o snd) (size o snd))") + apply(simp_all add: foo.size) + done end