# HG changeset patch # User huffman # Date 1358440276 28800 # Node ID ada575c605e1b493afbd85cb7a01f27f1bd7d402 # Parent 7bc58677860ed6d94b6c1bbf3c94210ca9146702 simplify proof of compact_imp_bounded diff -r 7bc58677860e -r ada575c605e1 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jan 17 19:20:56 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jan 17 08:31:16 2013 -0800 @@ -2205,6 +2205,9 @@ lemma bounded_Union[intro]: "finite F \ (\S\F. bounded S) \ bounded(\F)" by (induct rule: finite_induct[of F], auto) +lemma bounded_UN [intro]: "finite A \ \x\A. bounded (B x) \ bounded (\x\A. B x)" + by (induct set: finite, auto) + lemma bounded_insert [simp]: "bounded (insert x S) \ bounded S" proof - have "\y\{x}. dist x y \ 0" by simp @@ -2583,21 +2586,10 @@ have "compact U" "\x\U. open (ball x 1)" "U \ (\x\U. ball x 1)" using assms by auto then obtain D where D: "D \ U" "finite D" "U \ (\x\D. ball x 1)" by (elim compactE_image) - def d \ "SOME d. d \ D" - show "bounded U" - unfolding bounded_def - proof (intro exI, safe) - fix x assume "x \ U" - with D obtain d' where "d' \ D" "x \ ball d' 1" by auto - moreover have "dist d x \ dist d d' + dist d' x" - using dist_triangle[of d x d'] by (simp add: dist_commute) - moreover - from `x\U` D have "d \ D" - unfolding d_def by (rule_tac someI_ex) auto - ultimately - show "dist d x \ Max ((\d'. dist d d' + 1) ` D)" - using D by (subst Max_ge_iff) (auto intro!: bexI[of _ d']) - qed + from `finite D` have "bounded (\x\D. ball x 1)" + by (simp add: bounded_UN) + thus "bounded U" using `U \ (\x\D. ball x 1)` + by (rule bounded_subset) qed text{* In particular, some common special cases. *}