# HG changeset patch # User Cezary Kaliszyk # Date 1307521554 -32400 # Node ID fff147e99375edf97ca225f94c240c7d64cc0432 # Parent 297cff1d10483db3ac3c7d40cebb7d95a1ee6f5e Simplify fcb_res diff -r 297cff1d1048 -r fff147e99375 Nominal/Ex/TypeSchemes.thy --- a/Nominal/Ex/TypeSchemes.thy Wed Jun 08 09:56:39 2011 +0900 +++ b/Nominal/Ex/TypeSchemes.thy Wed Jun 08 17:25:54 2011 +0900 @@ -380,7 +380,7 @@ 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 T = supp S - atom ` ys \ supp S \ x \ atom ` ys \ x \ supp S \ 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" @@ -401,7 +401,9 @@ 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