# HG changeset patch # User Cezary Kaliszyk # Date 1308531558 -32400 # Node ID ab2aded5f7c91f3e30a435b15c6ce52ab5987628 # Parent 1628e47fa57c7dfa5c5a9cbcfdb3f0106774a6a2 Move lst_fcb to Nominal2_Abs diff -r 1628e47fa57c -r ab2aded5f7c9 Nominal/Ex/Let.thy --- a/Nominal/Ex/Let.thy Mon Jun 20 09:38:57 2011 +0900 +++ b/Nominal/Ex/Let.thy Mon Jun 20 09:59:18 2011 +0900 @@ -90,33 +90,6 @@ lemmas alpha_bn_inducts = alpha_bn_inducts_raw[quot_lifted] -lemma Abs_lst_fcb: - fixes xs ys :: "'a :: fs" - and S T :: "'b :: fs" - assumes e: "(Abs_lst (ba xs) T) = (Abs_lst (ba ys) S)" - and f1: "\x. x \ set (ba xs) \ x \ f xs T" - and f2: "\x. supp T - set (ba xs) = supp S - set (ba ys) \ x \ set (ba ys) \ x \ f xs T" - and eqv: "\p. p \ T = S \ p \ ba xs = ba ys \ supp p \ set (ba xs) \ set (ba ys) \ p \ (f xs T) = f ys S" - shows "f xs T = f ys S" - using e apply - - 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 "(set (ba xs) \ set (ba ys)) \* 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(simp add: fresh_star_def f2) - apply(simp add: eqv) - done - lemma max_eqvt[eqvt]: "p \ (max (a :: _ :: pure) b) = max (p \ a) (p \ b)" by (simp add: permute_pure) diff -r 1628e47fa57c -r ab2aded5f7c9 Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Mon Jun 20 09:38:57 2011 +0900 +++ b/Nominal/Nominal2_Abs.thy Mon Jun 20 09:59:18 2011 +0900 @@ -1036,6 +1036,33 @@ apply(simp add: Abs1_eq_iff[OF s s]) done +lemma Abs_lst_fcb: + fixes xs ys :: "'a :: fs" + and S T :: "'b :: fs" + assumes e: "(Abs_lst (ba xs) T) = (Abs_lst (ba ys) S)" + and f1: "\x. x \ set (ba xs) \ x \ f xs T" + and f2: "\x. supp T - set (ba xs) = supp S - set (ba ys) \ x \ set (ba ys) \ x \ f xs T" + and eqv: "\p. p \ T = S \ p \ ba xs = ba ys \ supp p \ set (ba xs) \ set (ba ys) \ p \ (f xs T) = f ys S" + shows "f xs T = f ys S" + using e apply - + 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 "(set (ba xs) \ set (ba ys)) \* 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(simp add: fresh_star_def f2) + apply(simp add: eqv) + done + lemma Abs_res_fcb: fixes xs ys :: "('a :: at_base) set" and S T :: "'b :: fs"