# HG changeset patch # User Christian Urban # Date 1291749369 0 # Node ID d6fe94028a5dd93270ffa380c1cdfc6003fbaa13 # Parent b136721eedb23ab958038d9c9bc83a59108babe6 moved general theorems into the libraries diff -r b136721eedb2 -r d6fe94028a5d Nominal/Ex/Foo2.thy --- a/Nominal/Ex/Foo2.thy Tue Dec 07 14:27:39 2010 +0000 +++ b/Nominal/Ex/Foo2.thy Tue Dec 07 19:16:09 2010 +0000 @@ -49,179 +49,9 @@ tests by cu *} -lemma set_renaming_perm: - assumes a: "p \ bs \ bs = {}" - and b: "finite bs" - shows "\q. q \ bs = p \ bs \ supp q \ bs \ (p \ bs)" -using b a -proof (induct) - case empty - have "0 \ {} = p \ {} \ supp (0::perm) \ {} \ p \ {}" - by (simp add: permute_set_eq supp_perm) - then show "\q. q \ {} = p \ {} \ supp q \ {} \ p \ {}" by blast -next - case (insert a bs) - then have " \q. q \ bs = p \ bs \ supp q \ bs \ p \ bs" - by (perm_simp) (auto) - then obtain q where *: "q \ bs = p \ bs" and **: "supp q \ bs \ p \ bs" by blast - { assume 1: "q \ a = p \ a" - have "q \ insert a bs = p \ insert a bs" using 1 * by (simp add: insert_eqvt) - moreover - have "supp q \ insert a bs \ p \ insert a bs" - using ** by (auto simp add: insert_eqvt) - ultimately - have "\q. q \ insert a bs = p \ insert a bs \ supp q \ insert a bs \ p \ insert a bs" by blast - } - moreover - { assume 2: "q \ a \ p \ a" - def q' \ "((q \ a) \ (p \ a)) + q" - { have "(q \ a) \ (p \ bs)" using `a \ bs` *[symmetric] by (simp add: mem_permute_iff) - moreover - have "(p \ a) \ (p \ bs)" using `a \ bs` by (simp add: mem_permute_iff) - ultimately - have "q' \ insert a bs = p \ insert a bs" using 2 * unfolding q'_def - by (simp add: insert_eqvt swap_set_not_in) - } - moreover - { have "{q \ a, p \ a} \ insert a bs \ p \ insert a bs" - using ** `a \ bs` `p \ insert a bs \ insert a bs = {}` - by (auto simp add: supp_perm insert_eqvt) - then have "supp (q \ a \ p \ a) \ insert a bs \ p \ insert a bs" by (simp add: supp_swap) - moreover - have "supp q \ insert a bs \ p \ insert a bs" - using ** by (auto simp add: insert_eqvt) - ultimately - have "supp q' \ insert a bs \ p \ insert a bs" - unfolding q'_def using supp_plus_perm by blast - } - ultimately - have "\q. q \ insert a bs = p \ insert a bs \ supp q \ insert a bs \ p \ insert a bs" by blast - } - ultimately show "\q. q \ insert a bs = p \ insert a bs \ supp q \ insert a bs \ p \ insert a bs" - by blast -qed - - -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 +text {* *} -lemma list_renaming_perm: - fixes bs::"atom list" - assumes a: "(p \ (set bs)) \ (set bs) = {}" - shows "\q. q \ bs = p \ bs \ supp q \ (set bs) \ (p \ (set bs))" -using a -proof (induct bs) - case Nil - have "0 \ [] = p \ [] \ supp (0::perm) \ set [] \ p \ set []" - by (simp add: permute_set_eq supp_perm) - then show "\q. q \ [] = p \ [] \ supp q \ set [] \ p \ (set [])" by blast -next - case (Cons a bs) - then have " \q. q \ bs = p \ bs \ supp q \ set bs \ p \ (set bs)" - by (perm_simp) (auto) - then obtain q where *: "q \ bs = p \ bs" and **: "supp q \ set bs \ p \ (set bs)" by blast - { assume 1: "a \ set bs" - have "q \ a = p \ a" using * 1 by (induct bs) (auto) - then have "q \ (a # bs) = p \ (a # bs)" using * by simp - moreover - have "supp q \ set (a # bs) \ p \ (set (a # bs))" using ** by (auto simp add: insert_eqvt) - ultimately - have "\q. q \ (a # bs) = p \ (a # bs) \ supp q \ set (a # bs) \ p \ (set (a # bs))" by blast - } - moreover - { assume 2: "a \ set bs" - def q' \ "((q \ a) \ (p \ a)) + q" - { have "(q \ a) \ (p \ bs)" using `a \ set bs` *[symmetric] - by (simp add: fresh_permute_iff) (simp add: fresh_def supp_of_atom_list) - moreover - have "(p \ a) \ (p \ bs)" using `a \ set bs` - by (simp add: fresh_permute_iff) (simp add: fresh_def supp_of_atom_list) - ultimately - have "q' \ (a # bs) = p \ (a # bs)" using 2 * unfolding q'_def - by (simp add: swap_fresh_fresh) - } - moreover - { have "{q \ a, p \ a} \ set (a # bs) \ p \ (set (a # bs))" - using ** `a \ set bs` `p \ (set (a # bs)) \ set (a # bs) = {}` - by (auto simp add: supp_perm insert_eqvt) - then have "supp (q \ a \ p \ a) \ set (a # bs) \ p \ set (a # bs)" by (simp add: supp_swap) - moreover - have "supp q \ set (a # bs) \ p \ (set (a # bs))" - using ** by (auto simp add: insert_eqvt) - ultimately - have "supp q' \ set (a # bs) \ p \ (set (a # bs))" - unfolding q'_def using supp_plus_perm by blast - } - ultimately - have "\q. q \ (a # bs) = p \ (a # bs) \ supp q \ set (a # bs) \ p \ (set (a # bs))" by blast - } - ultimately show "\q. q \ (a # bs) = p \ (a # bs) \ supp q \ set (a # bs) \ p \ (set (a # bs))" - 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 (* thm at_set_avoiding1[THEN exE] 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 diff -r b136721eedb2 -r d6fe94028a5d Nominal/Nominal2_Base.thy --- a/Nominal/Nominal2_Base.thy Tue Dec 07 14:27:39 2010 +0000 +++ b/Nominal/Nominal2_Base.thy Tue Dec 07 19:16:09 2010 +0000 @@ -1699,6 +1699,117 @@ then show "\p. (p \ a) \ c \ supp x \* p" by blast qed +section {* Renaming permutations *} + + +lemma set_renaming_perm: + assumes a: "p \ bs \ bs = {}" + and b: "finite bs" + shows "\q. q \ bs = p \ bs \ supp q \ bs \ (p \ bs)" +using b a +proof (induct) + case empty + have "0 \ {} = p \ {} \ supp (0::perm) \ {} \ p \ {}" + by (simp add: permute_set_eq supp_perm) + then show "\q. q \ {} = p \ {} \ supp q \ {} \ p \ {}" by blast +next + case (insert a bs) + then have " \q. q \ bs = p \ bs \ supp q \ bs \ p \ bs" + by (simp add: insert_eqvt) + then obtain q where *: "q \ bs = p \ bs" and **: "supp q \ bs \ p \ bs" by blast + { assume 1: "q \ a = p \ a" + have "q \ insert a bs = p \ insert a bs" using 1 * by (simp add: insert_eqvt) + moreover + have "supp q \ insert a bs \ p \ insert a bs" + using ** by (auto simp add: insert_eqvt) + ultimately + have "\q. q \ insert a bs = p \ insert a bs \ supp q \ insert a bs \ p \ insert a bs" by blast + } + moreover + { assume 2: "q \ a \ p \ a" + def q' \ "((q \ a) \ (p \ a)) + q" + { have "(q \ a) \ (p \ bs)" using `a \ bs` *[symmetric] by (simp add: mem_permute_iff) + moreover + have "(p \ a) \ (p \ bs)" using `a \ bs` by (simp add: mem_permute_iff) + ultimately + have "q' \ insert a bs = p \ insert a bs" using 2 * unfolding q'_def + by (simp add: insert_eqvt swap_set_not_in) + } + moreover + { have "{q \ a, p \ a} \ insert a bs \ p \ insert a bs" + using ** `a \ bs` `p \ insert a bs \ insert a bs = {}` + by (auto simp add: supp_perm insert_eqvt) + then have "supp (q \ a \ p \ a) \ insert a bs \ p \ insert a bs" by (simp add: supp_swap) + moreover + have "supp q \ insert a bs \ p \ insert a bs" + using ** by (auto simp add: insert_eqvt) + ultimately + have "supp q' \ insert a bs \ p \ insert a bs" + unfolding q'_def using supp_plus_perm by blast + } + ultimately + have "\q. q \ insert a bs = p \ insert a bs \ supp q \ insert a bs \ p \ insert a bs" by blast + } + ultimately show "\q. q \ insert a bs = p \ insert a bs \ supp q \ insert a bs \ p \ insert a bs" + by blast +qed + + +lemma list_renaming_perm: + fixes bs::"atom list" + assumes a: "(p \ (set bs)) \ (set bs) = {}" + shows "\q. q \ bs = p \ bs \ supp q \ (set bs) \ (p \ (set bs))" +using a +proof (induct bs) + case Nil + have "0 \ [] = p \ [] \ supp (0::perm) \ set [] \ p \ set []" + by (simp add: permute_set_eq supp_perm) + then show "\q. q \ [] = p \ [] \ supp q \ set [] \ p \ (set [])" by blast +next + case (Cons a bs) + then have " \q. q \ bs = p \ bs \ supp q \ set bs \ p \ (set bs)" + by (simp add: insert_eqvt) + then obtain q where *: "q \ bs = p \ bs" and **: "supp q \ set bs \ p \ (set bs)" by blast + { assume 1: "a \ set bs" + have "q \ a = p \ a" using * 1 by (induct bs) (auto) + then have "q \ (a # bs) = p \ (a # bs)" using * by simp + moreover + have "supp q \ set (a # bs) \ p \ (set (a # bs))" using ** by (auto simp add: insert_eqvt) + ultimately + have "\q. q \ (a # bs) = p \ (a # bs) \ supp q \ set (a # bs) \ p \ (set (a # bs))" by blast + } + moreover + { assume 2: "a \ set bs" + def q' \ "((q \ a) \ (p \ a)) + q" + { have "(q \ a) \ (p \ bs)" using `a \ set bs` *[symmetric] + by (simp add: fresh_permute_iff) (simp add: fresh_def supp_of_atom_list) + moreover + have "(p \ a) \ (p \ bs)" using `a \ set bs` + by (simp add: fresh_permute_iff) (simp add: fresh_def supp_of_atom_list) + ultimately + have "q' \ (a # bs) = p \ (a # bs)" using 2 * unfolding q'_def + by (simp add: swap_fresh_fresh) + } + moreover + { have "{q \ a, p \ a} \ set (a # bs) \ p \ (set (a # bs))" + using ** `a \ set bs` `p \ (set (a # bs)) \ set (a # bs) = {}` + by (auto simp add: supp_perm insert_eqvt) + then have "supp (q \ a \ p \ a) \ set (a # bs) \ p \ set (a # bs)" by (simp add: supp_swap) + moreover + have "supp q \ set (a # bs) \ p \ (set (a # bs))" + using ** by (auto simp add: insert_eqvt) + ultimately + have "supp q' \ set (a # bs) \ p \ (set (a # bs))" + unfolding q'_def using supp_plus_perm by blast + } + ultimately + have "\q. q \ (a # bs) = p \ (a # bs) \ supp q \ set (a # bs) \ p \ (set (a # bs))" by blast + } + ultimately show "\q. q \ (a # bs) = p \ (a # bs) \ supp q \ set (a # bs) \ p \ (set (a # bs))" + by blast +qed + + section {* Concrete Atoms Types *}