diff -r b136721eedb2 -r d6fe94028a5d Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Tue Dec 07 14:27:39 2010 +0000 +++ b/Nominal/Nominal2_Abs.thy Tue Dec 07 19:16:09 2010 +0000 @@ -540,6 +540,77 @@ by(auto simp add: Abs_fresh_iff) +subsection {* Renaming of bodies of abstractions *} + + +lemma Abs_rename_set: + fixes x::"'a::fs" + assumes a: "(p \ bs) \* (bs, x)" + and b: "finite bs" + shows "\y. [bs]set. x = [p \ bs]set. y" +proof - + from a b have "p \ bs \ bs = {}" using at_fresh_star_inter by (auto simp add: fresh_star_Pair) + with set_renaming_perm + obtain q where *: "q \ bs = p \ bs" and **: "supp q \ bs \ (p \ bs)" using b by metis + have "[bs]set. x = q \ ([bs]set. x)" + apply(rule perm_supp_eq[symmetric]) + using a ** + unfolding fresh_star_Pair + unfolding Abs_fresh_star_iff + unfolding fresh_star_def + by auto + also have "\ = [q \ bs]set. (q \ x)" by simp + also have "\ = [p \ bs]set. (q \ x)" using * by simp + finally have "[bs]set. x = [p \ bs]set. (q \ x)" . + then show "\y. [bs]set. x = [p \ bs]set. y" by blast +qed + +lemma Abs_rename_res: + fixes x::"'a::fs" + assumes a: "(p \ bs) \* (bs, x)" + and b: "finite bs" + shows "\y. [bs]res. x = [p \ bs]res. y" +proof - + from a b have "p \ bs \ bs = {}" using at_fresh_star_inter by (simp add: fresh_star_Pair) + with set_renaming_perm + obtain q where *: "q \ bs = p \ bs" and **: "supp q \ bs \ (p \ bs)" using b by metis + have "[bs]res. x = q \ ([bs]res. x)" + apply(rule perm_supp_eq[symmetric]) + using a ** + unfolding fresh_star_Pair + unfolding Abs_fresh_star_iff + unfolding fresh_star_def + by auto + also have "\ = [q \ bs]res. (q \ x)" by simp + also have "\ = [p \ bs]res. (q \ x)" using * by simp + finally have "[bs]res. x = [p \ bs]res. (q \ x)" . + then show "\y. [bs]res. x = [p \ bs]res. y" by blast +qed + +lemma Abs_rename_list: + fixes x::"'a::fs" + assumes a: "(p \ (set bs)) \* (bs, x)" + shows "\y. [bs]lst. x = [p \ bs]lst. y" +proof - + from a have "p \ (set bs) \ (set bs) = {}" using at_fresh_star_inter + by (simp add: fresh_star_Pair fresh_star_set) + with list_renaming_perm + obtain q where *: "q \ bs = p \ bs" and **: "supp q \ set bs \ (p \ set bs)" by metis + have "[bs]lst. x = q \ ([bs]lst. x)" + apply(rule perm_supp_eq[symmetric]) + using a ** + unfolding fresh_star_Pair + unfolding Abs_fresh_star_iff + unfolding fresh_star_def + by auto + also have "\ = [q \ bs]lst. (q \ x)" by simp + also have "\ = [p \ bs]lst. (q \ x)" using * by simp + finally have "[bs]lst. x = [p \ bs]lst. (q \ x)" . + then show "\y. [bs]lst. x = [p \ bs]lst. y" by blast +qed + + + section {* Infrastructure for building tuples of relations and functions *} fun