diff -r 7b9ff966974f -r da28fd2852ed src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy Wed Dec 04 19:55:30 2019 +0100 +++ b/src/HOL/Analysis/Homotopy.thy Wed Dec 04 23:11:29 2019 +0100 @@ -1478,6 +1478,38 @@ using homotopic_through_contractible [of S id S f T id g] by (simp add: assms contractible_imp_path_connected) +subsection\Starlike sets\ + +definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" + +lemma starlike_UNIV [simp]: "starlike UNIV" + by (simp add: starlike_def) + +lemma convex_imp_starlike: + "convex S \ S \ {} \ starlike S" + unfolding convex_contains_segment starlike_def by auto + +lemma starlike_convex_tweak_boundary_points: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "S \ {}" and ST: "rel_interior S \ T" and TS: "T \ closure S" + shows "starlike T" +proof - + have "rel_interior S \ {}" + by (simp add: assms rel_interior_eq_empty) + then obtain a where a: "a \ rel_interior S" by blast + with ST have "a \ T" by blast + have *: "\x. x \ T \ open_segment a x \ rel_interior S" + apply (rule rel_interior_closure_convex_segment [OF \convex S\ a]) + using assms by blast + show ?thesis + unfolding starlike_def + apply (rule bexI [OF _ \a \ T\]) + apply (simp add: closed_segment_eq_open) + apply (intro conjI ballI a \a \ T\ rel_interior_closure_convex_segment [OF \convex S\ a]) + apply (simp add: order_trans [OF * ST]) + done +qed + lemma starlike_imp_contractible_gen: fixes S :: "'a::real_normed_vector set" assumes S: "starlike S"