diff -r eae6549dbea2 -r a0088f1c049d src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Fri May 27 20:23:55 2016 +0200 +++ b/src/HOL/Topological_Spaces.thy Fri May 27 23:35:13 2016 +0200 @@ -331,7 +331,7 @@ have *: "{False <..} = {True}" "{..< True} = {False}" by auto have "A = UNIV \ A = {} \ A = {False <..} \ A = {..< True}" - using subset_UNIV[of A] unfolding UNIV_bool * by auto + using subset_UNIV[of A] unfolding UNIV_bool * by blast then show "open A" by auto qed @@ -1194,7 +1194,7 @@ fix S assume "open S" "x \ S" from incl[OF this] obtain i where "F i \ S" unfolding F_def by auto moreover have "\j. i \ j \ F j \ F i" - by (auto simp: F_def) + by (simp add: Inf_superset_mono F_def image_mono) ultimately show "eventually (\i. F i \ S) sequentially" by (auto simp: eventually_sequentially) qed @@ -1210,7 +1210,7 @@ show thesis proof show "decseq (\n. \i\n. A i)" - by (auto simp: decseq_def) + by (simp add: antimono_iff_le_Suc atMost_Suc) show "\n. x \ (\i\n. A i)" "\n. open (\i\n. A i)" using A by auto show "nhds x = (INF n. principal (\i\n. A i))" @@ -1735,7 +1735,7 @@ lemma continuous_within_compose3: "isCont g (f x) \ continuous (at x within s) f \ continuous (at x within s) (\x. g (f x))" - using continuous_within_compose2[of x s f g] by (simp add: continuous_at_imp_continuous_at_within) + using continuous_at_imp_continuous_at_within continuous_within_compose2 by blast lemma filtermap_nhds_open_map: assumes cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" @@ -2091,7 +2091,7 @@ unfolding continuous_on_open_invariant by (simp add: open_discrete) from this[of True] this[of False] obtain t f where "open t" "open f" and *: "f \ S = P -` {False} \ S" "t \ S = P -` {True} \ S" - by auto + by meson then have "t \ S = {} \ f \ S = {}" by (intro connectedD[OF \connected S\]) auto then show "\c. \s\S. P s = c" @@ -2253,7 +2253,7 @@ using open_right[OF \open A\ \?z \ A\ \?z < y\] by auto moreover obtain b where "b \ B" "x < b" "b < min a y" using cInf_less_iff[of "B \ {x <..}" "min a y"] \?z < a\ \?z < y\ \x < y\ \y \ B\ - by (auto intro: less_imp_le) + by auto moreover have "?z \ b" using \b \ B\ \x < b\ by (intro cInf_lower) auto @@ -2593,7 +2593,7 @@ by (auto simp add: totally_bounded_def) lemma totally_bounded_subset: "totally_bounded S \ T \ S \ totally_bounded T" - by (force simp add: totally_bounded_def) + by (fastforce simp add: totally_bounded_def) lemma totally_bounded_Union[intro]: assumes M: "finite M" "\S. S \ M \ totally_bounded S" shows "totally_bounded (\M)"