diff -r 8779fb01d8b4 -r 09acd7e116e8 Nominal/Nominal2_Base.thy --- a/Nominal/Nominal2_Base.thy Thu Dec 29 15:56:54 2011 +0000 +++ b/Nominal/Nominal2_Base.thy Thu Dec 29 18:05:13 2011 +0000 @@ -1,5 +1,5 @@ (* Title: Nominal2_Base - Authors: Brian Huffman, Christian Urban + Authors: Christian Urban, Brian Huffman, Cezary Kaliszyk Basic definitions and lemma infrastructure for Nominal Isabelle. @@ -1110,6 +1110,22 @@ shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" by (induct S) (simp_all) +lemma inter_fset_eqvt [eqvt]: + shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" + apply(descending) + unfolding list_eq_def inter_list_def + apply(perm_simp) + apply(simp) + done + +lemma subset_fset_eqvt [eqvt]: + shows "(p \ (S |\| T)) = ((p \ S) |\| (p \ T))" + apply(descending) + unfolding sub_list_def + apply(perm_simp) + apply(simp) + done + lemma map_fset_eqvt [eqvt]: shows "p \ (map_fset f S) = map_fset (p \ f) (p \ S)" by (lifting map_eqvt)