diff -r f8d660de0cf7 -r 1ae3c9b2d557 Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Thu Jun 09 15:34:51 2011 +0900 +++ b/Nominal/Nominal2_Abs.thy Fri Jun 10 15:30:09 2011 +0900 @@ -1011,9 +1011,61 @@ lemma prod_alpha_eq: shows "prod_alpha (op=) (op=) = (op=)" -unfolding prod_alpha_def -by (auto intro!: ext) + unfolding prod_alpha_def + by (auto intro!: ext) + +lemma Abs_lst1_fcb: + fixes x y :: "'a :: at_base" + and S T :: "'b :: fs" + assumes e: "(Abs_lst [atom x] T) = (Abs_lst [atom y] S)" + and f1: "x \ y \ atom y \ T \ atom x \ (atom y \ atom x) \ T \ atom x \ f x T" + and f2: "x \ y \ atom y \ T \ atom x \ (atom y \ atom x) \ T \ atom y \ f x T" + and p: "S = (atom x \ atom y) \ T \ x \ y \ atom y \ T \ atom x \ S \ (atom x \ atom y) \ (f x T) = f y S" + and s: "sort_of (atom x) = sort_of (atom y)" + shows "f x T = f y S" + using e + apply(case_tac "atom x \ S") + apply(simp add: Abs1_eq_iff'[OF s s]) + apply(elim conjE disjE) + apply(simp) + apply(rule trans) + apply(rule_tac p="(atom x \ atom y)" in supp_perm_eq[symmetric]) + apply(rule fresh_star_supp_conv) + apply(simp add: supp_swap fresh_star_def s f1 f2) + apply(simp add: swap_commute p) + apply(simp add: Abs1_eq_iff[OF s s]) + done +lemma Abs_res_fcb: + fixes xs ys :: "('a :: at_base) set" + and S T :: "'b :: fs" + assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)" + and f1: "\x. x \ atom ` xs \ x \ supp T \ x \ f xs T" + and f2: "\x. supp T - atom ` xs = supp S - atom ` ys \ x \ atom ` ys \ x \ supp S \ x \ f xs T" + and eqv: "\p. p \ T = S \ supp p \ atom ` xs \ supp T \ atom ` ys \ supp S + \ p \ (atom ` xs \ supp T) = atom ` ys \ supp S \ p \ (f xs T) = f ys S" + shows "f xs T = f ys S" + using e apply - + apply(subst (asm) Abs_eq_res_set) + apply(subst (asm) Abs_eq_iff2) + apply(simp add: alphas) + apply(elim exE conjE) + apply(rule trans) + apply(rule_tac p="p" in supp_perm_eq[symmetric]) + apply(rule fresh_star_supp_conv) + apply(drule fresh_star_perm_set_conv) + apply(rule finite_Diff) + apply(rule finite_supp) + apply(subgoal_tac "(atom ` xs \ supp T \ atom ` ys \ supp S) \* f xs T") + apply(metis Un_absorb2 fresh_star_Un) + apply(subst fresh_star_Un) + apply(rule conjI) + apply(simp add: fresh_star_def f1) + apply(subgoal_tac "supp T - atom ` xs = supp S - atom ` ys") + apply(simp add: fresh_star_def f2) + apply(blast) + apply(simp add: eqv) + done end