Nominal/Nominal2_Abs.thy
changeset 2878 06d91b7b5756
parent 2875 ab2aded5f7c9
child 2909 de5c9a0040ec
equal deleted inserted replaced
2876:99583bd6a7b2 2878:06d91b7b5756
  1061   apply(simp add: fresh_star_def f1)
  1061   apply(simp add: fresh_star_def f1)
  1062   apply(simp add: fresh_star_def f2)
  1062   apply(simp add: fresh_star_def f2)
  1063   apply(simp add: eqv)
  1063   apply(simp add: eqv)
  1064   done
  1064   done
  1065 
  1065 
       
  1066 lemma Abs_set_fcb:
       
  1067   fixes xs ys :: "'a :: fs"
       
  1068     and S T :: "'b :: fs"
       
  1069   assumes e: "(Abs_set (ba xs) T) = (Abs_set (ba ys) S)"
       
  1070     and f1: "\<And>x. x \<in> ba xs \<Longrightarrow> x \<sharp> f xs T"
       
  1071     and f2: "\<And>x. supp T - ba xs = supp S - ba ys \<Longrightarrow> x \<in> ba ys \<Longrightarrow> x \<sharp> f xs T"
       
  1072     and eqv: "\<And>p. p \<bullet> T = S \<Longrightarrow> p \<bullet> ba xs = ba ys \<Longrightarrow> supp p \<subseteq> ba xs \<union> ba ys \<Longrightarrow> p \<bullet> (f xs T) = f ys S"
       
  1073   shows "f xs T = f ys S"
       
  1074   using e apply -
       
  1075   apply(subst (asm) Abs_eq_iff2)
       
  1076   apply(simp add: alphas)
       
  1077   apply(elim exE conjE)
       
  1078   apply(rule trans)
       
  1079   apply(rule_tac p="p" in supp_perm_eq[symmetric])
       
  1080   apply(rule fresh_star_supp_conv)
       
  1081   apply(drule fresh_star_perm_set_conv)
       
  1082   apply(rule finite_Diff)
       
  1083   apply(rule finite_supp)
       
  1084   apply(subgoal_tac "(ba xs \<union> ba ys) \<sharp>* f xs T")
       
  1085   apply(metis Un_absorb2 fresh_star_Un)
       
  1086   apply(subst fresh_star_Un)
       
  1087   apply(rule conjI)
       
  1088   apply(simp add: fresh_star_def f1)
       
  1089   apply(simp add: fresh_star_def f2)
       
  1090   apply(simp add: eqv)
       
  1091   done
       
  1092 
  1066 lemma Abs_res_fcb:
  1093 lemma Abs_res_fcb:
  1067   fixes xs ys :: "('a :: at_base) set"
  1094   fixes xs ys :: "('a :: at_base) set"
  1068     and S T :: "'b :: fs"
  1095     and S T :: "'b :: fs"
  1069   assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)"
  1096   assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)"
  1070     and f1: "\<And>x. x \<in> atom ` xs \<Longrightarrow> x \<in> supp T \<Longrightarrow> x \<sharp> f xs T"
  1097     and f1: "\<And>x. x \<in> atom ` xs \<Longrightarrow> x \<in> supp T \<Longrightarrow> x \<sharp> f xs T"