diff -r 92c001d93225 -r 1d1772a89026 Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Tue Jan 18 06:55:18 2011 +0100 +++ b/Nominal/Nominal2_Abs.thy Tue Jan 18 11:02:57 2011 +0100 @@ -342,6 +342,41 @@ by (auto dest: disjoint_right_eq) qed +lemma alpha_abs_res_stronger1: + fixes x::"'a::fs" + assumes asm: "(as, x) \res (op =) supp p' (as', x')" + shows "\p. (as, x) \res (op =) supp p (as', x') \ supp p \ (supp x \ as) \ (supp x' \ as')" +proof - + from asm have 0: "(supp x - as) \* p'" by (auto simp only: alphas) + then have #: "p' \ (supp x - as) = (supp x - as)" + by (simp add: atom_set_perm_eq) + have "finite (supp x)" by (simp add: finite_supp) + then obtain p where *: "\b \ supp x. p \ b = p' \ b" and **: "supp p \ supp x \ p' \ supp x" + using set_renaming_perm by blast + from * have a: "p \ x = p' \ x" using supp_perm_perm_eq by auto + from 0 have 1: "(supp x - as) \* p" using * + by (auto simp add: fresh_star_def fresh_perm) + then have 2: "(supp x - as) \ supp p = {}" + by (auto simp add: fresh_star_def fresh_def) + have b: "supp x = (supp x - as) \ (supp x \ as)" by auto + have "supp p \ supp x \ p' \ supp x" using ** by simp + also have "\ = (supp x - as) \ (supp x \ as) \ (p' \ ((supp x - as) \ (supp x \ as)))" + using b by simp + also have "\ = (supp x - as) \ (supp x \ as) \ ((p' \ (supp x - as)) \ (p' \ (supp x \ as)))" + by (simp add: union_eqvt) + also have "\ = (supp x - as) \ (supp x \ as) \ (p' \ (supp x \ as))" + using # by auto + also have "\ = (supp x - as) \ (supp x \ as) \ (supp x' \ as')" using asm + by (simp add: supp_property_res) + finally have "supp p \ (supp x - as) \ (supp x \ as) \ (supp x' \ as')" . + then + have "supp p \ (supp x \ as) \ (supp x' \ as')" using 2 by auto + moreover + have "(as, x) \res (op =) supp p (as', x')" using asm 1 a by (simp add: alphas) + ultimately + show "\p. (as, x) \res (op =) supp p (as', x') \ supp p \ (supp x \ as) \ (supp x' \ as')" by blast +qed + section {* Quotient types *}