# HG changeset patch # User berghofe # Date 1233264483 -3600 # Node ID 9f03b5f847cde54c239f667d5b4d3d9b85b09fea # Parent c81f8b2967e1097c63937b9d11d8760791ed4ebe Added strong congruence rule for UN. diff -r c81f8b2967e1 -r 9f03b5f847cd src/HOL/Set.thy --- a/src/HOL/Set.thy Thu Jan 29 22:27:07 2009 +0100 +++ b/src/HOL/Set.thy Thu Jan 29 22:28:03 2009 +0100 @@ -885,6 +885,10 @@ "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" by (simp add: UNION_def) +lemma strong_UN_cong: + "A = B ==> (!!x. x:B =simp=> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" + by (simp add: UNION_def simp_implies_def) + subsubsection {* Intersections of families *}