diff -r 425b4c406d80 -r b3d97424b130 Nominal/Nominal2_Base.thy --- a/Nominal/Nominal2_Base.thy Mon Jun 11 14:02:57 2012 +0100 +++ b/Nominal/Nominal2_Base.thy Tue Jun 12 01:23:52 2012 +0100 @@ -8,6 +8,7 @@ imports Main "~~/src/HOL/Library/Infinite_Set" "~~/src/HOL/Quotient_Examples/FSet" + "~~/src/HOL/Library/FinFun_Syntax" keywords "atom_decl" "equivariance" :: thy_decl uses ("nominal_basics.ML") @@ -654,6 +655,30 @@ by (lifting set_eqvt) +subsection {* Permutations for @{typ "'a \f 'b"} (FinFuns) *} + +instantiation finfun :: (pt, pt) pt +begin + +definition "p \ f = Abs_finfun (p \ (finfun_apply f))" + +lemma Rep_finfun_permute: + shows "p \ finfun_apply f \ finfun" +apply(simp add: permute_fun_comp) +apply(rule finfun_right_compose) +apply(rule finfun_left_compose) +apply(rule finfun_apply) +apply(simp) +done + +instance +apply(default) +apply(simp_all add: permute_finfun_def finfun_apply_inverse Rep_finfun_permute Abs_finfun_inverse) +done + +end + + subsection {* Permutations for @{typ char}, @{typ nat}, and @{typ int} *} instantiation char :: pt @@ -1190,6 +1215,25 @@ shows "p \ (map_fset f S) = map_fset (p \ f) (p \ S)" by (lifting map_eqvt) +subsubsection {* Equivariance for @{typ "'a \f 'b"} *} + +lemma permute_finfun_update[simp, eqvt]: + "(p \ (finfun_update f a b)) = finfun_update (p \ f) (p \ a) (p \ b)" +unfolding finfun_update_def +unfolding permute_finfun_def +apply(simp add: Abs_finfun_inverse fun_upd_finfun finfun_apply finfun_apply_inverse Rep_finfun_permute) +apply(simp add: fun_upd_def) +apply(perm_simp exclude: finfun_apply) +apply(rule refl) +done + +lemma permute_finfun_const[simp, eqvt]: + shows "(p \ (K$ b)) = (K$ (p \ b))" +unfolding finfun_const_def +unfolding permute_finfun_def +by (simp add: permute_finfun_def const_finfun finfun_apply_inverse Rep_finfun_permute Abs_finfun_inverse) + + section {* Supp, Freshness and Supports *} @@ -2187,6 +2231,37 @@ done +subsection {* Type @{typ "'a \f 'b"} is finitely supported *} + +lemma fresh_finfun_const: + shows "a \ (K$ b) \ a \ b" + by (simp add: fresh_def supp_def) + +lemma fresh_finfun_update: + shows "\a \ f; a \ b; a \ x\ \ a \ f(b $:= x)" + unfolding fresh_conv_MOST + unfolding permute_finfun_update + by (elim MOST_rev_mp) (simp) + +lemma supp_finfun_const: + "supp (K$ b) = supp(b)" + by (simp add: supp_def) + +lemma supp_finfun_update: + "supp (f(a $:= b)) \ supp(f, a, b)" +using fresh_finfun_update +by (auto simp add: fresh_def supp_Pair) + +instance finfun :: (fs, fs) fs + apply(default) + apply(induct_tac x rule: finfun_weak_induct) + apply(simp add: supp_finfun_const finite_supp) + apply(rule finite_subset) + apply(rule supp_finfun_update) + apply(simp add: supp_Pair finite_supp) + done + + section {* Freshness and Fresh-Star *} lemma fresh_Unit_elim: