diff -r a3db645bbfda -r 9761840f2d5c Nominal-General/Nominal2_Supp.thy --- a/Nominal-General/Nominal2_Supp.thy Mon Apr 19 11:32:33 2010 +0200 +++ b/Nominal-General/Nominal2_Supp.thy Mon Apr 19 11:38:43 2010 +0200 @@ -141,20 +141,38 @@ using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \ supp c"] unfolding fresh_star_def fresh_def by blast +lemma at_set_avoiding2: + assumes "finite xs" + and "finite (supp c)" "finite (supp x)" + and "xs \* x" + shows "\p. (p \ xs) \* c \ supp x \* p" +using assms +apply(erule_tac c="(c, x)" in at_set_avoiding) +apply(simp add: supp_Pair) +apply(rule_tac x="p" in exI) +apply(simp add: fresh_star_prod) +apply(subgoal_tac "\a \ supp p. a \ x") +apply(auto simp add: fresh_star_def fresh_def supp_perm)[1] +apply(auto simp add: fresh_star_def fresh_def) +done + +lemma at_set_avoiding2_atom: + assumes "finite (supp c)" "finite (supp x)" + and b: "xa \ x" + shows "\p. (p \ xa) \ c \ supp x \* p" +proof - + have a: "{xa} \* x" unfolding fresh_star_def by (simp add: b) + obtain p where p1: "(p \ {xa}) \* c" and p2: "supp x \* p" + using at_set_avoiding2[of "{xa}" "c" "x"] assms a by blast + have c: "(p \ xa) \ c" using p1 + unfolding fresh_star_def Ball_def + by (erule_tac x="p \ xa" in allE) (simp add: eqvts) + hence "p \ xa \ c \ supp x \* p" using p2 by blast + then show ?thesis by blast +qed section {* The freshness lemma according to Andrew Pitts *} -lemma fresh_conv_MOST: - shows "a \ x \ (MOST b. (a \ b) \ x = x)" - unfolding fresh_def supp_def MOST_iff_cofinite by simp - -lemma fresh_apply: - assumes "a \ f" and "a \ x" - shows "a \ f x" - using assms unfolding fresh_conv_MOST - unfolding permute_fun_app_eq [where f=f] - by (elim MOST_rev_mp, simp) - lemma freshness_lemma: fixes h :: "'a::at \ 'b::pt" assumes a: "\a. atom a \ (h, h a)" @@ -173,7 +191,8 @@ next assume "a \ b" hence "atom a \ b" by (simp add: fresh_at_base) - with a3 have "atom a \ h b" by (rule fresh_apply) + with a3 have "atom a \ h b" + by (rule fresh_fun_app) with a2 have d1: "(atom b \ atom a) \ (h b) = (h b)" by (rule swap_fresh_fresh) from a1 a3 have d2: "(atom b \ atom a) \ h = h" @@ -358,7 +377,7 @@ using not_fresh_nat_of [unfolded fresh_def] by auto -section {* Support for sets of atoms *} +section {* Support for finite sets of atoms *} lemma supp_finite_atom_set: fixes S::"atom set" @@ -447,36 +466,4 @@ qed qed -section {* at_set_avoiding2 *} - -lemma at_set_avoiding2: - assumes "finite xs" - and "finite (supp c)" "finite (supp x)" - and "xs \* x" - shows "\p. (p \ xs) \* c \ supp x \* p" -using assms -apply(erule_tac c="(c, x)" in at_set_avoiding) -apply(simp add: supp_Pair) -apply(rule_tac x="p" in exI) -apply(simp add: fresh_star_prod) -apply(subgoal_tac "\a \ supp p. a \ x") -apply(auto simp add: fresh_star_def fresh_def supp_perm)[1] -apply(auto simp add: fresh_star_def fresh_def) -done - -lemma at_set_avoiding2_atom: - assumes "finite (supp c)" "finite (supp x)" - and b: "xa \ x" - shows "\p. (p \ xa) \ c \ supp x \* p" -proof - - have a: "{xa} \* x" unfolding fresh_star_def by (simp add: b) - obtain p where p1: "(p \ {xa}) \* c" and p2: "supp x \* p" - using at_set_avoiding2[of "{xa}" "c" "x"] assms a by blast - have c: "(p \ xa) \ c" using p1 - unfolding fresh_star_def Ball_def - by (erule_tac x="p \ xa" in allE) (simp add: eqvts) - hence "p \ xa \ c \ supp x \* p" using p2 by blast - then show ?thesis by blast -qed - end