1034 apply(simp add: supp_swap fresh_star_def s f1 f2) |
1034 apply(simp add: supp_swap fresh_star_def s f1 f2) |
1035 apply(simp add: swap_commute p) |
1035 apply(simp add: swap_commute p) |
1036 apply(simp add: Abs1_eq_iff[OF s s]) |
1036 apply(simp add: Abs1_eq_iff[OF s s]) |
1037 done |
1037 done |
1038 |
1038 |
|
1039 lemma Abs_lst_fcb: |
|
1040 fixes xs ys :: "'a :: fs" |
|
1041 and S T :: "'b :: fs" |
|
1042 assumes e: "(Abs_lst (ba xs) T) = (Abs_lst (ba ys) S)" |
|
1043 and f1: "\<And>x. x \<in> set (ba xs) \<Longrightarrow> x \<sharp> f xs T" |
|
1044 and f2: "\<And>x. supp T - set (ba xs) = supp S - set (ba ys) \<Longrightarrow> x \<in> set (ba ys) \<Longrightarrow> x \<sharp> f xs T" |
|
1045 and eqv: "\<And>p. p \<bullet> T = S \<Longrightarrow> p \<bullet> ba xs = ba ys \<Longrightarrow> supp p \<subseteq> set (ba xs) \<union> set (ba ys) \<Longrightarrow> p \<bullet> (f xs T) = f ys S" |
|
1046 shows "f xs T = f ys S" |
|
1047 using e apply - |
|
1048 apply(subst (asm) Abs_eq_iff2) |
|
1049 apply(simp add: alphas) |
|
1050 apply(elim exE conjE) |
|
1051 apply(rule trans) |
|
1052 apply(rule_tac p="p" in supp_perm_eq[symmetric]) |
|
1053 apply(rule fresh_star_supp_conv) |
|
1054 apply(drule fresh_star_perm_set_conv) |
|
1055 apply(rule finite_Diff) |
|
1056 apply(rule finite_supp) |
|
1057 apply(subgoal_tac "(set (ba xs) \<union> set (ba ys)) \<sharp>* f xs T") |
|
1058 apply(metis Un_absorb2 fresh_star_Un) |
|
1059 apply(subst fresh_star_Un) |
|
1060 apply(rule conjI) |
|
1061 apply(simp add: fresh_star_def f1) |
|
1062 apply(simp add: fresh_star_def f2) |
|
1063 apply(simp add: eqv) |
|
1064 done |
|
1065 |
1039 lemma Abs_res_fcb: |
1066 lemma Abs_res_fcb: |
1040 fixes xs ys :: "('a :: at_base) set" |
1067 fixes xs ys :: "('a :: at_base) set" |
1041 and S T :: "'b :: fs" |
1068 and S T :: "'b :: fs" |
1042 assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)" |
1069 assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)" |
1043 and f1: "\<And>x. x \<in> atom ` xs \<Longrightarrow> x \<in> supp T \<Longrightarrow> x \<sharp> f xs T" |
1070 and f1: "\<And>x. x \<in> atom ` xs \<Longrightarrow> x \<in> supp T \<Longrightarrow> x \<sharp> f xs T" |