# HG changeset patch # User Christian Urban # Date 1272638066 -3600 # Node ID b53e98bfb29850fe0935ebd7dfb984fefe22e8e2 # Parent 74d869595fed8e987e12cdf7f3ecb069d2136054 added lemmas establishing the support of finite sets of finitely supported elements diff -r 74d869595fed -r b53e98bfb298 Nominal-General/Nominal2_Base.thy --- a/Nominal-General/Nominal2_Base.thy Fri Apr 30 14:21:18 2010 +0100 +++ b/Nominal-General/Nominal2_Base.thy Fri Apr 30 15:34:26 2010 +0100 @@ -1067,10 +1067,17 @@ unfolding fresh_def supp_def unfolding MOST_iff_cofinite by simp +lemma supp_subset_fresh: + assumes a: "\a. a \ x \ a \ y" + shows "supp y \ supp x" + using a + unfolding fresh_def + by blast + lemma fresh_fun_app: assumes "a \ f" and "a \ x" shows "a \ f x" - using assms + using assms unfolding fresh_conv_MOST unfolding permute_fun_app_eq by (elim MOST_rev_mp, simp) @@ -1081,22 +1088,22 @@ unfolding fresh_def by auto +text {* support of equivariant functions *} + lemma supp_fun_eqvt: - assumes a: "\p. p \ f = f" + assumes a: "\p. p \ f = f" shows "supp f = {}" unfolding supp_def using a by simp - lemma fresh_fun_eqvt_app: - assumes a: "\p. p \ f = f" + assumes a: "\p. p \ f = f" shows "a \ x \ a \ f x" proof - from a have "supp f = {}" by (simp add: supp_fun_eqvt) then show "a \ x \ a \ f x" unfolding fresh_def - using supp_fun_app - by auto + using supp_fun_app by auto qed diff -r 74d869595fed -r b53e98bfb298 Nominal-General/Nominal2_Supp.thy --- a/Nominal-General/Nominal2_Supp.thy Fri Apr 30 14:21:18 2010 +0100 +++ b/Nominal-General/Nominal2_Supp.thy Fri Apr 30 15:34:26 2010 +0100 @@ -467,4 +467,85 @@ qed qed + +section {* Support of Finite Sets of Finitely Supported Elements *} + +lemma Union_fresh: + shows "a \ S \ a \ (\x \ S. supp x)" + unfolding Union_image_eq[symmetric] + apply(rule_tac f="\S. \ supp ` S" in fresh_fun_eqvt_app) + apply(perm_simp) + apply(rule refl) + apply(assumption) + done + +lemma Union_supports_set: + shows "(\x \ S. supp x) supports S" + apply(simp add: supports_def fresh_def[symmetric]) + apply(rule allI)+ + apply(rule impI) + apply(erule conjE) + apply(simp add: permute_set_eq) + apply(auto) + apply(subgoal_tac "(a \ b) \ xa = xa")(*A*) + apply(simp) + apply(rule swap_fresh_fresh) + apply(force) + apply(force) + apply(rule_tac x="x" in exI) + apply(simp) + apply(rule sym) + apply(rule swap_fresh_fresh) + apply(auto) + done + +lemma Union_of_fin_supp_sets: + fixes S::"('a::fs set)" + assumes fin: "finite S" + shows "finite (\x\S. supp x)" + using fin by (induct) (auto simp add: finite_supp) + +lemma Union_included_in_supp: + fixes S::"('a::fs set)" + assumes fin: "finite S" + shows "(\x\S. supp x) \ supp S" +proof - + have "(\x\S. supp x) = supp (\x\S. supp x)" + apply(rule supp_finite_atom_set[symmetric]) + apply(rule Union_of_fin_supp_sets) + apply(rule fin) + done + also have "\ \ supp S" + apply(rule supp_subset_fresh) + apply(simp add: Union_fresh) + done + finally show ?thesis . +qed + +lemma supp_of_fin_sets: + fixes S::"('a::fs set)" + assumes fin: "finite S" + shows "(supp S) = (\x\S. supp x)" +apply(rule subset_antisym) +apply(rule supp_is_subset) +apply(rule Union_supports_set) +apply(rule Union_of_fin_supp_sets[OF fin]) +apply(rule Union_included_in_supp[OF fin]) +done + +lemma supp_of_fin_union: + fixes S T::"('a::fs) set" + assumes fin1: "finite S" + and fin2: "finite T" + shows "supp (S \ T) = supp S \ supp T" + using fin1 fin2 + by (simp add: supp_of_fin_sets) + +lemma supp_of_fin_insert: + fixes S::"('a::fs) set" + assumes fin: "finite S" + shows "supp (insert x S) = supp x \ supp S" + using fin + by (simp add: supp_of_fin_sets) + end