diff -r d5713db7e146 -r dd7490fdd998 Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Sun Dec 19 07:50:37 2010 +0000 +++ b/Nominal/Nominal2_Abs.thy Tue Dec 21 10:28:08 2010 +0000 @@ -547,7 +547,7 @@ fixes x::"'a::fs" assumes a: "(p \ bs) \* (bs, x)" and b: "finite bs" - shows "\y. [bs]set. x = [p \ bs]set. y" + shows "\q. [bs]set. x = [p \ bs]set. (q \ x) \ q \ bs = p \ bs" proof - from a b have "p \ bs \ bs = {}" using at_fresh_star_inter by (auto simp add: fresh_star_Pair) with set_renaming_perm @@ -560,16 +560,15 @@ 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 + finally have "[bs]set. x = [p \ bs]set. (q \ x)" by (simp add: *) + then show "\q. [bs]set. x = [p \ bs]set. (q \ x) \ q \ bs = p \ bs" using * by metis 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" + shows "\q. [bs]res. x = [p \ bs]res. (q \ x) \ q \ bs = p \ bs" proof - from a b have "p \ bs \ bs = {}" using at_fresh_star_inter by (simp add: fresh_star_Pair) with set_renaming_perm @@ -582,15 +581,14 @@ 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 + finally have "[bs]res. x = [p \ bs]res. (q \ x)" by (simp add: *) + then show "\q. [bs]res. x = [p \ bs]res. (q \ x) \ q \ bs = p \ bs" using * by metis qed lemma Abs_rename_lst: fixes x::"'a::fs" assumes a: "(p \ (set bs)) \* (bs, x)" - shows "\y. [bs]lst. x = [p \ bs]lst. y" + shows "\q. [bs]lst. x = [p \ bs]lst. (q \ x) \ q \ bs = p \ bs" proof - from a have "p \ (set bs) \ (set bs) = {}" using at_fresh_star_inter by (simp add: fresh_star_Pair fresh_star_set) @@ -604,12 +602,32 @@ 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 + finally have "[bs]lst. x = [p \ bs]lst. (q \ x)" by (simp add: *) + then show "\q. [bs]lst. x = [p \ bs]lst. (q \ x) \ q \ bs = p \ bs" using * by metis qed +text {* for deep recursive binders *} + +lemma Abs_rename_set': + fixes x::"'a::fs" + assumes a: "(p \ bs) \* (bs, x)" + and b: "finite bs" + shows "\q. [bs]set. x = [q \ bs]set. (q \ x) \ q \ bs = p \ bs" +using Abs_rename_set[OF a b] by metis + +lemma Abs_rename_res': + fixes x::"'a::fs" + assumes a: "(p \ bs) \* (bs, x)" + and b: "finite bs" + shows "\q. [bs]res. x = [q \ bs]res. (q \ x) \ q \ bs = p \ bs" +using Abs_rename_res[OF a b] by metis + +lemma Abs_rename_lst': + fixes x::"'a::fs" + assumes a: "(p \ (set bs)) \* (bs, x)" + shows "\q. [bs]lst. x = [q \ bs]lst. (q \ x) \ q \ bs = p \ bs" +using Abs_rename_lst[OF a] by metis section {* Infrastructure for building tuples of relations and functions *}