src/HOL/Wellfounded.thy
changeset 45970 b6d0cff57d96
parent 45139 bdcaa3f3a2f4
child 46177 adac34829e10
--- a/src/HOL/Wellfounded.thy	Sat Dec 24 15:53:10 2011 +0100
+++ b/src/HOL/Wellfounded.thy	Sat Dec 24 15:53:10 2011 +0100
@@ -298,8 +298,10 @@
 
 lemma wfP_SUP:
   "\<forall>i. wfP (r i) \<Longrightarrow> \<forall>i j. r i \<noteq> r j \<longrightarrow> inf (DomainP (r i)) (RangeP (r j)) = bot \<Longrightarrow> wfP (SUPR UNIV r)"
-  by (rule wf_UN [where I=UNIV and r="\<lambda>i. {(x, y). r i x y}", to_pred SUP_UN_eq2])
-    (simp_all add: Collect_def)
+  apply (rule wf_UN [where I=UNIV and r="\<lambda>i. {(x, y). r i x y}", to_pred SUP_UN_eq2])
+  apply (simp_all add: inf_set_def)
+  apply auto
+  done
 
 lemma wf_Union: 
  "[| ALL r:R. wf r;