diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Fri Apr 12 19:48:29 2019 +0200 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Fri Apr 12 22:09:25 2019 +0200 @@ -231,16 +231,16 @@ definitions, but we need to split them into two implications because of the lack of type quantifiers. Then ENR turns out to be equivalent to ANR plus local compactness."\ -definition%important AR :: "'a::topological_space set \ bool" where +definition\<^marker>\tag important\ AR :: "'a::topological_space set \ bool" where "AR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ S' retract_of U" -definition%important ANR :: "'a::topological_space set \ bool" where +definition\<^marker>\tag important\ ANR :: "'a::topological_space set \ bool" where "ANR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ (\T. openin (top_of_set U) T \ S' retract_of T)" -definition%important ENR :: "'a::topological_space set \ bool" where +definition\<^marker>\tag important\ ENR :: "'a::topological_space set \ bool" where "ENR S \ \U. open U \ S retract_of U" text \First, show that we do indeed get the "usual" properties of ARs and ANRs.\ @@ -4072,7 +4072,7 @@ by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial) qed -corollary%unimportant ANR_sphere: +corollary\<^marker>\tag unimportant\ ANR_sphere: fixes a :: "'a::euclidean_space" shows "ANR(sphere a r)" by (simp add: ENR_imp_ANR ENR_sphere) @@ -4254,7 +4254,7 @@ qed -corollary%unimportant nullhomotopic_into_ANR_extension: +corollary\<^marker>\tag unimportant\ nullhomotopic_into_ANR_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" @@ -4288,7 +4288,7 @@ by (force elim: homotopic_with_eq [of _ _ _ g "\x. c"] simp: \\x. x \ S \ g x = f x\) qed -corollary%unimportant nullhomotopic_into_rel_frontier_extension: +corollary\<^marker>\tag unimportant\ nullhomotopic_into_rel_frontier_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" @@ -4299,7 +4299,7 @@ (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x \ S. g x = f x))" by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex) -corollary%unimportant nullhomotopic_into_sphere_extension: +corollary\<^marker>\tag unimportant\ nullhomotopic_into_sphere_extension: fixes f :: "'a::euclidean_space \ 'b :: euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "S \ {}" and fim: "f ` S \ sphere a r" @@ -4321,7 +4321,7 @@ done qed -proposition%unimportant Borsuk_map_essential_bounded_component: +proposition\<^marker>\tag unimportant\ Borsuk_map_essential_bounded_component: fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S" shows "bounded (connected_component_set (- S) a) \ @@ -4578,7 +4578,7 @@ qed qed -proposition%unimportant homotopic_neighbourhood_extension: +proposition\<^marker>\tag unimportant\ homotopic_neighbourhood_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and fim: "f ` S \ U" and contg: "continuous_on S g" and gim: "g ` S \ U" @@ -4643,7 +4643,7 @@ qed text\ Homotopy on a union of closed-open sets.\ -proposition%unimportant homotopic_on_clopen_Union: +proposition\<^marker>\tag unimportant\ homotopic_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S"