diff -r dd3b9c046c7d -r 4da5c5c29009 Nominal/Abs.thy --- a/Nominal/Abs.thy Thu Jun 03 15:02:52 2010 +0200 +++ b/Nominal/Abs.thy Mon Jun 07 11:43:01 2010 +0200 @@ -51,6 +51,73 @@ by (case_tac [!] bs, case_tac [!] cs) (auto simp add: le_fun_def le_bool_def alphas) +(* equivariance *) +lemma alpha_gen_eqvt[eqvt]: + shows "(bs, x) \gen R f q (cs, y) \ (p \ bs, p \ x) \gen (p \ R) (p \ f) (p \ q) (p \ cs, p \ y)" + and "(bs, x) \res R f q (cs, y) \ (p \ bs, p \ x) \res (p \ R) (p \ f) (p \ q) (p \ cs, p \ y)" + and "(ds, x) \lst R f q (es, y) \ (p \ ds, p \ x) \lst (p \ R) (p \ f) (p \ q) (p \ es, p \ y)" + unfolding alphas + unfolding permute_eqvt[symmetric] + unfolding set_eqvt[symmetric] + unfolding permute_fun_app_eq[symmetric] + unfolding Diff_eqvt[symmetric] + by (auto simp add: permute_bool_def fresh_star_permute_iff) + +(* equivalence *) +lemma alpha_gen_refl: + assumes a: "R x x" + shows "(bs, x) \gen R f 0 (bs, x)" + and "(bs, x) \res R f 0 (bs, x)" + and "(cs, x) \lst R f 0 (cs, x)" + using a + unfolding alphas + unfolding fresh_star_def + by (simp_all add: fresh_zero_perm) + +lemma alpha_gen_sym: + assumes a: "R (p \ x) y \ R (- p \ y) x" + shows "(bs, x) \gen R f p (cs, y) \ (cs, y) \gen R f (- p) (bs, x)" + and "(bs, x) \res R f p (cs, y) \ (cs, y) \res R f (- p) (bs, x)" + and "(ds, x) \lst R f p (es, y) \ (es, y) \lst R f (- p) (ds, x)" + unfolding alphas fresh_star_def + using a + by (auto simp add: fresh_minus_perm) + +lemma alpha_gen_sym_eqvt: + assumes a: "R (p \ x) y \ R y (p \ x)" + and b: "p \ R = R" + shows "(bs, x) \gen R f p (cs, y) \ (cs, y) \gen R f (- p) (bs, x)" + and "(bs, x) \res R f p (cs, y) \ (cs, y) \res R f (- p) (bs, x)" + and "(ds, x) \lst R f p (es, y) \ (es, y) \lst R f (- p) (ds, x)" +proof - + { assume "R (p \ x) y" + then have "R y (p \ x)" using a by simp + then have "R (- p \ y) x" + apply(rule_tac p="p" in permute_boolE) + apply(perm_simp add: permute_minus_cancel b) + apply(assumption) + done + } + then show "(bs, x) \gen R f p (cs, y) \ (cs, y) \gen R f (- p) (bs, x)" + and "(bs, x) \res R f p (cs, y) \ (cs, y) \res R f (- p) (bs, x)" + and "(ds, x) \lst R f p (es, y) \ (es, y) \lst R f (- p) (ds, x)" + unfolding alphas fresh_star_def + by (auto simp add: fresh_minus_perm) +qed + +lemma alpha_gen_trans: + assumes a: "\R (p \ x) y; R (q \ y) z\ \ R ((q + p) \ x) z" + shows "\(bs, x) \gen R f p (cs, y); (cs, y) \gen R f q (ds, z)\ \ (bs, x) \gen R f (q + p) (ds, z)" + and "\(bs, x) \res R f p (cs, y); (cs, y) \res R f q (ds, z)\ \ (bs, x) \res R f (q + p) (ds, z)" + and "\(es, x) \lst R f p (gs, y); (gs, y) \lst R f q (hs, z)\ \ (es, x) \lst R f (q + p) (hs, z)" + using a + unfolding alphas fresh_star_def + by (simp_all add: fresh_plus_perm) + + + +section {* General Abstractions *} + fun alpha_abs where @@ -731,45 +798,7 @@ lemmas alphas_compose_trans2 = alpha_gen_compose_trans2 alpha_res_compose_trans2 alpha_lst_compose_trans2 -lemma alpha_gen_refl: - assumes a: "R x x" - shows "(bs, x) \gen R f 0 (bs, x)" - and "(bs, x) \res R f 0 (bs, x)" - and "(cs, x) \lst R f 0 (cs, x)" - using a - unfolding alphas - unfolding fresh_star_def - by (simp_all add: fresh_zero_perm) -lemma alpha_gen_sym: - assumes a: "R (p \ x) y \ R (- p \ y) x" - shows "(bs, x) \gen R f p (cs, y) \ (cs, y) \gen R f (- p) (bs, x)" - and "(bs, x) \res R f p (cs, y) \ (cs, y) \res R f (- p) (bs, x)" - and "(ds, x) \lst R f p (es, y) \ (es, y) \lst R f (- p) (ds, x)" - unfolding alphas fresh_star_def - using a - by (auto simp add: fresh_minus_perm) - -lemma alpha_gen_trans: - assumes a: "\R (p \ x) y; R (q \ y) z\ \ R ((q + p) \ x) z" - shows "\(bs, x) \gen R f p (cs, y); (cs, y) \gen R f q (ds, z)\ \ (bs, x) \gen R f (q + p) (ds, z)" - and "\(bs, x) \res R f p (cs, y); (cs, y) \res R f q (ds, z)\ \ (bs, x) \res R f (q + p) (ds, z)" - and "\(es, x) \lst R f p (gs, y); (gs, y) \lst R f q (hs, z)\ \ (es, x) \lst R f (q + p) (hs, z)" - using a - unfolding alphas fresh_star_def - by (simp_all add: fresh_plus_perm) - - -lemma alpha_gen_eqvt[eqvt]: - shows "(bs, x) \gen R f q (cs, y) \ (p \ bs, p \ x) \gen (p \ R) (p \ f) (p \ q) (p \ cs, p \ y)" - and "(bs, x) \res R f q (cs, y) \ (p \ bs, p \ x) \res (p \ R) (p \ f) (p \ q) (p \ cs, p \ y)" - and "(ds, x) \lst R f q (es, y) \ (p \ ds, p \ x) \lst (p \ R) (p \ f) (p \ q) (p \ es, p \ y)" - unfolding alphas - unfolding permute_eqvt[symmetric] - unfolding set_eqvt[symmetric] - unfolding permute_fun_app_eq[symmetric] - unfolding Diff_eqvt[symmetric] - by (auto simp add: permute_bool_def fresh_star_permute_iff) lemma alpha_gen_simpler: assumes fv_rsp: "\x y. R y x \ f x = f y"