# HG changeset patch # User nipkow # Date 1531387426 -7200 # Node ID 75129a73aca3b521857565944e8dd324889bf8cb # Parent cedf3480fdadd252f1750e834438716b8784a73d more economic tagging diff -r cedf3480fdad -r 75129a73aca3 src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Thu Jul 12 11:23:46 2018 +0200 @@ -12,12 +12,1748 @@ (* *) (* (c) Copyright, John Harrison 1998-2008 *) -section \Results connected with topological dimension\ +section \Brouwer's Fixed Point Theorem\ theory Brouwer_Fixpoint imports Path_Connected Homeomorphism begin +subsection \Unit cubes\ + +(* FIXME mv euclidean topological space *) +definition unit_cube :: "'a::euclidean_space set" + where "unit_cube = {x. \i\Basis. 0 \ x \ i \ x \ i \ 1}" + +lemma mem_unit_cube: "x \ unit_cube \ (\i\Basis. 0 \ x \ i \ x \ i \ 1)" + unfolding unit_cube_def by simp + +lemma bounded_unit_cube: "bounded unit_cube" + unfolding bounded_def +proof (intro exI ballI) + fix y :: 'a assume y: "y \ unit_cube" + have "dist 0 y = norm y" by (rule dist_0_norm) + also have "\ = norm (\i\Basis. (y \ i) *\<^sub>R i)" unfolding euclidean_representation .. + also have "\ \ (\i\Basis. norm ((y \ i) *\<^sub>R i))" by (rule norm_sum) + also have "\ \ (\i::'a\Basis. 1)" + by (rule sum_mono, simp add: y [unfolded mem_unit_cube]) + finally show "dist 0 y \ (\i::'a\Basis. 1)" . +qed + +lemma closed_unit_cube: "closed unit_cube" + unfolding unit_cube_def Collect_ball_eq Collect_conj_eq + by (rule closed_INT, auto intro!: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma compact_unit_cube: "compact unit_cube" (is "compact ?C") + unfolding compact_eq_seq_compact_metric + using bounded_unit_cube closed_unit_cube + by (rule bounded_closed_imp_seq_compact) + +lemma convex_unit_cube: "convex unit_cube" + by (rule is_interval_convex) (fastforce simp add: is_interval_def mem_unit_cube) + + +(* FIXME mv topology euclidean space *) +subsection \Retractions\ + +definition "retraction S T r \ T \ S \ continuous_on S r \ r ` S \ T \ (\x\T. r x = x)" + +definition retract_of (infixl "retract'_of" 50) + where "(T retract_of S) \ (\r. retraction S T r)" + +lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" + unfolding retraction_def by auto + +text \Preservation of fixpoints under (more general notion of) retraction\ + +lemma invertible_fixpoint_property: + fixes S :: "'a::euclidean_space set" + and T :: "'b::euclidean_space set" + assumes contt: "continuous_on T i" + and "i ` T \ S" + and contr: "continuous_on S r" + and "r ` S \ T" + and ri: "\y. y \ T \ r (i y) = y" + and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" + and contg: "continuous_on T g" + and "g ` T \ T" + obtains y where "y \ T" and "g y = y" +proof - + have "\x\S. (i \ g \ r) x = x" + proof (rule FP) + show "continuous_on S (i \ g \ r)" + by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) + show "(i \ g \ r) ` S \ S" + using assms(2,4,8) by force + qed + then obtain x where x: "x \ S" "(i \ g \ r) x = x" .. + then have *: "g (r x) \ T" + using assms(4,8) by auto + have "r ((i \ g \ r) x) = r x" + using x by auto + then show ?thesis + using "*" ri that by auto +qed + +lemma homeomorphic_fixpoint_property: + fixes S :: "'a::euclidean_space set" + and T :: "'b::euclidean_space set" + assumes "S homeomorphic T" + shows "(\f. continuous_on S f \ f ` S \ S \ (\x\S. f x = x)) \ + (\g. continuous_on T g \ g ` T \ T \ (\y\T. g y = y))" + (is "?lhs = ?rhs") +proof - + obtain r i where r: + "\x\S. i (r x) = x" "r ` S = T" "continuous_on S r" + "\y\T. r (i y) = y" "i ` T = S" "continuous_on T i" + using assms unfolding homeomorphic_def homeomorphism_def by blast + show ?thesis + proof + assume ?lhs + with r show ?rhs + by (metis invertible_fixpoint_property[of T i S r] order_refl) + next + assume ?rhs + with r show ?lhs + by (metis invertible_fixpoint_property[of S r T i] order_refl) + qed +qed + +lemma retract_fixpoint_property: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + and S :: "'a set" + assumes "T retract_of S" + and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" + and contg: "continuous_on T g" + and "g ` T \ T" + obtains y where "y \ T" and "g y = y" +proof - + obtain h where "retraction S T h" + using assms(1) unfolding retract_of_def .. + then show ?thesis + unfolding retraction_def + using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] + by (metis assms(4) contg image_ident that) +qed + +lemma retraction: + "retraction S T r \ + T \ S \ continuous_on S r \ r ` S = T \ (\x \ T. r x = x)" +by (force simp: retraction_def) + +lemma retract_of_imp_extensible: + assumes "S retract_of T" and "continuous_on S f" and "f ` S \ U" + obtains g where "continuous_on T g" "g ` T \ U" "\x. x \ S \ g x = f x" +using assms +apply (clarsimp simp add: retract_of_def retraction) +apply (rule_tac g = "f \ r" in that) +apply (auto simp: continuous_on_compose2) +done + +lemma idempotent_imp_retraction: + assumes "continuous_on S f" and "f ` S \ S" and "\x. x \ S \ f(f x) = f x" + shows "retraction S (f ` S) f" +by (simp add: assms retraction) + +lemma retraction_subset: + assumes "retraction S T r" and "T \ s'" and "s' \ S" + shows "retraction s' T r" + unfolding retraction_def + by (metis assms continuous_on_subset image_mono retraction) + +lemma retract_of_subset: + assumes "T retract_of S" and "T \ s'" and "s' \ S" + shows "T retract_of s'" +by (meson assms retract_of_def retraction_subset) + +lemma retraction_refl [simp]: "retraction S S (\x. x)" +by (simp add: continuous_on_id retraction) + +lemma retract_of_refl [iff]: "S retract_of S" + unfolding retract_of_def retraction_def + using continuous_on_id by blast + +lemma retract_of_imp_subset: + "S retract_of T \ S \ T" +by (simp add: retract_of_def retraction_def) + +lemma retract_of_empty [simp]: + "({} retract_of S) \ S = {}" "(S retract_of {}) \ S = {}" +by (auto simp: retract_of_def retraction_def) + +lemma retract_of_singleton [iff]: "({x} retract_of S) \ x \ S" + unfolding retract_of_def retraction_def by force + +lemma retraction_comp: + "\retraction S T f; retraction T U g\ + \ retraction S U (g \ f)" +apply (auto simp: retraction_def intro: continuous_on_compose2) +by blast + +lemma retract_of_trans [trans]: + assumes "S retract_of T" and "T retract_of U" + shows "S retract_of U" +using assms by (auto simp: retract_of_def intro: retraction_comp) + +lemma closedin_retract: + fixes S :: "'a :: real_normed_vector set" + assumes "S retract_of T" + shows "closedin (subtopology euclidean T) S" +proof - + obtain r where "S \ T" "continuous_on T r" "r ` T \ S" "\x. x \ S \ r x = x" + using assms by (auto simp: retract_of_def retraction_def) + then have S: "S = {x \ T. (norm(r x - x)) = 0}" by auto + show ?thesis + apply (subst S) + apply (rule continuous_closedin_preimage_constant) + by (simp add: \continuous_on T r\ continuous_on_diff continuous_on_id continuous_on_norm) +qed + +lemma closedin_self [simp]: + fixes S :: "'a :: real_normed_vector set" + shows "closedin (subtopology euclidean S) S" + by (simp add: closedin_retract) + +lemma retract_of_contractible: + assumes "contractible T" "S retract_of T" + shows "contractible S" +using assms +apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with) +apply (rule_tac x="r a" in exI) +apply (rule_tac x="r \ h" in exI) +apply (intro conjI continuous_intros continuous_on_compose) +apply (erule continuous_on_subset | force)+ +done + +lemma retract_of_compact: + "\compact T; S retract_of T\ \ compact S" + by (metis compact_continuous_image retract_of_def retraction) + +lemma retract_of_closed: + fixes S :: "'a :: real_normed_vector set" + shows "\closed T; S retract_of T\ \ closed S" + by (metis closedin_retract closedin_closed_eq) + +lemma retract_of_connected: + "\connected T; S retract_of T\ \ connected S" + by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) + +lemma retract_of_path_connected: + "\path_connected T; S retract_of T\ \ path_connected S" + by (metis path_connected_continuous_image retract_of_def retraction) + +lemma retract_of_simply_connected: + "\simply_connected T; S retract_of T\ \ simply_connected S" +apply (simp add: retract_of_def retraction_def, clarify) +apply (rule simply_connected_retraction_gen) +apply (force simp: continuous_on_id elim!: continuous_on_subset)+ +done + +lemma retract_of_homotopically_trivial: + assumes ts: "T retract_of S" + and hom: "\f g. \continuous_on U f; f ` U \ S; + continuous_on U g; g ` U \ S\ + \ homotopic_with (\x. True) U S f g" + and "continuous_on U f" "f ` U \ T" + and "continuous_on U g" "g ` U \ T" + shows "homotopic_with (\x. True) U T f g" +proof - + obtain r where "r ` S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" + using ts by (auto simp: retract_of_def retraction) + then obtain k where "Retracts S r T k" + unfolding Retracts_def + by (metis continuous_on_subset dual_order.trans image_iff image_mono) + then show ?thesis + apply (rule Retracts.homotopically_trivial_retraction_gen) + using assms + apply (force simp: hom)+ + done +qed + +lemma retract_of_homotopically_trivial_null: + assumes ts: "T retract_of S" + and hom: "\f. \continuous_on U f; f ` U \ S\ + \ \c. homotopic_with (\x. True) U S f (\x. c)" + and "continuous_on U f" "f ` U \ T" + obtains c where "homotopic_with (\x. True) U T f (\x. c)" +proof - + obtain r where "r ` S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" + using ts by (auto simp: retract_of_def retraction) + then obtain k where "Retracts S r T k" + unfolding Retracts_def + by (metis continuous_on_subset dual_order.trans image_iff image_mono) + then show ?thesis + apply (rule Retracts.homotopically_trivial_retraction_null_gen) + apply (rule TrueI refl assms that | assumption)+ + done +qed + +lemma retraction_imp_quotient_map: + "retraction S T r + \ U \ T + \ (openin (subtopology euclidean S) (S \ r -` U) \ + openin (subtopology euclidean T) U)" +apply (clarsimp simp add: retraction) +apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) +apply (auto simp: elim: continuous_on_subset) +done + +lemma retract_of_locally_compact: + fixes S :: "'a :: {heine_borel,real_normed_vector} set" + shows "\ locally compact S; T retract_of S\ \ locally compact T" + by (metis locally_compact_closedin closedin_retract) + +lemma retract_of_Times: + "\S retract_of s'; T retract_of t'\ \ (S \ T) retract_of (s' \ t')" +apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) +apply (rename_tac f g) +apply (rule_tac x="\z. ((f \ fst) z, (g \ snd) z)" in exI) +apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ +done + +lemma homotopic_into_retract: + "\f ` S \ T; g ` S \ T; T retract_of U; homotopic_with (\x. True) S U f g\ + \ homotopic_with (\x. True) S T f g" +apply (subst (asm) homotopic_with_def) +apply (simp add: homotopic_with retract_of_def retraction_def, clarify) +apply (rule_tac x="r \ h" in exI) +apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ +done + +lemma retract_of_locally_connected: + assumes "locally connected T" "S retract_of T" + shows "locally connected S" + using assms + by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_connected_quotient_image) + +lemma retract_of_locally_path_connected: + assumes "locally path_connected T" "S retract_of T" + shows "locally path_connected S" + using assms + by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image) + +text \A few simple lemmas about deformation retracts\ + +lemma deformation_retract_imp_homotopy_eqv: + fixes S :: "'a::euclidean_space set" + assumes "homotopic_with (\x. True) S S id r" and r: "retraction S T r" + shows "S homotopy_eqv T" +proof - + have "homotopic_with (\x. True) S S (id \ r) id" + by (simp add: assms(1) homotopic_with_symD) + moreover have "homotopic_with (\x. True) T T (r \ id) id" + using r unfolding retraction_def + by (metis (no_types, lifting) comp_id continuous_on_id' homotopic_with_equal homotopic_with_symD id_def image_id order_refl) + ultimately + show ?thesis + unfolding homotopy_eqv_def + by (metis continuous_on_id' id_def image_id r retraction_def) +qed + +lemma deformation_retract: + fixes S :: "'a::euclidean_space set" + shows "(\r. homotopic_with (\x. True) S S id r \ retraction S T r) \ + T retract_of S \ (\f. homotopic_with (\x. True) S S id f \ f ` S \ T)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (auto simp: retract_of_def retraction_def) +next + assume ?rhs + then show ?lhs + apply (clarsimp simp add: retract_of_def retraction_def) + apply (rule_tac x=r in exI, simp) + apply (rule homotopic_with_trans, assumption) + apply (rule_tac f = "r \ f" and g="r \ id" in homotopic_with_eq) + apply (rule_tac Y=S in homotopic_compose_continuous_left) + apply (auto simp: homotopic_with_sym) + done +qed + +lemma deformation_retract_of_contractible_sing: + fixes S :: "'a::euclidean_space set" + assumes "contractible S" "a \ S" + obtains r where "homotopic_with (\x. True) S S id r" "retraction S {a} r" +proof - + have "{a} retract_of S" + by (simp add: \a \ S\) + moreover have "homotopic_with (\x. True) S S id (\x. a)" + using assms + by (auto simp: contractible_def continuous_on_const continuous_on_id homotopic_into_contractible image_subset_iff) + moreover have "(\x. a) ` S \ {a}" + by (simp add: image_subsetI) + ultimately show ?thesis + using that deformation_retract by metis +qed + + +lemma continuous_on_compact_surface_projection_aux: + fixes S :: "'a::t2_space set" + assumes "compact S" "S \ T" "image q T \ S" + and contp: "continuous_on T p" + and "\x. x \ S \ q x = x" + and [simp]: "\x. x \ T \ q(p x) = q x" + and "\x. x \ T \ p(q x) = p x" + shows "continuous_on T q" +proof - + have *: "image p T = image p S" + using assms by auto (metis imageI subset_iff) + have contp': "continuous_on S p" + by (rule continuous_on_subset [OF contp \S \ T\]) + have "continuous_on (p ` T) q" + by (simp add: "*" assms(1) assms(2) assms(5) continuous_on_inv contp' rev_subsetD) + then have "continuous_on T (q \ p)" + by (rule continuous_on_compose [OF contp]) + then show ?thesis + by (rule continuous_on_eq [of _ "q \ p"]) (simp add: o_def) +qed + +lemma continuous_on_compact_surface_projection: + fixes S :: "'a::real_normed_vector set" + assumes "compact S" + and S: "S \ V - {0}" and "cone V" + and iff: "\x k. x \ V - {0} \ 0 < k \ (k *\<^sub>R x) \ S \ d x = k" + shows "continuous_on (V - {0}) (\x. d x *\<^sub>R x)" +proof (rule continuous_on_compact_surface_projection_aux [OF \compact S\ S]) + show "(\x. d x *\<^sub>R x) ` (V - {0}) \ S" + using iff by auto + show "continuous_on (V - {0}) (\x. inverse(norm x) *\<^sub>R x)" + by (intro continuous_intros) force + show "\x. x \ S \ d x *\<^sub>R x = x" + by (metis S zero_less_one local.iff scaleR_one subset_eq) + show "d (x /\<^sub>R norm x) *\<^sub>R (x /\<^sub>R norm x) = d x *\<^sub>R x" if "x \ V - {0}" for x + using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \cone V\ + by (simp add: field_simps cone_def zero_less_mult_iff) + show "d x *\<^sub>R x /\<^sub>R norm (d x *\<^sub>R x) = x /\<^sub>R norm x" if "x \ V - {0}" for x + proof - + have "0 < d x" + using local.iff that by blast + then show ?thesis + by simp + qed +qed + +subsection \Absolute retracts, absolute neighbourhood retracts (ANR) and Euclidean neighbourhood retracts (ENR)\ + +text \Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also Euclidean neighbourhood +retracts (ENR). We define AR and ANR by specializing the standard definitions for a set to embedding +in spaces of higher dimension. + +John Harrison writes: "This turns out to be sufficient (since any set in $\mathbb{R}^n$ can be +embedded as a closed subset of a convex subset of $\mathbb{R}^{n+1}$) to derive the usual +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 AR :: "'a::topological_space set => bool" + where + "AR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (subtopology euclidean U) S' + \ S' retract_of U" + +definition ANR :: "'a::topological_space set => bool" + where + "ANR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (subtopology euclidean U) S' + \ (\T. openin (subtopology euclidean U) T \ + S' retract_of T)" + +definition 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.\ + +lemma AR_imp_absolute_extensor: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "AR S" and contf: "continuous_on T f" and "f ` T \ S" + and cloUT: "closedin (subtopology euclidean U) T" + obtains g where "continuous_on U g" "g ` U \ S" "\x. x \ T \ g x = f x" +proof - + have "aff_dim S < int (DIM('b \ real))" + using aff_dim_le_DIM [of S] by simp + then obtain C and S' :: "('b * real) set" + where C: "convex C" "C \ {}" + and cloCS: "closedin (subtopology euclidean C) S'" + and hom: "S homeomorphic S'" + by (metis that homeomorphic_closedin_convex) + then have "S' retract_of C" + using \AR S\ by (simp add: AR_def) + then obtain r where "S' \ C" and contr: "continuous_on C r" + and "r ` C \ S'" and rid: "\x. x\S' \ r x = x" + by (auto simp: retraction_def retract_of_def) + obtain g h where "homeomorphism S S' g h" + using hom by (force simp: homeomorphic_def) + then have "continuous_on (f ` T) g" + by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def) + then have contgf: "continuous_on T (g \ f)" + by (metis continuous_on_compose contf) + have gfTC: "(g \ f) ` T \ C" + proof - + have "g ` S = S'" + by (metis (no_types) \homeomorphism S S' g h\ homeomorphism_def) + with \S' \ C\ \f ` T \ S\ show ?thesis by force + qed + obtain f' where f': "continuous_on U f'" "f' ` U \ C" + "\x. x \ T \ f' x = (g \ f) x" + by (metis Dugundji [OF C cloUT contgf gfTC]) + show ?thesis + proof (rule_tac g = "h \ r \ f'" in that) + show "continuous_on U (h \ r \ f')" + apply (intro continuous_on_compose f') + using continuous_on_subset contr f' apply blast + by (meson \homeomorphism S S' g h\ \r ` C \ S'\ continuous_on_subset \f' ` U \ C\ homeomorphism_def image_mono) + show "(h \ r \ f') ` U \ S" + using \homeomorphism S S' g h\ \r ` C \ S'\ \f' ` U \ C\ + by (fastforce simp: homeomorphism_def) + show "\x. x \ T \ (h \ r \ f') x = f x" + using \homeomorphism S S' g h\ \f ` T \ S\ f' + by (auto simp: rid homeomorphism_def) + qed +qed + +lemma AR_imp_absolute_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "AR S" "S homeomorphic S'" + and clo: "closedin (subtopology euclidean U) S'" + shows "S' retract_of U" +proof - + obtain g h where hom: "homeomorphism S S' g h" + using assms by (force simp: homeomorphic_def) + have h: "continuous_on S' h" " h ` S' \ S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + obtain h' where h': "continuous_on U h'" "h' ` U \ S" + and h'h: "\x. x \ S' \ h' x = h x" + by (blast intro: AR_imp_absolute_extensor [OF \AR S\ h clo]) + have [simp]: "S' \ U" using clo closedin_limpt by blast + show ?thesis + proof (simp add: retraction_def retract_of_def, intro exI conjI) + show "continuous_on U (g \ h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \ h') ` U \ S'" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\x\S'. (g \ h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed +qed + +lemma AR_imp_absolute_retract_UNIV: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "AR S" and hom: "S homeomorphic S'" + and clo: "closed S'" + shows "S' retract_of UNIV" +apply (rule AR_imp_absolute_retract [OF \AR S\ hom]) +using clo closed_closedin by auto + +lemma absolute_extensor_imp_AR: + fixes S :: "'a::euclidean_space set" + assumes "\f :: 'a * real \ 'a. + \U T. \continuous_on T f; f ` T \ S; + closedin (subtopology euclidean U) T\ + \ \g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)" + shows "AR S" +proof (clarsimp simp: AR_def) + fix U and T :: "('a * real) set" + assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" + then obtain g h where hom: "homeomorphism S T g h" + by (force simp: homeomorphic_def) + have h: "continuous_on T h" " h ` T \ S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + obtain h' where h': "continuous_on U h'" "h' ` U \ S" + and h'h: "\x\T. h' x = h x" + using assms [OF h clo] by blast + have [simp]: "T \ U" + using clo closedin_imp_subset by auto + show "T retract_of U" + proof (simp add: retraction_def retract_of_def, intro exI conjI) + show "continuous_on U (g \ h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \ h') ` U \ T" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\x\T. (g \ h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed +qed + +lemma AR_eq_absolute_extensor: + fixes S :: "'a::euclidean_space set" + shows "AR S \ + (\f :: 'a * real \ 'a. + \U T. continuous_on T f \ f ` T \ S \ + closedin (subtopology euclidean U) T \ + (\g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)))" +apply (rule iffI) + apply (metis AR_imp_absolute_extensor) +apply (simp add: absolute_extensor_imp_AR) +done + +lemma AR_imp_retract: + fixes S :: "'a::euclidean_space set" + assumes "AR S \ closedin (subtopology euclidean U) S" + shows "S retract_of U" +using AR_imp_absolute_retract assms homeomorphic_refl by blast + +lemma AR_homeomorphic_AR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "AR T" "S homeomorphic T" + shows "AR S" +unfolding AR_def +by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym) + +lemma homeomorphic_AR_iff_AR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + shows "S homeomorphic T \ AR S \ AR T" +by (metis AR_homeomorphic_AR homeomorphic_sym) + + +lemma ANR_imp_absolute_neighbourhood_extensor: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "ANR S" and contf: "continuous_on T f" and "f ` T \ S" + and cloUT: "closedin (subtopology euclidean U) T" + obtains V g where "T \ V" "openin (subtopology euclidean U) V" + "continuous_on V g" + "g ` V \ S" "\x. x \ T \ g x = f x" +proof - + have "aff_dim S < int (DIM('b \ real))" + using aff_dim_le_DIM [of S] by simp + then obtain C and S' :: "('b * real) set" + where C: "convex C" "C \ {}" + and cloCS: "closedin (subtopology euclidean C) S'" + and hom: "S homeomorphic S'" + by (metis that homeomorphic_closedin_convex) + then obtain D where opD: "openin (subtopology euclidean C) D" and "S' retract_of D" + using \ANR S\ by (auto simp: ANR_def) + then obtain r where "S' \ D" and contr: "continuous_on D r" + and "r ` D \ S'" and rid: "\x. x \ S' \ r x = x" + by (auto simp: retraction_def retract_of_def) + obtain g h where homgh: "homeomorphism S S' g h" + using hom by (force simp: homeomorphic_def) + have "continuous_on (f ` T) g" + by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def homgh) + then have contgf: "continuous_on T (g \ f)" + by (intro continuous_on_compose contf) + have gfTC: "(g \ f) ` T \ C" + proof - + have "g ` S = S'" + by (metis (no_types) homeomorphism_def homgh) + then show ?thesis + by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology) + qed + obtain f' where contf': "continuous_on U f'" + and "f' ` U \ C" + and eq: "\x. x \ T \ f' x = (g \ f) x" + by (metis Dugundji [OF C cloUT contgf gfTC]) + show ?thesis + proof (rule_tac V = "U \ f' -` D" and g = "h \ r \ f'" in that) + show "T \ U \ f' -` D" + using cloUT closedin_imp_subset \S' \ D\ \f ` T \ S\ eq homeomorphism_image1 homgh + by fastforce + show ope: "openin (subtopology euclidean U) (U \ f' -` D)" + using \f' ` U \ C\ by (auto simp: opD contf' continuous_openin_preimage) + have conth: "continuous_on (r ` f' ` (U \ f' -` D)) h" + apply (rule continuous_on_subset [of S']) + using homeomorphism_def homgh apply blast + using \r ` D \ S'\ by blast + show "continuous_on (U \ f' -` D) (h \ r \ f')" + apply (intro continuous_on_compose conth + continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto) + done + show "(h \ r \ f') ` (U \ f' -` D) \ S" + using \homeomorphism S S' g h\ \f' ` U \ C\ \r ` D \ S'\ + by (auto simp: homeomorphism_def) + show "\x. x \ T \ (h \ r \ f') x = f x" + using \homeomorphism S S' g h\ \f ` T \ S\ eq + by (auto simp: rid homeomorphism_def) + qed +qed + + +corollary ANR_imp_absolute_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ANR S" "S homeomorphic S'" + and clo: "closedin (subtopology euclidean U) S'" + obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" +proof - + obtain g h where hom: "homeomorphism S S' g h" + using assms by (force simp: homeomorphic_def) + have h: "continuous_on S' h" " h ` S' \ S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + from ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo] + obtain V h' where "S' \ V" and opUV: "openin (subtopology euclidean U) V" + and h': "continuous_on V h'" "h' ` V \ S" + and h'h:"\x. x \ S' \ h' x = h x" + by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo]) + have "S' retract_of V" + proof (simp add: retraction_def retract_of_def, intro exI conjI \S' \ V\) + show "continuous_on V (g \ h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \ h') ` V \ S'" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\x\S'. (g \ h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed + then show ?thesis + by (rule that [OF opUV]) +qed + +corollary ANR_imp_absolute_neighbourhood_retract_UNIV: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'" + obtains V where "open V" "S' retract_of V" + using ANR_imp_absolute_neighbourhood_retract [OF \ANR S\ hom] +by (metis clo closed_closedin open_openin subtopology_UNIV) + +corollary neighbourhood_extension_into_ANR: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes contf: "continuous_on S f" and fim: "f ` S \ T" and "ANR T" "closed S" + obtains V g where "S \ V" "open V" "continuous_on V g" + "g ` V \ T" "\x. x \ S \ g x = f x" + using ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf fim] + by (metis \closed S\ closed_closedin open_openin subtopology_UNIV) + +lemma absolute_neighbourhood_extensor_imp_ANR: + fixes S :: "'a::euclidean_space set" + assumes "\f :: 'a * real \ 'a. + \U T. \continuous_on T f; f ` T \ S; + closedin (subtopology euclidean U) T\ + \ \V g. T \ V \ openin (subtopology euclidean U) V \ + continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)" + shows "ANR S" +proof (clarsimp simp: ANR_def) + fix U and T :: "('a * real) set" + assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" + then obtain g h where hom: "homeomorphism S T g h" + by (force simp: homeomorphic_def) + have h: "continuous_on T h" " h ` T \ S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + obtain V h' where "T \ V" and opV: "openin (subtopology euclidean U) V" + and h': "continuous_on V h'" "h' ` V \ S" + and h'h: "\x\T. h' x = h x" + using assms [OF h clo] by blast + have [simp]: "T \ U" + using clo closedin_imp_subset by auto + have "T retract_of V" + proof (simp add: retraction_def retract_of_def, intro exI conjI \T \ V\) + show "continuous_on V (g \ h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \ h') ` V \ T" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\x\T. (g \ h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed + then show "\V. openin (subtopology euclidean U) V \ T retract_of V" + using opV by blast +qed + +lemma ANR_eq_absolute_neighbourhood_extensor: + fixes S :: "'a::euclidean_space set" + shows "ANR S \ + (\f :: 'a * real \ 'a. + \U T. continuous_on T f \ f ` T \ S \ + closedin (subtopology euclidean U) T \ + (\V g. T \ V \ openin (subtopology euclidean U) V \ + continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)))" +apply (rule iffI) + apply (metis ANR_imp_absolute_neighbourhood_extensor) +apply (simp add: absolute_neighbourhood_extensor_imp_ANR) +done + +lemma ANR_imp_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" "closedin (subtopology euclidean U) S" + obtains V where "openin (subtopology euclidean U) V" "S retract_of V" +using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast + +lemma ANR_imp_absolute_closed_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ANR S" "S homeomorphic S'" and US': "closedin (subtopology euclidean U) S'" + obtains V W + where "openin (subtopology euclidean U) V" + "closedin (subtopology euclidean U) W" + "S' \ V" "V \ W" "S' retract_of W" +proof - + obtain Z where "openin (subtopology euclidean U) Z" and S'Z: "S' retract_of Z" + by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) + then have UUZ: "closedin (subtopology euclidean U) (U - Z)" + by auto + have "S' \ (U - Z) = {}" + using \S' retract_of Z\ closedin_retract closedin_subtopology by fastforce + then obtain V W + where "openin (subtopology euclidean U) V" + and "openin (subtopology euclidean U) W" + and "S' \ V" "U - Z \ W" "V \ W = {}" + using separation_normal_local [OF US' UUZ] by auto + moreover have "S' retract_of U - W" + apply (rule retract_of_subset [OF S'Z]) + using US' \S' \ V\ \V \ W = {}\ closedin_subset apply fastforce + using Diff_subset_conv \U - Z \ W\ by blast + ultimately show ?thesis + apply (rule_tac V=V and W = "U-W" in that) + using openin_imp_subset apply force+ + done +qed + +lemma ANR_imp_closed_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" "closedin (subtopology euclidean U) S" + obtains V W where "openin (subtopology euclidean U) V" + "closedin (subtopology euclidean U) W" + "S \ V" "V \ W" "S retract_of W" +by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) + +lemma ANR_homeomorphic_ANR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "ANR T" "S homeomorphic T" + shows "ANR S" +unfolding ANR_def +by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym) + +lemma homeomorphic_ANR_iff_ANR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + shows "S homeomorphic T \ ANR S \ ANR T" +by (metis ANR_homeomorphic_ANR homeomorphic_sym) + +subsubsection \Analogous properties of ENRs\ + +lemma ENR_imp_absolute_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ENR S" and hom: "S homeomorphic S'" + and "S' \ U" + obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" +proof - + obtain X where "open X" "S retract_of X" + using \ENR S\ by (auto simp: ENR_def) + then obtain r where "retraction X S r" + by (auto simp: retract_of_def) + have "locally compact S'" + using retract_of_locally_compact open_imp_locally_compact + homeomorphic_local_compactness \S retract_of X\ \open X\ hom by blast + then obtain W where UW: "openin (subtopology euclidean U) W" + and WS': "closedin (subtopology euclidean W) S'" + apply (rule locally_compact_closedin_open) + apply (rename_tac W) + apply (rule_tac W = "U \ W" in that, blast) + by (simp add: \S' \ U\ closedin_limpt) + obtain f g where hom: "homeomorphism S S' f g" + using assms by (force simp: homeomorphic_def) + have contg: "continuous_on S' g" + using hom homeomorphism_def by blast + moreover have "g ` S' \ S" by (metis hom equalityE homeomorphism_def) + ultimately obtain h where conth: "continuous_on W h" and hg: "\x. x \ S' \ h x = g x" + using Tietze_unbounded [of S' g W] WS' by blast + have "W \ U" using UW openin_open by auto + have "S' \ W" using WS' closedin_closed by auto + have him: "\x. x \ S' \ h x \ X" + by (metis (no_types) \S retract_of X\ hg hom homeomorphism_def image_insert insert_absorb insert_iff retract_of_imp_subset subset_eq) + have "S' retract_of (W \ h -` X)" + proof (simp add: retraction_def retract_of_def, intro exI conjI) + show "S' \ W" "S' \ h -` X" + using him WS' closedin_imp_subset by blast+ + show "continuous_on (W \ h -` X) (f \ r \ h)" + proof (intro continuous_on_compose) + show "continuous_on (W \ h -` X) h" + by (meson conth continuous_on_subset inf_le1) + show "continuous_on (h ` (W \ h -` X)) r" + proof - + have "h ` (W \ h -` X) \ X" + by blast + then show "continuous_on (h ` (W \ h -` X)) r" + by (meson \retraction X S r\ continuous_on_subset retraction) + qed + show "continuous_on (r ` h ` (W \ h -` X)) f" + apply (rule continuous_on_subset [of S]) + using hom homeomorphism_def apply blast + apply clarify + apply (meson \retraction X S r\ subsetD imageI retraction_def) + done + qed + show "(f \ r \ h) ` (W \ h -` X) \ S'" + using \retraction X S r\ hom + by (auto simp: retraction_def homeomorphism_def) + show "\x\S'. (f \ r \ h) x = x" + using \retraction X S r\ hom by (auto simp: retraction_def homeomorphism_def hg) + qed + then show ?thesis + apply (rule_tac V = "W \ h -` X" in that) + apply (rule openin_trans [OF _ UW]) + using \continuous_on W h\ \open X\ continuous_openin_preimage_eq apply blast+ + done +qed + +corollary ENR_imp_absolute_neighbourhood_retract_UNIV: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ENR S" "S homeomorphic S'" + obtains T' where "open T'" "S' retract_of T'" +by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV) + +lemma ENR_homeomorphic_ENR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "ENR T" "S homeomorphic T" + shows "ENR S" +unfolding ENR_def +by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym) + +lemma homeomorphic_ENR_iff_ENR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "S homeomorphic T" + shows "ENR S \ ENR T" +by (meson ENR_homeomorphic_ENR assms homeomorphic_sym) + +lemma ENR_translation: + fixes S :: "'a::euclidean_space set" + shows "ENR(image (\x. a + x) S) \ ENR S" +by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR) + +lemma ENR_linear_image_eq: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "linear f" "inj f" + shows "ENR (image f S) \ ENR S" +apply (rule homeomorphic_ENR_iff_ENR) +using assms homeomorphic_sym linear_homeomorphic_image by auto + +text \Some relations among the concepts. We also relate AR to being a retract of UNIV, which is +often a more convenient proxy in the closed case.\ + +lemma AR_imp_ANR: "AR S \ ANR S" + using ANR_def AR_def by fastforce + +lemma ENR_imp_ANR: + fixes S :: "'a::euclidean_space set" + shows "ENR S \ ANR S" +apply (simp add: ANR_def) +by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset) + +lemma ENR_ANR: + fixes S :: "'a::euclidean_space set" + shows "ENR S \ ANR S \ locally compact S" +proof + assume "ENR S" + then have "locally compact S" + using ENR_def open_imp_locally_compact retract_of_locally_compact by auto + then show "ANR S \ locally compact S" + using ENR_imp_ANR \ENR S\ by blast +next + assume "ANR S \ locally compact S" + then have "ANR S" "locally compact S" by auto + then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T" + using locally_compact_homeomorphic_closed + by (metis DIM_prod DIM_real Suc_eq_plus1 lessI) + then show "ENR S" + using \ANR S\ + apply (simp add: ANR_def) + apply (drule_tac x=UNIV in spec) + apply (drule_tac x=T in spec, clarsimp) + apply (meson ENR_def ENR_homeomorphic_ENR open_openin) + done +qed + + +lemma AR_ANR: + fixes S :: "'a::euclidean_space set" + shows "AR S \ ANR S \ contractible S \ S \ {}" + (is "?lhs = ?rhs") +proof + assume ?lhs + obtain C and S' :: "('a * real) set" + where "convex C" "C \ {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'" + apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) + using aff_dim_le_DIM [of S] by auto + with \AR S\ have "contractible S" + apply (simp add: AR_def) + apply (drule_tac x=C in spec) + apply (drule_tac x="S'" in spec, simp) + using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce + with \AR S\ show ?rhs + apply (auto simp: AR_imp_ANR) + apply (force simp: AR_def) + done +next + assume ?rhs + then obtain a and h:: "real \ 'a \ 'a" + where conth: "continuous_on ({0..1} \ S) h" + and hS: "h ` ({0..1} \ S) \ S" + and [simp]: "\x. h(0, x) = x" + and [simp]: "\x. h(1, x) = a" + and "ANR S" "S \ {}" + by (auto simp: contractible_def homotopic_with_def) + then have "a \ S" + by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) + have "\g. continuous_on W g \ g ` W \ S \ (\x\T. g x = f x)" + if f: "continuous_on T f" "f ` T \ S" + and WT: "closedin (subtopology euclidean W) T" + for W T and f :: "'a \ real \ 'a" + proof - + obtain U g + where "T \ U" and WU: "openin (subtopology euclidean W) U" + and contg: "continuous_on U g" + and "g ` U \ S" and gf: "\x. x \ T \ g x = f x" + using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \ANR S\, rule_format, OF f WT] + by auto + have WWU: "closedin (subtopology euclidean W) (W - U)" + using WU closedin_diff by fastforce + moreover have "(W - U) \ T = {}" + using \T \ U\ by auto + ultimately obtain V V' + where WV': "openin (subtopology euclidean W) V'" + and WV: "openin (subtopology euclidean W) V" + and "W - U \ V'" "T \ V" "V' \ V = {}" + using separation_normal_local [of W "W-U" T] WT by blast + then have WVT: "T \ (W - V) = {}" + by auto + have WWV: "closedin (subtopology euclidean W) (W - V)" + using WV closedin_diff by fastforce + obtain j :: " 'a \ real \ real" + where contj: "continuous_on W j" + and j: "\x. x \ W \ j x \ {0..1}" + and j0: "\x. x \ W - V \ j x = 1" + and j1: "\x. x \ T \ j x = 0" + by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment) + have Weq: "W = (W - V) \ (W - V')" + using \V' \ V = {}\ by force + show ?thesis + proof (intro conjI exI) + have *: "continuous_on (W - V') (\x. h (j x, g x))" + apply (rule continuous_on_compose2 [OF conth continuous_on_Pair]) + apply (rule continuous_on_subset [OF contj Diff_subset]) + apply (rule continuous_on_subset [OF contg]) + apply (metis Diff_subset_conv Un_commute \W - U \ V'\) + using j \g ` U \ S\ \W - U \ V'\ apply fastforce + done + show "continuous_on W (\x. if x \ W - V then a else h (j x, g x))" + apply (subst Weq) + apply (rule continuous_on_cases_local) + apply (simp_all add: Weq [symmetric] WWV continuous_on_const *) + using WV' closedin_diff apply fastforce + apply (auto simp: j0 j1) + done + next + have "h (j (x, y), g (x, y)) \ S" if "(x, y) \ W" "(x, y) \ V" for x y + proof - + have "j(x, y) \ {0..1}" + using j that by blast + moreover have "g(x, y) \ S" + using \V' \ V = {}\ \W - U \ V'\ \g ` U \ S\ that by fastforce + ultimately show ?thesis + using hS by blast + qed + with \a \ S\ \g ` U \ S\ + show "(\x. if x \ W - V then a else h (j x, g x)) ` W \ S" + by auto + next + show "\x\T. (if x \ W - V then a else h (j x, g x)) = f x" + using \T \ V\ by (auto simp: j0 j1 gf) + qed + qed + then show ?lhs + by (simp add: AR_eq_absolute_extensor) +qed + + +lemma ANR_retract_of_ANR: + fixes S :: "'a::euclidean_space set" + assumes "ANR T" "S retract_of T" + shows "ANR S" +using assms +apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def) +apply (clarsimp elim!: all_forward) +apply (erule impCE, metis subset_trans) +apply (clarsimp elim!: ex_forward) +apply (rule_tac x="r \ g" in exI) +by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans) + +lemma AR_retract_of_AR: + fixes S :: "'a::euclidean_space set" + shows "\AR T; S retract_of T\ \ AR S" +using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce + +lemma ENR_retract_of_ENR: + "\ENR T; S retract_of T\ \ ENR S" +by (meson ENR_def retract_of_trans) + +lemma retract_of_UNIV: + fixes S :: "'a::euclidean_space set" + shows "S retract_of UNIV \ AR S \ closed S" +by (metis AR_ANR AR_imp_retract ENR_def ENR_imp_ANR closed_UNIV closed_closedin contractible_UNIV empty_not_UNIV open_UNIV retract_of_closed retract_of_contractible retract_of_empty(1) subtopology_UNIV) + +lemma compact_AR: + fixes S :: "'a::euclidean_space set" + shows "compact S \ AR S \ compact S \ S retract_of UNIV" +using compact_imp_closed retract_of_UNIV by blast + +text \More properties of ARs, ANRs and ENRs\ + +lemma not_AR_empty [simp]: "~ AR({})" + by (auto simp: AR_def) + +lemma ENR_empty [simp]: "ENR {}" + by (simp add: ENR_def) + +lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)" + by (simp add: ENR_imp_ANR) + +lemma convex_imp_AR: + fixes S :: "'a::euclidean_space set" + shows "\convex S; S \ {}\ \ AR S" +apply (rule absolute_extensor_imp_AR) +apply (rule Dugundji, assumption+) +by blast + +lemma convex_imp_ANR: + fixes S :: "'a::euclidean_space set" + shows "convex S \ ANR S" +using ANR_empty AR_imp_ANR convex_imp_AR by blast + +lemma ENR_convex_closed: + fixes S :: "'a::euclidean_space set" + shows "\closed S; convex S\ \ ENR S" +using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast + +lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)" + using retract_of_UNIV by auto + +lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)" + by (simp add: AR_imp_ANR) + +lemma ENR_UNIV [simp]:"ENR UNIV" + using ENR_def by blast + +lemma AR_singleton: + fixes a :: "'a::euclidean_space" + shows "AR {a}" + using retract_of_UNIV by blast + +lemma ANR_singleton: + fixes a :: "'a::euclidean_space" + shows "ANR {a}" + by (simp add: AR_imp_ANR AR_singleton) + +lemma ENR_singleton: "ENR {a}" + using ENR_def by blast + +text \ARs closed under union\ + +lemma AR_closed_Un_local_aux: + fixes U :: "'a::euclidean_space set" + assumes "closedin (subtopology euclidean U) S" + "closedin (subtopology euclidean U) T" + "AR S" "AR T" "AR(S \ T)" + shows "(S \ T) retract_of U" +proof - + have "S \ T \ {}" + using assms AR_def by fastforce + have "S \ U" "T \ U" + using assms by (auto simp: closedin_imp_subset) + define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" + define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" + define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" + have US': "closedin (subtopology euclidean U) S'" + using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] + by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) + have UT': "closedin (subtopology euclidean U) T'" + using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] + by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) + have "S \ S'" + using S'_def \S \ U\ setdist_sing_in_set by fastforce + have "T \ T'" + using T'_def \T \ U\ setdist_sing_in_set by fastforce + have "S \ T \ W" "W \ U" + using \S \ U\ by (auto simp: W_def setdist_sing_in_set) + have "(S \ T) retract_of W" + apply (rule AR_imp_absolute_retract [OF \AR(S \ T)\]) + apply (simp add: homeomorphic_refl) + apply (rule closedin_subset_trans [of U]) + apply (simp_all add: assms closedin_Int \S \ T \ W\ \W \ U\) + done + then obtain r0 + where "S \ T \ W" and contr0: "continuous_on W r0" + and "r0 ` W \ S \ T" + and r0 [simp]: "\x. x \ S \ T \ r0 x = x" + by (auto simp: retract_of_def retraction_def) + have ST: "x \ W \ x \ S \ x \ T" for x + using setdist_eq_0_closedin \S \ T \ {}\ assms + by (force simp: W_def setdist_sing_in_set) + have "S' \ T' = W" + by (auto simp: S'_def T'_def W_def) + then have cloUW: "closedin (subtopology euclidean U) W" + using closedin_Int US' UT' by blast + define r where "r \ \x. if x \ W then r0 x else x" + have "r ` (W \ S) \ S" "r ` (W \ T) \ T" + using \r0 ` W \ S \ T\ r_def by auto + have contr: "continuous_on (W \ (S \ T)) r" + unfolding r_def + proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) + show "closedin (subtopology euclidean (W \ (S \ T))) W" + using \S \ U\ \T \ U\ \W \ U\ \closedin (subtopology euclidean U) W\ closedin_subset_trans by fastforce + show "closedin (subtopology euclidean (W \ (S \ T))) (S \ T)" + by (meson \S \ U\ \T \ U\ \W \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) + show "\x. x \ W \ x \ W \ x \ S \ T \ x \ W \ r0 x = x" + by (auto simp: ST) + qed + have cloUWS: "closedin (subtopology euclidean U) (W \ S)" + by (simp add: cloUW assms closedin_Un) + obtain g where contg: "continuous_on U g" + and "g ` U \ S" and geqr: "\x. x \ W \ S \ g x = r x" + apply (rule AR_imp_absolute_extensor [OF \AR S\ _ _ cloUWS]) + apply (rule continuous_on_subset [OF contr]) + using \r ` (W \ S) \ S\ apply auto + done + have cloUWT: "closedin (subtopology euclidean U) (W \ T)" + by (simp add: cloUW assms closedin_Un) + obtain h where conth: "continuous_on U h" + and "h ` U \ T" and heqr: "\x. x \ W \ T \ h x = r x" + apply (rule AR_imp_absolute_extensor [OF \AR T\ _ _ cloUWT]) + apply (rule continuous_on_subset [OF contr]) + using \r ` (W \ T) \ T\ apply auto + done + have "U = S' \ T'" + by (force simp: S'_def T'_def) + then have cont: "continuous_on U (\x. if x \ S' then g x else h x)" + apply (rule ssubst) + apply (rule continuous_on_cases_local) + using US' UT' \S' \ T' = W\ \U = S' \ T'\ + contg conth continuous_on_subset geqr heqr apply auto + done + have UST: "(\x. if x \ S' then g x else h x) ` U \ S \ T" + using \g ` U \ S\ \h ` U \ T\ by auto + show ?thesis + apply (simp add: retract_of_def retraction_def \S \ U\ \T \ U\) + apply (rule_tac x="\x. if x \ S' then g x else h x" in exI) + apply (intro conjI cont UST) + by (metis IntI ST Un_iff \S \ S'\ \S' \ T' = W\ \T \ T'\ subsetD geqr heqr r0 r_def) +qed + + +lemma AR_closed_Un_local: + fixes S :: "'a::euclidean_space set" + assumes STS: "closedin (subtopology euclidean (S \ T)) S" + and STT: "closedin (subtopology euclidean (S \ T)) T" + and "AR S" "AR T" "AR(S \ T)" + shows "AR(S \ T)" +proof - + have "C retract_of U" + if hom: "S \ T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" + for U and C :: "('a * real) set" + proof - + obtain f g where hom: "homeomorphism (S \ T) C f g" + using hom by (force simp: homeomorphic_def) + have US: "closedin (subtopology euclidean U) (C \ g -` S)" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) + using hom homeomorphism_def apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have UT: "closedin (subtopology euclidean U) (C \ g -` T)" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) + using hom homeomorphism_def apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have ARS: "AR (C \ g -` S)" + apply (rule AR_homeomorphic_AR [OF \AR S\]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ART: "AR (C \ g -` T)" + apply (rule AR_homeomorphic_AR [OF \AR T\]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ARI: "AR ((C \ g -` S) \ (C \ g -` T))" + apply (rule AR_homeomorphic_AR [OF \AR (S \ T)\]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom + apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have "C = (C \ g -` S) \ (C \ g -` T)" + using hom by (auto simp: homeomorphism_def) + then show ?thesis + by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI]) + qed + then show ?thesis + by (force simp: AR_def) +qed + +corollary AR_closed_Un: + fixes S :: "'a::euclidean_space set" + shows "\closed S; closed T; AR S; AR T; AR (S \ T)\ \ AR (S \ T)" +by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV) + +text \ANRs closed under union\ + +lemma ANR_closed_Un_local_aux: + fixes U :: "'a::euclidean_space set" + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "ANR S" "ANR T" "ANR(S \ T)" + obtains V where "openin (subtopology euclidean U) V" "(S \ T) retract_of V" +proof (cases "S = {} \ T = {}") + case True with assms that show ?thesis + by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb) +next + case False + then have [simp]: "S \ {}" "T \ {}" by auto + have "S \ U" "T \ U" + using assms by (auto simp: closedin_imp_subset) + define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" + define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" + define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" + have cloUS': "closedin (subtopology euclidean U) S'" + using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] + by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) + have cloUT': "closedin (subtopology euclidean U) T'" + using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] + by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) + have "S \ S'" + using S'_def \S \ U\ setdist_sing_in_set by fastforce + have "T \ T'" + using T'_def \T \ U\ setdist_sing_in_set by fastforce + have "S' \ T' = U" + by (auto simp: S'_def T'_def) + have "W \ S'" + by (simp add: Collect_mono S'_def W_def) + have "W \ T'" + by (simp add: Collect_mono T'_def W_def) + have ST_W: "S \ T \ W" and "W \ U" + using \S \ U\ by (force simp: W_def setdist_sing_in_set)+ + have "S' \ T' = W" + by (auto simp: S'_def T'_def W_def) + then have cloUW: "closedin (subtopology euclidean U) W" + using closedin_Int cloUS' cloUT' by blast + obtain W' W0 where "openin (subtopology euclidean W) W'" + and cloWW0: "closedin (subtopology euclidean W) W0" + and "S \ T \ W'" "W' \ W0" + and ret: "(S \ T) retract_of W0" + apply (rule ANR_imp_closed_neighbourhood_retract [OF \ANR(S \ T)\]) + apply (rule closedin_subset_trans [of U, OF _ ST_W \W \ U\]) + apply (blast intro: assms)+ + done + then obtain U0 where opeUU0: "openin (subtopology euclidean U) U0" + and U0: "S \ T \ U0" "U0 \ W \ W0" + unfolding openin_open using \W \ U\ by blast + have "W0 \ U" + using \W \ U\ cloWW0 closedin_subset by fastforce + obtain r0 + where "S \ T \ W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \ S \ T" + and r0 [simp]: "\x. x \ S \ T \ r0 x = x" + using ret by (force simp: retract_of_def retraction_def) + have ST: "x \ W \ x \ S \ x \ T" for x + using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin) + define r where "r \ \x. if x \ W0 then r0 x else x" + have "r ` (W0 \ S) \ S" "r ` (W0 \ T) \ T" + using \r0 ` W0 \ S \ T\ r_def by auto + have contr: "continuous_on (W0 \ (S \ T)) r" + unfolding r_def + proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) + show "closedin (subtopology euclidean (W0 \ (S \ T))) W0" + apply (rule closedin_subset_trans [of U]) + using cloWW0 cloUW closedin_trans \W0 \ U\ \S \ U\ \T \ U\ apply blast+ + done + show "closedin (subtopology euclidean (W0 \ (S \ T))) (S \ T)" + by (meson \S \ U\ \T \ U\ \W0 \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) + show "\x. x \ W0 \ x \ W0 \ x \ S \ T \ x \ W0 \ r0 x = x" + using ST cloWW0 closedin_subset by fastforce + qed + have cloS'WS: "closedin (subtopology euclidean S') (W0 \ S)" + by (meson closedin_subset_trans US cloUS' \S \ S'\ \W \ S'\ cloUW cloWW0 + closedin_Un closedin_imp_subset closedin_trans) + obtain W1 g where "W0 \ S \ W1" and contg: "continuous_on W1 g" + and opeSW1: "openin (subtopology euclidean S') W1" + and "g ` W1 \ S" and geqr: "\x. x \ W0 \ S \ g x = r x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ _ \r ` (W0 \ S) \ S\ cloS'WS]) + apply (rule continuous_on_subset [OF contr], blast+) + done + have cloT'WT: "closedin (subtopology euclidean T') (W0 \ T)" + by (meson closedin_subset_trans UT cloUT' \T \ T'\ \W \ T'\ cloUW cloWW0 + closedin_Un closedin_imp_subset closedin_trans) + obtain W2 h where "W0 \ T \ W2" and conth: "continuous_on W2 h" + and opeSW2: "openin (subtopology euclidean T') W2" + and "h ` W2 \ T" and heqr: "\x. x \ W0 \ T \ h x = r x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ _ \r ` (W0 \ T) \ T\ cloT'WT]) + apply (rule continuous_on_subset [OF contr], blast+) + done + have "S' \ T' = W" + by (force simp: S'_def T'_def W_def) + obtain O1 O2 where "open O1" "W1 = S' \ O1" "open O2" "W2 = T' \ O2" + using opeSW1 opeSW2 by (force simp: openin_open) + show ?thesis + proof + have eq: "W1 - (W - U0) \ (W2 - (W - U0)) = + ((U - T') \ O1 \ (U - S') \ O2 \ U \ O1 \ O2) - (W - U0)" + using \U0 \ W \ W0\ \W0 \ S \ W1\ \W0 \ T \ W2\ + by (auto simp: \S' \ T' = U\ [symmetric] \S' \ T' = W\ [symmetric] \W1 = S' \ O1\ \W2 = T' \ O2\) + show "openin (subtopology euclidean U) (W1 - (W - U0) \ (W2 - (W - U0)))" + apply (subst eq) + apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \open O1\ \open O2\, simp_all) + done + have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \ (W2 - (W - U0)))) (W1 - (W - U0))" + using cloUS' apply (simp add: closedin_closed) + apply (erule ex_forward) + using U0 \W0 \ S \ W1\ + apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) + done + have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \ (W2 - (W - U0)))) (W2 - (W - U0))" + using cloUT' apply (simp add: closedin_closed) + apply (erule ex_forward) + using U0 \W0 \ T \ W2\ + apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) + done + have *: "\x\S \ T. (if x \ S' then g x else h x) = x" + using ST \S' \ T' = W\ cloT'WT closedin_subset geqr heqr + apply (auto simp: r_def, fastforce) + using \S \ S'\ \T \ T'\ \W0 \ S \ W1\ \W1 = S' \ O1\ by auto + have "\r. continuous_on (W1 - (W - U0) \ (W2 - (W - U0))) r \ + r ` (W1 - (W - U0) \ (W2 - (W - U0))) \ S \ T \ + (\x\S \ T. r x = x)" + apply (rule_tac x = "\x. if x \ S' then g x else h x" in exI) + apply (intro conjI *) + apply (rule continuous_on_cases_local + [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]]) + using \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = W\ + \g ` W1 \ S\ \h ` W2 \ T\ apply auto + using \U0 \ W \ W0\ \W0 \ S \ W1\ apply (fastforce simp add: geqr heqr)+ + done + then show "S \ T retract_of W1 - (W - U0) \ (W2 - (W - U0))" + using \W0 \ S \ W1\ \W0 \ T \ W2\ ST opeUU0 U0 + by (auto simp: retract_of_def retraction_def) + qed +qed + + +lemma ANR_closed_Un_local: + fixes S :: "'a::euclidean_space set" + assumes STS: "closedin (subtopology euclidean (S \ T)) S" + and STT: "closedin (subtopology euclidean (S \ T)) T" + and "ANR S" "ANR T" "ANR(S \ T)" + shows "ANR(S \ T)" +proof - + have "\T. openin (subtopology euclidean U) T \ C retract_of T" + if hom: "S \ T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" + for U and C :: "('a * real) set" + proof - + obtain f g where hom: "homeomorphism (S \ T) C f g" + using hom by (force simp: homeomorphic_def) + have US: "closedin (subtopology euclidean U) (C \ g -` S)" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) + using hom [unfolded homeomorphism_def] apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have UT: "closedin (subtopology euclidean U) (C \ g -` T)" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) + using hom [unfolded homeomorphism_def] apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have ANRS: "ANR (C \ g -` S)" + apply (rule ANR_homeomorphic_ANR [OF \ANR S\]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ANRT: "ANR (C \ g -` T)" + apply (rule ANR_homeomorphic_ANR [OF \ANR T\]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ANRI: "ANR ((C \ g -` S) \ (C \ g -` T))" + apply (rule ANR_homeomorphic_ANR [OF \ANR (S \ T)\]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom + apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have "C = (C \ g -` S) \ (C \ g -` T)" + using hom by (auto simp: homeomorphism_def) + then show ?thesis + by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI]) + qed + then show ?thesis + by (auto simp: ANR_def) +qed + +corollary ANR_closed_Un: + fixes S :: "'a::euclidean_space set" + shows "\closed S; closed T; ANR S; ANR T; ANR (S \ T)\ \ ANR (S \ T)" +by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int) + +lemma ANR_openin: + fixes S :: "'a::euclidean_space set" + assumes "ANR T" and opeTS: "openin (subtopology euclidean T) S" + shows "ANR S" +proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) + fix f :: "'a \ real \ 'a" and U C + assume contf: "continuous_on C f" and fim: "f ` C \ S" + and cloUC: "closedin (subtopology euclidean U) C" + have "f ` C \ T" + using fim opeTS openin_imp_subset by blast + obtain W g where "C \ W" + and UW: "openin (subtopology euclidean U) W" + and contg: "continuous_on W g" + and gim: "g ` W \ T" + and geq: "\x. x \ C \ g x = f x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf \f ` C \ T\ cloUC]) + using fim by auto + show "\V g. C \ V \ openin (subtopology euclidean U) V \ continuous_on V g \ g ` V \ S \ (\x\C. g x = f x)" + proof (intro exI conjI) + show "C \ W \ g -` S" + using \C \ W\ fim geq by blast + show "openin (subtopology euclidean U) (W \ g -` S)" + by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) + show "continuous_on (W \ g -` S) g" + by (blast intro: continuous_on_subset [OF contg]) + show "g ` (W \ g -` S) \ S" + using gim by blast + show "\x\C. g x = f x" + using geq by blast + qed +qed + +lemma ENR_openin: + fixes S :: "'a::euclidean_space set" + assumes "ENR T" and opeTS: "openin (subtopology euclidean T) S" + shows "ENR S" + using assms apply (simp add: ENR_ANR) + using ANR_openin locally_open_subset by blast + +lemma ANR_neighborhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ANR U" "S retract_of T" "openin (subtopology euclidean U) T" + shows "ANR S" + using ANR_openin ANR_retract_of_ANR assms by blast + +lemma ENR_neighborhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ENR U" "S retract_of T" "openin (subtopology euclidean U) T" + shows "ENR S" + using ENR_openin ENR_retract_of_ENR assms by blast + +lemma ANR_rel_interior: + fixes S :: "'a::euclidean_space set" + shows "ANR S \ ANR(rel_interior S)" + by (blast intro: ANR_openin openin_set_rel_interior) + +lemma ANR_delete: + fixes S :: "'a::euclidean_space set" + shows "ANR S \ ANR(S - {a})" + by (blast intro: ANR_openin openin_delete openin_subtopology_self) + +lemma ENR_rel_interior: + fixes S :: "'a::euclidean_space set" + shows "ENR S \ ENR(rel_interior S)" + by (blast intro: ENR_openin openin_set_rel_interior) + +lemma ENR_delete: + fixes S :: "'a::euclidean_space set" + shows "ENR S \ ENR(S - {a})" + by (blast intro: ENR_openin openin_delete openin_subtopology_self) + +lemma open_imp_ENR: "open S \ ENR S" + using ENR_def by blast + +lemma open_imp_ANR: + fixes S :: "'a::euclidean_space set" + shows "open S \ ANR S" + by (simp add: ENR_imp_ANR open_imp_ENR) + +lemma ANR_ball [iff]: + fixes a :: "'a::euclidean_space" + shows "ANR(ball a r)" + by (simp add: convex_imp_ANR) + +lemma ENR_ball [iff]: "ENR(ball a r)" + by (simp add: open_imp_ENR) + +lemma AR_ball [simp]: + fixes a :: "'a::euclidean_space" + shows "AR(ball a r) \ 0 < r" + by (auto simp: AR_ANR convex_imp_contractible) + +lemma ANR_cball [iff]: + fixes a :: "'a::euclidean_space" + shows "ANR(cball a r)" + by (simp add: convex_imp_ANR) + +lemma ENR_cball: + fixes a :: "'a::euclidean_space" + shows "ENR(cball a r)" + using ENR_convex_closed by blast + +lemma AR_cball [simp]: + fixes a :: "'a::euclidean_space" + shows "AR(cball a r) \ 0 \ r" + by (auto simp: AR_ANR convex_imp_contractible) + +lemma ANR_box [iff]: + fixes a :: "'a::euclidean_space" + shows "ANR(cbox a b)" "ANR(box a b)" + by (auto simp: convex_imp_ANR open_imp_ANR) + +lemma ENR_box [iff]: + fixes a :: "'a::euclidean_space" + shows "ENR(cbox a b)" "ENR(box a b)" +apply (simp add: ENR_convex_closed closed_cbox) +by (simp add: open_box open_imp_ENR) + +lemma AR_box [simp]: + "AR(cbox a b) \ cbox a b \ {}" "AR(box a b) \ box a b \ {}" + by (auto simp: AR_ANR convex_imp_contractible) + +lemma ANR_interior: + fixes S :: "'a::euclidean_space set" + shows "ANR(interior S)" + by (simp add: open_imp_ANR) + +lemma ENR_interior: + fixes S :: "'a::euclidean_space set" + shows "ENR(interior S)" + by (simp add: open_imp_ENR) + +lemma AR_imp_contractible: + fixes S :: "'a::euclidean_space set" + shows "AR S \ contractible S" + by (simp add: AR_ANR) + +lemma ENR_imp_locally_compact: + fixes S :: "'a::euclidean_space set" + shows "ENR S \ locally compact S" + by (simp add: ENR_ANR) + +lemma ANR_imp_locally_path_connected: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" + shows "locally path_connected S" +proof - + obtain U and T :: "('a \ real) set" + where "convex U" "U \ {}" + and UT: "closedin (subtopology euclidean U) T" + and "S homeomorphic T" + apply (rule homeomorphic_closedin_convex [of S]) + using aff_dim_le_DIM [of S] apply auto + done + then have "locally path_connected T" + by (meson ANR_imp_absolute_neighbourhood_retract + assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) + then have S: "locally path_connected S" + if "openin (subtopology euclidean U) V" "T retract_of V" "U \ {}" for V + using \S homeomorphic T\ homeomorphic_locally homeomorphic_path_connectedness by blast + show ?thesis + using assms + apply (clarsimp simp: ANR_def) + apply (drule_tac x=U in spec) + apply (drule_tac x=T in spec) + using \S homeomorphic T\ \U \ {}\ UT apply (blast intro: S) + done +qed + +lemma ANR_imp_locally_connected: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" + shows "locally connected S" +using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto + +lemma AR_imp_locally_path_connected: + fixes S :: "'a::euclidean_space set" + assumes "AR S" + shows "locally path_connected S" +by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms) + +lemma AR_imp_locally_connected: + fixes S :: "'a::euclidean_space set" + assumes "AR S" + shows "locally connected S" +using ANR_imp_locally_connected AR_ANR assms by blast + +lemma ENR_imp_locally_path_connected: + fixes S :: "'a::euclidean_space set" + assumes "ENR S" + shows "locally path_connected S" +by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms) + +lemma ENR_imp_locally_connected: + fixes S :: "'a::euclidean_space set" + assumes "ENR S" + shows "locally connected S" +using ANR_imp_locally_connected ENR_ANR assms by blast + +lemma ANR_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "ANR S" "ANR T" shows "ANR(S \ T)" +proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) + fix f :: " ('a \ 'b) \ real \ 'a \ 'b" and U C + assume "continuous_on C f" and fim: "f ` C \ S \ T" + and cloUC: "closedin (subtopology euclidean U) C" + have contf1: "continuous_on C (fst \ f)" + by (simp add: \continuous_on C f\ continuous_on_fst) + obtain W1 g where "C \ W1" + and UW1: "openin (subtopology euclidean U) W1" + and contg: "continuous_on W1 g" + and gim: "g ` W1 \ S" + and geq: "\x. x \ C \ g x = (fst \ f) x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ contf1 _ cloUC]) + using fim apply auto + done + have contf2: "continuous_on C (snd \ f)" + by (simp add: \continuous_on C f\ continuous_on_snd) + obtain W2 h where "C \ W2" + and UW2: "openin (subtopology euclidean U) W2" + and conth: "continuous_on W2 h" + and him: "h ` W2 \ T" + and heq: "\x. x \ C \ h x = (snd \ f) x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf2 _ cloUC]) + using fim apply auto + done + show "\V g. C \ V \ + openin (subtopology euclidean U) V \ + continuous_on V g \ g ` V \ S \ T \ (\x\C. g x = f x)" + proof (intro exI conjI) + show "C \ W1 \ W2" + by (simp add: \C \ W1\ \C \ W2\) + show "openin (subtopology euclidean U) (W1 \ W2)" + by (simp add: UW1 UW2 openin_Int) + show "continuous_on (W1 \ W2) (\x. (g x, h x))" + by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) + show "(\x. (g x, h x)) ` (W1 \ W2) \ S \ T" + using gim him by blast + show "(\x\C. (g x, h x) = f x)" + using geq heq by auto + qed +qed + +lemma AR_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "AR S" "AR T" shows "AR(S \ T)" +using assms by (simp add: AR_ANR ANR_Times contractible_Times) + +subsection \Kuhn Simplices\ + lemma bij_betw_singleton_eq: assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \ A" assumes eq: "(\x. x \ A \ x \ a \ f x = g x)" @@ -59,6 +1795,7 @@ using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+ qed auto +(* FIXME mv *) lemma brouwer_compactness_lemma: fixes f :: "'a::metric_space \ 'b::real_normed_vector" assumes "compact s" @@ -107,7 +1844,7 @@ qed -subsection \The key "counting" observation, somewhat abstracted\ +subsubsection \The key "counting" observation, somewhat abstracted\ lemma kuhn_counting_lemma: fixes bnd compo compo' face S F @@ -138,7 +1875,7 @@ by auto qed -subsection \The odd/even result for faces of complete vertices, generalized\ +subsubsection \The odd/even result for faces of complete vertices, generalized\ lemma kuhn_complete_lemma: assumes [simp]: "finite simplices" @@ -1102,7 +2839,7 @@ by (subst (asm) eq_commute) auto } qed -subsection \Reduced labelling\ +subsubsection \Reduced labelling\ definition reduced :: "nat \ (nat \ nat) \ nat" where "reduced n x = (LEAST k. k = n \ x k \ 0)" @@ -1348,7 +3085,7 @@ qed qed -subsection \The main result for the unit cube\ +subsubsection \Main result for the unit cube\ lemma kuhn_labelling_lemma': assumes "(\x::nat\real. P x \ P (f x))" @@ -1387,32 +3124,9 @@ qed qed -definition unit_cube :: "'a::euclidean_space set" - where "unit_cube = {x. \i\Basis. 0 \ x \ i \ x \ i \ 1}" - -lemma mem_unit_cube: "x \ unit_cube \ (\i\Basis. 0 \ x \ i \ x \ i \ 1)" - unfolding unit_cube_def by simp - -lemma bounded_unit_cube: "bounded unit_cube" - unfolding bounded_def -proof (intro exI ballI) - fix y :: 'a assume y: "y \ unit_cube" - have "dist 0 y = norm y" by (rule dist_0_norm) - also have "\ = norm (\i\Basis. (y \ i) *\<^sub>R i)" unfolding euclidean_representation .. - also have "\ \ (\i\Basis. norm ((y \ i) *\<^sub>R i))" by (rule norm_sum) - also have "\ \ (\i::'a\Basis. 1)" - by (rule sum_mono, simp add: y [unfolded mem_unit_cube]) - finally show "dist 0 y \ (\i::'a\Basis. 1)" . -qed - -lemma closed_unit_cube: "closed unit_cube" - unfolding unit_cube_def Collect_ball_eq Collect_conj_eq - by (rule closed_INT, auto intro!: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) - -lemma compact_unit_cube: "compact unit_cube" (is "compact ?C") - unfolding compact_eq_seq_compact_metric - using bounded_unit_cube closed_unit_cube - by (rule bounded_closed_imp_seq_compact) +subsection \Brouwer's fixed point theorem\ + +text \We start proving Brouwer's fixed point theorem for unit cubes.\ lemma brouwer_cube: fixes f :: "'a::euclidean_space \ 'a" @@ -1685,94 +3399,7 @@ using i by auto qed - -subsection \Retractions\ - -definition "retraction S T r \ T \ S \ continuous_on S r \ r ` S \ T \ (\x\T. r x = x)" - -definition retract_of (infixl "retract'_of" 50) - where "(T retract_of S) \ (\r. retraction S T r)" - -lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" - unfolding retraction_def by auto - -subsection \Preservation of fixpoints under (more general notion of) retraction\ - -lemma invertible_fixpoint_property: - fixes S :: "'a::euclidean_space set" - and T :: "'b::euclidean_space set" - assumes contt: "continuous_on T i" - and "i ` T \ S" - and contr: "continuous_on S r" - and "r ` S \ T" - and ri: "\y. y \ T \ r (i y) = y" - and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" - and contg: "continuous_on T g" - and "g ` T \ T" - obtains y where "y \ T" and "g y = y" -proof - - have "\x\S. (i \ g \ r) x = x" - proof (rule FP) - show "continuous_on S (i \ g \ r)" - by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) - show "(i \ g \ r) ` S \ S" - using assms(2,4,8) by force - qed - then obtain x where x: "x \ S" "(i \ g \ r) x = x" .. - then have *: "g (r x) \ T" - using assms(4,8) by auto - have "r ((i \ g \ r) x) = r x" - using x by auto - then show ?thesis - using "*" ri that by auto -qed - -lemma homeomorphic_fixpoint_property: - fixes S :: "'a::euclidean_space set" - and T :: "'b::euclidean_space set" - assumes "S homeomorphic T" - shows "(\f. continuous_on S f \ f ` S \ S \ (\x\S. f x = x)) \ - (\g. continuous_on T g \ g ` T \ T \ (\y\T. g y = y))" - (is "?lhs = ?rhs") -proof - - obtain r i where r: - "\x\S. i (r x) = x" "r ` S = T" "continuous_on S r" - "\y\T. r (i y) = y" "i ` T = S" "continuous_on T i" - using assms unfolding homeomorphic_def homeomorphism_def by blast - show ?thesis - proof - assume ?lhs - with r show ?rhs - by (metis invertible_fixpoint_property[of T i S r] order_refl) - next - assume ?rhs - with r show ?lhs - by (metis invertible_fixpoint_property[of S r T i] order_refl) - qed -qed - -lemma retract_fixpoint_property: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - and S :: "'a set" - assumes "T retract_of S" - and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" - and contg: "continuous_on T g" - and "g ` T \ T" - obtains y where "y \ T" and "g y = y" -proof - - obtain h where "retraction S T h" - using assms(1) unfolding retract_of_def .. - then show ?thesis - unfolding retraction_def - using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] - by (metis assms(4) contg image_ident that) -qed - - -subsection \The Brouwer theorem for any set with nonempty interior\ - -lemma convex_unit_cube: "convex unit_cube" - by (rule is_interval_convex) (fastforce simp add: is_interval_def mem_unit_cube) +text \Next step is to prove it for nonempty interiors.\ lemma brouwer_weak: fixes f :: "'a::euclidean_space \ 'a" @@ -1806,8 +3433,7 @@ by (auto intro: that) qed - -text \And in particular for a closed ball.\ +text \Then the particular case for closed balls.\ lemma brouwer_ball: fixes f :: "'a::euclidean_space \ 'a" @@ -1819,11 +3445,9 @@ unfolding interior_cball ball_eq_empty using assms by auto -text \Still more general form; could derive this directly without using the - rather involved \HOMEOMORPHIC_CONVEX_COMPACT\ theorem, just using - a scaling and translation to put the set inside the unit cube.\ - -lemma brouwer: +text \And finally we prove Brouwer's fixed point theorem in its general version.\ + +theorem brouwer: fixes f :: "'a::euclidean_space \ 'a" assumes S: "compact S" "convex S" "S \ {}" and contf: "continuous_on S f" @@ -1858,9 +3482,11 @@ qed qed +subsection \Applications\ + text \So we get the no-retraction theorem.\ -theorem no_retraction_cball: +corollary no_retraction_cball: fixes a :: "'a::euclidean_space" assumes "e > 0" shows "\ (frontier (cball a e) retract_of (cball a e))" @@ -1900,7 +3526,7 @@ using continuous_on_const less_eq_real_def by auto qed -lemma connected_sphere_eq: +corollary connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") @@ -1934,7 +3560,7 @@ qed qed -lemma path_connected_sphere_eq: +corollary path_connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "path_connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") @@ -1998,309 +3624,9 @@ ultimately show False by simp qed -subsection\More Properties of Retractions\ - -lemma retraction: - "retraction S T r \ - T \ S \ continuous_on S r \ r ` S = T \ (\x \ T. r x = x)" -by (force simp: retraction_def) - -lemma retract_of_imp_extensible: - assumes "S retract_of T" and "continuous_on S f" and "f ` S \ U" - obtains g where "continuous_on T g" "g ` T \ U" "\x. x \ S \ g x = f x" -using assms -apply (clarsimp simp add: retract_of_def retraction) -apply (rule_tac g = "f \ r" in that) -apply (auto simp: continuous_on_compose2) -done - -lemma idempotent_imp_retraction: - assumes "continuous_on S f" and "f ` S \ S" and "\x. x \ S \ f(f x) = f x" - shows "retraction S (f ` S) f" -by (simp add: assms retraction) - -lemma retraction_subset: - assumes "retraction S T r" and "T \ s'" and "s' \ S" - shows "retraction s' T r" - unfolding retraction_def - by (metis assms continuous_on_subset image_mono retraction) - -lemma retract_of_subset: - assumes "T retract_of S" and "T \ s'" and "s' \ S" - shows "T retract_of s'" -by (meson assms retract_of_def retraction_subset) - -lemma retraction_refl [simp]: "retraction S S (\x. x)" -by (simp add: continuous_on_id retraction) - -lemma retract_of_refl [iff]: "S retract_of S" - unfolding retract_of_def retraction_def - using continuous_on_id by blast - -lemma retract_of_imp_subset: - "S retract_of T \ S \ T" -by (simp add: retract_of_def retraction_def) - -lemma retract_of_empty [simp]: - "({} retract_of S) \ S = {}" "(S retract_of {}) \ S = {}" -by (auto simp: retract_of_def retraction_def) - -lemma retract_of_singleton [iff]: "({x} retract_of S) \ x \ S" - unfolding retract_of_def retraction_def by force - -lemma retraction_comp: - "\retraction S T f; retraction T U g\ - \ retraction S U (g \ f)" -apply (auto simp: retraction_def intro: continuous_on_compose2) -by blast - -lemma retract_of_trans [trans]: - assumes "S retract_of T" and "T retract_of U" - shows "S retract_of U" -using assms by (auto simp: retract_of_def intro: retraction_comp) - -lemma closedin_retract: - fixes S :: "'a :: real_normed_vector set" - assumes "S retract_of T" - shows "closedin (subtopology euclidean T) S" -proof - - obtain r where "S \ T" "continuous_on T r" "r ` T \ S" "\x. x \ S \ r x = x" - using assms by (auto simp: retract_of_def retraction_def) - then have S: "S = {x \ T. (norm(r x - x)) = 0}" by auto - show ?thesis - apply (subst S) - apply (rule continuous_closedin_preimage_constant) - by (simp add: \continuous_on T r\ continuous_on_diff continuous_on_id continuous_on_norm) -qed - -lemma closedin_self [simp]: - fixes S :: "'a :: real_normed_vector set" - shows "closedin (subtopology euclidean S) S" - by (simp add: closedin_retract) - -lemma retract_of_contractible: - assumes "contractible T" "S retract_of T" - shows "contractible S" -using assms -apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with) -apply (rule_tac x="r a" in exI) -apply (rule_tac x="r \ h" in exI) -apply (intro conjI continuous_intros continuous_on_compose) -apply (erule continuous_on_subset | force)+ -done - -lemma retract_of_compact: - "\compact T; S retract_of T\ \ compact S" - by (metis compact_continuous_image retract_of_def retraction) - -lemma retract_of_closed: - fixes S :: "'a :: real_normed_vector set" - shows "\closed T; S retract_of T\ \ closed S" - by (metis closedin_retract closedin_closed_eq) - -lemma retract_of_connected: - "\connected T; S retract_of T\ \ connected S" - by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) - -lemma retract_of_path_connected: - "\path_connected T; S retract_of T\ \ path_connected S" - by (metis path_connected_continuous_image retract_of_def retraction) - -lemma retract_of_simply_connected: - "\simply_connected T; S retract_of T\ \ simply_connected S" -apply (simp add: retract_of_def retraction_def, clarify) -apply (rule simply_connected_retraction_gen) -apply (force simp: continuous_on_id elim!: continuous_on_subset)+ -done - -lemma retract_of_homotopically_trivial: - assumes ts: "T retract_of S" - and hom: "\f g. \continuous_on U f; f ` U \ S; - continuous_on U g; g ` U \ S\ - \ homotopic_with (\x. True) U S f g" - and "continuous_on U f" "f ` U \ T" - and "continuous_on U g" "g ` U \ T" - shows "homotopic_with (\x. True) U T f g" -proof - - obtain r where "r ` S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" - using ts by (auto simp: retract_of_def retraction) - then obtain k where "Retracts S r T k" - unfolding Retracts_def - by (metis continuous_on_subset dual_order.trans image_iff image_mono) - then show ?thesis - apply (rule Retracts.homotopically_trivial_retraction_gen) - using assms - apply (force simp: hom)+ - done -qed - -lemma retract_of_homotopically_trivial_null: - assumes ts: "T retract_of S" - and hom: "\f. \continuous_on U f; f ` U \ S\ - \ \c. homotopic_with (\x. True) U S f (\x. c)" - and "continuous_on U f" "f ` U \ T" - obtains c where "homotopic_with (\x. True) U T f (\x. c)" -proof - - obtain r where "r ` S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" - using ts by (auto simp: retract_of_def retraction) - then obtain k where "Retracts S r T k" - unfolding Retracts_def - by (metis continuous_on_subset dual_order.trans image_iff image_mono) - then show ?thesis - apply (rule Retracts.homotopically_trivial_retraction_null_gen) - apply (rule TrueI refl assms that | assumption)+ - done -qed - -lemma retraction_imp_quotient_map: - "retraction S T r - \ U \ T - \ (openin (subtopology euclidean S) (S \ r -` U) \ - openin (subtopology euclidean T) U)" -apply (clarsimp simp add: retraction) -apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) -apply (auto simp: elim: continuous_on_subset) -done - -lemma retract_of_locally_compact: - fixes S :: "'a :: {heine_borel,real_normed_vector} set" - shows "\ locally compact S; T retract_of S\ \ locally compact T" - by (metis locally_compact_closedin closedin_retract) - -lemma retract_of_Times: - "\S retract_of s'; T retract_of t'\ \ (S \ T) retract_of (s' \ t')" -apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) -apply (rename_tac f g) -apply (rule_tac x="\z. ((f \ fst) z, (g \ snd) z)" in exI) -apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ -done - -lemma homotopic_into_retract: - "\f ` S \ T; g ` S \ T; T retract_of U; homotopic_with (\x. True) S U f g\ - \ homotopic_with (\x. True) S T f g" -apply (subst (asm) homotopic_with_def) -apply (simp add: homotopic_with retract_of_def retraction_def, clarify) -apply (rule_tac x="r \ h" in exI) -apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ -done - -lemma retract_of_locally_connected: - assumes "locally connected T" "S retract_of T" - shows "locally connected S" - using assms - by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_connected_quotient_image) - -lemma retract_of_locally_path_connected: - assumes "locally path_connected T" "S retract_of T" - shows "locally path_connected S" - using assms - by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image) - -subsubsection\A few simple lemmas about deformation retracts\ - -lemma deformation_retract_imp_homotopy_eqv: - fixes S :: "'a::euclidean_space set" - assumes "homotopic_with (\x. True) S S id r" and r: "retraction S T r" - shows "S homotopy_eqv T" -proof - - have "homotopic_with (\x. True) S S (id \ r) id" - by (simp add: assms(1) homotopic_with_symD) - moreover have "homotopic_with (\x. True) T T (r \ id) id" - using r unfolding retraction_def - by (metis (no_types, lifting) comp_id continuous_on_id' homotopic_with_equal homotopic_with_symD id_def image_id order_refl) - ultimately - show ?thesis - unfolding homotopy_eqv_def - by (metis continuous_on_id' id_def image_id r retraction_def) -qed - -lemma deformation_retract: - fixes S :: "'a::euclidean_space set" - shows "(\r. homotopic_with (\x. True) S S id r \ retraction S T r) \ - T retract_of S \ (\f. homotopic_with (\x. True) S S id f \ f ` S \ T)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - by (auto simp: retract_of_def retraction_def) -next - assume ?rhs - then show ?lhs - apply (clarsimp simp add: retract_of_def retraction_def) - apply (rule_tac x=r in exI, simp) - apply (rule homotopic_with_trans, assumption) - apply (rule_tac f = "r \ f" and g="r \ id" in homotopic_with_eq) - apply (rule_tac Y=S in homotopic_compose_continuous_left) - apply (auto simp: homotopic_with_sym) - done -qed - -lemma deformation_retract_of_contractible_sing: - fixes S :: "'a::euclidean_space set" - assumes "contractible S" "a \ S" - obtains r where "homotopic_with (\x. True) S S id r" "retraction S {a} r" -proof - - have "{a} retract_of S" - by (simp add: \a \ S\) - moreover have "homotopic_with (\x. True) S S id (\x. a)" - using assms - by (auto simp: contractible_def continuous_on_const continuous_on_id homotopic_into_contractible image_subset_iff) - moreover have "(\x. a) ` S \ {a}" - by (simp add: image_subsetI) - ultimately show ?thesis - using that deformation_retract by metis -qed - - -subsection\Punctured affine hulls, etc\ - -lemma continuous_on_compact_surface_projection_aux: - fixes S :: "'a::t2_space set" - assumes "compact S" "S \ T" "image q T \ S" - and contp: "continuous_on T p" - and "\x. x \ S \ q x = x" - and [simp]: "\x. x \ T \ q(p x) = q x" - and "\x. x \ T \ p(q x) = p x" - shows "continuous_on T q" -proof - - have *: "image p T = image p S" - using assms by auto (metis imageI subset_iff) - have contp': "continuous_on S p" - by (rule continuous_on_subset [OF contp \S \ T\]) - have "continuous_on (p ` T) q" - by (simp add: "*" assms(1) assms(2) assms(5) continuous_on_inv contp' rev_subsetD) - then have "continuous_on T (q \ p)" - by (rule continuous_on_compose [OF contp]) - then show ?thesis - by (rule continuous_on_eq [of _ "q \ p"]) (simp add: o_def) -qed - -lemma continuous_on_compact_surface_projection: - fixes S :: "'a::real_normed_vector set" - assumes "compact S" - and S: "S \ V - {0}" and "cone V" - and iff: "\x k. x \ V - {0} \ 0 < k \ (k *\<^sub>R x) \ S \ d x = k" - shows "continuous_on (V - {0}) (\x. d x *\<^sub>R x)" -proof (rule continuous_on_compact_surface_projection_aux [OF \compact S\ S]) - show "(\x. d x *\<^sub>R x) ` (V - {0}) \ S" - using iff by auto - show "continuous_on (V - {0}) (\x. inverse(norm x) *\<^sub>R x)" - by (intro continuous_intros) force - show "\x. x \ S \ d x *\<^sub>R x = x" - by (metis S zero_less_one local.iff scaleR_one subset_eq) - show "d (x /\<^sub>R norm x) *\<^sub>R (x /\<^sub>R norm x) = d x *\<^sub>R x" if "x \ V - {0}" for x - using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \cone V\ - by (simp add: field_simps cone_def zero_less_mult_iff) - show "d x *\<^sub>R x /\<^sub>R norm (d x *\<^sub>R x) = x /\<^sub>R norm x" if "x \ V - {0}" for x - proof - - have "0 < d x" - using local.iff that by blast - then show ?thesis - by simp - qed -qed - -proposition rel_frontier_deformation_retract_of_punctured_convex: +subsubsection \Punctured affine hulls, etc\ + +lemma rel_frontier_deformation_retract_of_punctured_convex: fixes S :: "'a::euclidean_space set" assumes "convex S" "convex T" "bounded S" and arelS: "a \ rel_interior S" @@ -2502,7 +3828,7 @@ shows "connected(rel_frontier S)" by (simp add: assms path_connected_imp_connected path_connected_sphere_gen) -subsection\Borsuk-style characterization of separation\ +subsubsection\Borsuk-style characterization of separation\ lemma continuous_on_Borsuk_map: "a \ s \ continuous_on s (\x. inverse(norm (x - a)) *\<^sub>R (x - a))" @@ -2595,1327 +3921,7 @@ by blast qed -subsection\Absolute retracts, etc\ - -text\Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also - Euclidean neighbourhood retracts (ENR). We define AR and ANR by - specializing the standard definitions for a set to embedding in -spaces of higher dimension. \ - -(*This turns out to be sufficient (since any set in -R^n can be embedded as a closed subset of a convex subset of R^{n+1}) to -derive the usual 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. -- JRH*) - -definition AR :: "'a::topological_space set => bool" - where - "AR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (subtopology euclidean U) S' - \ S' retract_of U" - -definition ANR :: "'a::topological_space set => bool" - where - "ANR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (subtopology euclidean U) S' - \ (\T. openin (subtopology euclidean U) T \ - S' retract_of T)" - -definition 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.\ - -proposition AR_imp_absolute_extensor: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "AR S" and contf: "continuous_on T f" and "f ` T \ S" - and cloUT: "closedin (subtopology euclidean U) T" - obtains g where "continuous_on U g" "g ` U \ S" "\x. x \ T \ g x = f x" -proof - - have "aff_dim S < int (DIM('b \ real))" - using aff_dim_le_DIM [of S] by simp - then obtain C and S' :: "('b * real) set" - where C: "convex C" "C \ {}" - and cloCS: "closedin (subtopology euclidean C) S'" - and hom: "S homeomorphic S'" - by (metis that homeomorphic_closedin_convex) - then have "S' retract_of C" - using \AR S\ by (simp add: AR_def) - then obtain r where "S' \ C" and contr: "continuous_on C r" - and "r ` C \ S'" and rid: "\x. x\S' \ r x = x" - by (auto simp: retraction_def retract_of_def) - obtain g h where "homeomorphism S S' g h" - using hom by (force simp: homeomorphic_def) - then have "continuous_on (f ` T) g" - by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def) - then have contgf: "continuous_on T (g \ f)" - by (metis continuous_on_compose contf) - have gfTC: "(g \ f) ` T \ C" - proof - - have "g ` S = S'" - by (metis (no_types) \homeomorphism S S' g h\ homeomorphism_def) - with \S' \ C\ \f ` T \ S\ show ?thesis by force - qed - obtain f' where f': "continuous_on U f'" "f' ` U \ C" - "\x. x \ T \ f' x = (g \ f) x" - by (metis Dugundji [OF C cloUT contgf gfTC]) - show ?thesis - proof (rule_tac g = "h \ r \ f'" in that) - show "continuous_on U (h \ r \ f')" - apply (intro continuous_on_compose f') - using continuous_on_subset contr f' apply blast - by (meson \homeomorphism S S' g h\ \r ` C \ S'\ continuous_on_subset \f' ` U \ C\ homeomorphism_def image_mono) - show "(h \ r \ f') ` U \ S" - using \homeomorphism S S' g h\ \r ` C \ S'\ \f' ` U \ C\ - by (fastforce simp: homeomorphism_def) - show "\x. x \ T \ (h \ r \ f') x = f x" - using \homeomorphism S S' g h\ \f ` T \ S\ f' - by (auto simp: rid homeomorphism_def) - qed -qed - -lemma AR_imp_absolute_retract: - fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "AR S" "S homeomorphic S'" - and clo: "closedin (subtopology euclidean U) S'" - shows "S' retract_of U" -proof - - obtain g h where hom: "homeomorphism S S' g h" - using assms by (force simp: homeomorphic_def) - have h: "continuous_on S' h" " h ` S' \ S" - using hom homeomorphism_def apply blast - apply (metis hom equalityE homeomorphism_def) - done - obtain h' where h': "continuous_on U h'" "h' ` U \ S" - and h'h: "\x. x \ S' \ h' x = h x" - by (blast intro: AR_imp_absolute_extensor [OF \AR S\ h clo]) - have [simp]: "S' \ U" using clo closedin_limpt by blast - show ?thesis - proof (simp add: retraction_def retract_of_def, intro exI conjI) - show "continuous_on U (g \ h')" - apply (intro continuous_on_compose h') - apply (meson hom continuous_on_subset h' homeomorphism_cont1) - done - show "(g \ h') ` U \ S'" - using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) - show "\x\S'. (g \ h') x = x" - by clarsimp (metis h'h hom homeomorphism_def) - qed -qed - -lemma AR_imp_absolute_retract_UNIV: - fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "AR S" and hom: "S homeomorphic S'" - and clo: "closed S'" - shows "S' retract_of UNIV" -apply (rule AR_imp_absolute_retract [OF \AR S\ hom]) -using clo closed_closedin by auto - -lemma absolute_extensor_imp_AR: - fixes S :: "'a::euclidean_space set" - assumes "\f :: 'a * real \ 'a. - \U T. \continuous_on T f; f ` T \ S; - closedin (subtopology euclidean U) T\ - \ \g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)" - shows "AR S" -proof (clarsimp simp: AR_def) - fix U and T :: "('a * real) set" - assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" - then obtain g h where hom: "homeomorphism S T g h" - by (force simp: homeomorphic_def) - have h: "continuous_on T h" " h ` T \ S" - using hom homeomorphism_def apply blast - apply (metis hom equalityE homeomorphism_def) - done - obtain h' where h': "continuous_on U h'" "h' ` U \ S" - and h'h: "\x\T. h' x = h x" - using assms [OF h clo] by blast - have [simp]: "T \ U" - using clo closedin_imp_subset by auto - show "T retract_of U" - proof (simp add: retraction_def retract_of_def, intro exI conjI) - show "continuous_on U (g \ h')" - apply (intro continuous_on_compose h') - apply (meson hom continuous_on_subset h' homeomorphism_cont1) - done - show "(g \ h') ` U \ T" - using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) - show "\x\T. (g \ h') x = x" - by clarsimp (metis h'h hom homeomorphism_def) - qed -qed - -lemma AR_eq_absolute_extensor: - fixes S :: "'a::euclidean_space set" - shows "AR S \ - (\f :: 'a * real \ 'a. - \U T. continuous_on T f \ f ` T \ S \ - closedin (subtopology euclidean U) T \ - (\g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)))" -apply (rule iffI) - apply (metis AR_imp_absolute_extensor) -apply (simp add: absolute_extensor_imp_AR) -done - -lemma AR_imp_retract: - fixes S :: "'a::euclidean_space set" - assumes "AR S \ closedin (subtopology euclidean U) S" - shows "S retract_of U" -using AR_imp_absolute_retract assms homeomorphic_refl by blast - -lemma AR_homeomorphic_AR: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "AR T" "S homeomorphic T" - shows "AR S" -unfolding AR_def -by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym) - -lemma homeomorphic_AR_iff_AR: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - shows "S homeomorphic T \ AR S \ AR T" -by (metis AR_homeomorphic_AR homeomorphic_sym) - - -proposition ANR_imp_absolute_neighbourhood_extensor: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "ANR S" and contf: "continuous_on T f" and "f ` T \ S" - and cloUT: "closedin (subtopology euclidean U) T" - obtains V g where "T \ V" "openin (subtopology euclidean U) V" - "continuous_on V g" - "g ` V \ S" "\x. x \ T \ g x = f x" -proof - - have "aff_dim S < int (DIM('b \ real))" - using aff_dim_le_DIM [of S] by simp - then obtain C and S' :: "('b * real) set" - where C: "convex C" "C \ {}" - and cloCS: "closedin (subtopology euclidean C) S'" - and hom: "S homeomorphic S'" - by (metis that homeomorphic_closedin_convex) - then obtain D where opD: "openin (subtopology euclidean C) D" and "S' retract_of D" - using \ANR S\ by (auto simp: ANR_def) - then obtain r where "S' \ D" and contr: "continuous_on D r" - and "r ` D \ S'" and rid: "\x. x \ S' \ r x = x" - by (auto simp: retraction_def retract_of_def) - obtain g h where homgh: "homeomorphism S S' g h" - using hom by (force simp: homeomorphic_def) - have "continuous_on (f ` T) g" - by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def homgh) - then have contgf: "continuous_on T (g \ f)" - by (intro continuous_on_compose contf) - have gfTC: "(g \ f) ` T \ C" - proof - - have "g ` S = S'" - by (metis (no_types) homeomorphism_def homgh) - then show ?thesis - by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology) - qed - obtain f' where contf': "continuous_on U f'" - and "f' ` U \ C" - and eq: "\x. x \ T \ f' x = (g \ f) x" - by (metis Dugundji [OF C cloUT contgf gfTC]) - show ?thesis - proof (rule_tac V = "U \ f' -` D" and g = "h \ r \ f'" in that) - show "T \ U \ f' -` D" - using cloUT closedin_imp_subset \S' \ D\ \f ` T \ S\ eq homeomorphism_image1 homgh - by fastforce - show ope: "openin (subtopology euclidean U) (U \ f' -` D)" - using \f' ` U \ C\ by (auto simp: opD contf' continuous_openin_preimage) - have conth: "continuous_on (r ` f' ` (U \ f' -` D)) h" - apply (rule continuous_on_subset [of S']) - using homeomorphism_def homgh apply blast - using \r ` D \ S'\ by blast - show "continuous_on (U \ f' -` D) (h \ r \ f')" - apply (intro continuous_on_compose conth - continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto) - done - show "(h \ r \ f') ` (U \ f' -` D) \ S" - using \homeomorphism S S' g h\ \f' ` U \ C\ \r ` D \ S'\ - by (auto simp: homeomorphism_def) - show "\x. x \ T \ (h \ r \ f') x = f x" - using \homeomorphism S S' g h\ \f ` T \ S\ eq - by (auto simp: rid homeomorphism_def) - qed -qed - - -corollary ANR_imp_absolute_neighbourhood_retract: - fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "ANR S" "S homeomorphic S'" - and clo: "closedin (subtopology euclidean U) S'" - obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" -proof - - obtain g h where hom: "homeomorphism S S' g h" - using assms by (force simp: homeomorphic_def) - have h: "continuous_on S' h" " h ` S' \ S" - using hom homeomorphism_def apply blast - apply (metis hom equalityE homeomorphism_def) - done - from ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo] - obtain V h' where "S' \ V" and opUV: "openin (subtopology euclidean U) V" - and h': "continuous_on V h'" "h' ` V \ S" - and h'h:"\x. x \ S' \ h' x = h x" - by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo]) - have "S' retract_of V" - proof (simp add: retraction_def retract_of_def, intro exI conjI \S' \ V\) - show "continuous_on V (g \ h')" - apply (intro continuous_on_compose h') - apply (meson hom continuous_on_subset h' homeomorphism_cont1) - done - show "(g \ h') ` V \ S'" - using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) - show "\x\S'. (g \ h') x = x" - by clarsimp (metis h'h hom homeomorphism_def) - qed - then show ?thesis - by (rule that [OF opUV]) -qed - -corollary ANR_imp_absolute_neighbourhood_retract_UNIV: - fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'" - obtains V where "open V" "S' retract_of V" - using ANR_imp_absolute_neighbourhood_retract [OF \ANR S\ hom] -by (metis clo closed_closedin open_openin subtopology_UNIV) - -corollary neighbourhood_extension_into_ANR: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes contf: "continuous_on S f" and fim: "f ` S \ T" and "ANR T" "closed S" - obtains V g where "S \ V" "open V" "continuous_on V g" - "g ` V \ T" "\x. x \ S \ g x = f x" - using ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf fim] - by (metis \closed S\ closed_closedin open_openin subtopology_UNIV) - -lemma absolute_neighbourhood_extensor_imp_ANR: - fixes S :: "'a::euclidean_space set" - assumes "\f :: 'a * real \ 'a. - \U T. \continuous_on T f; f ` T \ S; - closedin (subtopology euclidean U) T\ - \ \V g. T \ V \ openin (subtopology euclidean U) V \ - continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)" - shows "ANR S" -proof (clarsimp simp: ANR_def) - fix U and T :: "('a * real) set" - assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" - then obtain g h where hom: "homeomorphism S T g h" - by (force simp: homeomorphic_def) - have h: "continuous_on T h" " h ` T \ S" - using hom homeomorphism_def apply blast - apply (metis hom equalityE homeomorphism_def) - done - obtain V h' where "T \ V" and opV: "openin (subtopology euclidean U) V" - and h': "continuous_on V h'" "h' ` V \ S" - and h'h: "\x\T. h' x = h x" - using assms [OF h clo] by blast - have [simp]: "T \ U" - using clo closedin_imp_subset by auto - have "T retract_of V" - proof (simp add: retraction_def retract_of_def, intro exI conjI \T \ V\) - show "continuous_on V (g \ h')" - apply (intro continuous_on_compose h') - apply (meson hom continuous_on_subset h' homeomorphism_cont1) - done - show "(g \ h') ` V \ T" - using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) - show "\x\T. (g \ h') x = x" - by clarsimp (metis h'h hom homeomorphism_def) - qed - then show "\V. openin (subtopology euclidean U) V \ T retract_of V" - using opV by blast -qed - -lemma ANR_eq_absolute_neighbourhood_extensor: - fixes S :: "'a::euclidean_space set" - shows "ANR S \ - (\f :: 'a * real \ 'a. - \U T. continuous_on T f \ f ` T \ S \ - closedin (subtopology euclidean U) T \ - (\V g. T \ V \ openin (subtopology euclidean U) V \ - continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)))" -apply (rule iffI) - apply (metis ANR_imp_absolute_neighbourhood_extensor) -apply (simp add: absolute_neighbourhood_extensor_imp_ANR) -done - -lemma ANR_imp_neighbourhood_retract: - fixes S :: "'a::euclidean_space set" - assumes "ANR S" "closedin (subtopology euclidean U) S" - obtains V where "openin (subtopology euclidean U) V" "S retract_of V" -using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast - -lemma ANR_imp_absolute_closed_neighbourhood_retract: - fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "ANR S" "S homeomorphic S'" and US': "closedin (subtopology euclidean U) S'" - obtains V W - where "openin (subtopology euclidean U) V" - "closedin (subtopology euclidean U) W" - "S' \ V" "V \ W" "S' retract_of W" -proof - - obtain Z where "openin (subtopology euclidean U) Z" and S'Z: "S' retract_of Z" - by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) - then have UUZ: "closedin (subtopology euclidean U) (U - Z)" - by auto - have "S' \ (U - Z) = {}" - using \S' retract_of Z\ closedin_retract closedin_subtopology by fastforce - then obtain V W - where "openin (subtopology euclidean U) V" - and "openin (subtopology euclidean U) W" - and "S' \ V" "U - Z \ W" "V \ W = {}" - using separation_normal_local [OF US' UUZ] by auto - moreover have "S' retract_of U - W" - apply (rule retract_of_subset [OF S'Z]) - using US' \S' \ V\ \V \ W = {}\ closedin_subset apply fastforce - using Diff_subset_conv \U - Z \ W\ by blast - ultimately show ?thesis - apply (rule_tac V=V and W = "U-W" in that) - using openin_imp_subset apply force+ - done -qed - -lemma ANR_imp_closed_neighbourhood_retract: - fixes S :: "'a::euclidean_space set" - assumes "ANR S" "closedin (subtopology euclidean U) S" - obtains V W where "openin (subtopology euclidean U) V" - "closedin (subtopology euclidean U) W" - "S \ V" "V \ W" "S retract_of W" -by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) - -lemma ANR_homeomorphic_ANR: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "ANR T" "S homeomorphic T" - shows "ANR S" -unfolding ANR_def -by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym) - -lemma homeomorphic_ANR_iff_ANR: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - shows "S homeomorphic T \ ANR S \ ANR T" -by (metis ANR_homeomorphic_ANR homeomorphic_sym) - -subsection\ Analogous properties of ENRs\ - -proposition ENR_imp_absolute_neighbourhood_retract: - fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "ENR S" and hom: "S homeomorphic S'" - and "S' \ U" - obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" -proof - - obtain X where "open X" "S retract_of X" - using \ENR S\ by (auto simp: ENR_def) - then obtain r where "retraction X S r" - by (auto simp: retract_of_def) - have "locally compact S'" - using retract_of_locally_compact open_imp_locally_compact - homeomorphic_local_compactness \S retract_of X\ \open X\ hom by blast - then obtain W where UW: "openin (subtopology euclidean U) W" - and WS': "closedin (subtopology euclidean W) S'" - apply (rule locally_compact_closedin_open) - apply (rename_tac W) - apply (rule_tac W = "U \ W" in that, blast) - by (simp add: \S' \ U\ closedin_limpt) - obtain f g where hom: "homeomorphism S S' f g" - using assms by (force simp: homeomorphic_def) - have contg: "continuous_on S' g" - using hom homeomorphism_def by blast - moreover have "g ` S' \ S" by (metis hom equalityE homeomorphism_def) - ultimately obtain h where conth: "continuous_on W h" and hg: "\x. x \ S' \ h x = g x" - using Tietze_unbounded [of S' g W] WS' by blast - have "W \ U" using UW openin_open by auto - have "S' \ W" using WS' closedin_closed by auto - have him: "\x. x \ S' \ h x \ X" - by (metis (no_types) \S retract_of X\ hg hom homeomorphism_def image_insert insert_absorb insert_iff retract_of_imp_subset subset_eq) - have "S' retract_of (W \ h -` X)" - proof (simp add: retraction_def retract_of_def, intro exI conjI) - show "S' \ W" "S' \ h -` X" - using him WS' closedin_imp_subset by blast+ - show "continuous_on (W \ h -` X) (f \ r \ h)" - proof (intro continuous_on_compose) - show "continuous_on (W \ h -` X) h" - by (meson conth continuous_on_subset inf_le1) - show "continuous_on (h ` (W \ h -` X)) r" - proof - - have "h ` (W \ h -` X) \ X" - by blast - then show "continuous_on (h ` (W \ h -` X)) r" - by (meson \retraction X S r\ continuous_on_subset retraction) - qed - show "continuous_on (r ` h ` (W \ h -` X)) f" - apply (rule continuous_on_subset [of S]) - using hom homeomorphism_def apply blast - apply clarify - apply (meson \retraction X S r\ subsetD imageI retraction_def) - done - qed - show "(f \ r \ h) ` (W \ h -` X) \ S'" - using \retraction X S r\ hom - by (auto simp: retraction_def homeomorphism_def) - show "\x\S'. (f \ r \ h) x = x" - using \retraction X S r\ hom by (auto simp: retraction_def homeomorphism_def hg) - qed - then show ?thesis - apply (rule_tac V = "W \ h -` X" in that) - apply (rule openin_trans [OF _ UW]) - using \continuous_on W h\ \open X\ continuous_openin_preimage_eq apply blast+ - done -qed - -corollary ENR_imp_absolute_neighbourhood_retract_UNIV: - fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "ENR S" "S homeomorphic S'" - obtains T' where "open T'" "S' retract_of T'" -by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV) - -lemma ENR_homeomorphic_ENR: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "ENR T" "S homeomorphic T" - shows "ENR S" -unfolding ENR_def -by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym) - -lemma homeomorphic_ENR_iff_ENR: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "S homeomorphic T" - shows "ENR S \ ENR T" -by (meson ENR_homeomorphic_ENR assms homeomorphic_sym) - -lemma ENR_translation: - fixes S :: "'a::euclidean_space set" - shows "ENR(image (\x. a + x) S) \ ENR S" -by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR) - -lemma ENR_linear_image_eq: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "linear f" "inj f" - shows "ENR (image f S) \ ENR S" -apply (rule homeomorphic_ENR_iff_ENR) -using assms homeomorphic_sym linear_homeomorphic_image by auto - -subsection\Some relations among the concepts\ - -text\We also relate AR to being a retract of UNIV, which is often a more convenient proxy in the closed case.\ - -lemma AR_imp_ANR: "AR S \ ANR S" - using ANR_def AR_def by fastforce - -lemma ENR_imp_ANR: - fixes S :: "'a::euclidean_space set" - shows "ENR S \ ANR S" -apply (simp add: ANR_def) -by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset) - -lemma ENR_ANR: - fixes S :: "'a::euclidean_space set" - shows "ENR S \ ANR S \ locally compact S" -proof - assume "ENR S" - then have "locally compact S" - using ENR_def open_imp_locally_compact retract_of_locally_compact by auto - then show "ANR S \ locally compact S" - using ENR_imp_ANR \ENR S\ by blast -next - assume "ANR S \ locally compact S" - then have "ANR S" "locally compact S" by auto - then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T" - using locally_compact_homeomorphic_closed - by (metis DIM_prod DIM_real Suc_eq_plus1 lessI) - then show "ENR S" - using \ANR S\ - apply (simp add: ANR_def) - apply (drule_tac x=UNIV in spec) - apply (drule_tac x=T in spec, clarsimp) - apply (meson ENR_def ENR_homeomorphic_ENR open_openin) - done -qed - - -proposition AR_ANR: - fixes S :: "'a::euclidean_space set" - shows "AR S \ ANR S \ contractible S \ S \ {}" - (is "?lhs = ?rhs") -proof - assume ?lhs - obtain C and S' :: "('a * real) set" - where "convex C" "C \ {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'" - apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) - using aff_dim_le_DIM [of S] by auto - with \AR S\ have "contractible S" - apply (simp add: AR_def) - apply (drule_tac x=C in spec) - apply (drule_tac x="S'" in spec, simp) - using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce - with \AR S\ show ?rhs - apply (auto simp: AR_imp_ANR) - apply (force simp: AR_def) - done -next - assume ?rhs - then obtain a and h:: "real \ 'a \ 'a" - where conth: "continuous_on ({0..1} \ S) h" - and hS: "h ` ({0..1} \ S) \ S" - and [simp]: "\x. h(0, x) = x" - and [simp]: "\x. h(1, x) = a" - and "ANR S" "S \ {}" - by (auto simp: contractible_def homotopic_with_def) - then have "a \ S" - by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) - have "\g. continuous_on W g \ g ` W \ S \ (\x\T. g x = f x)" - if f: "continuous_on T f" "f ` T \ S" - and WT: "closedin (subtopology euclidean W) T" - for W T and f :: "'a \ real \ 'a" - proof - - obtain U g - where "T \ U" and WU: "openin (subtopology euclidean W) U" - and contg: "continuous_on U g" - and "g ` U \ S" and gf: "\x. x \ T \ g x = f x" - using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \ANR S\, rule_format, OF f WT] - by auto - have WWU: "closedin (subtopology euclidean W) (W - U)" - using WU closedin_diff by fastforce - moreover have "(W - U) \ T = {}" - using \T \ U\ by auto - ultimately obtain V V' - where WV': "openin (subtopology euclidean W) V'" - and WV: "openin (subtopology euclidean W) V" - and "W - U \ V'" "T \ V" "V' \ V = {}" - using separation_normal_local [of W "W-U" T] WT by blast - then have WVT: "T \ (W - V) = {}" - by auto - have WWV: "closedin (subtopology euclidean W) (W - V)" - using WV closedin_diff by fastforce - obtain j :: " 'a \ real \ real" - where contj: "continuous_on W j" - and j: "\x. x \ W \ j x \ {0..1}" - and j0: "\x. x \ W - V \ j x = 1" - and j1: "\x. x \ T \ j x = 0" - by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment) - have Weq: "W = (W - V) \ (W - V')" - using \V' \ V = {}\ by force - show ?thesis - proof (intro conjI exI) - have *: "continuous_on (W - V') (\x. h (j x, g x))" - apply (rule continuous_on_compose2 [OF conth continuous_on_Pair]) - apply (rule continuous_on_subset [OF contj Diff_subset]) - apply (rule continuous_on_subset [OF contg]) - apply (metis Diff_subset_conv Un_commute \W - U \ V'\) - using j \g ` U \ S\ \W - U \ V'\ apply fastforce - done - show "continuous_on W (\x. if x \ W - V then a else h (j x, g x))" - apply (subst Weq) - apply (rule continuous_on_cases_local) - apply (simp_all add: Weq [symmetric] WWV continuous_on_const *) - using WV' closedin_diff apply fastforce - apply (auto simp: j0 j1) - done - next - have "h (j (x, y), g (x, y)) \ S" if "(x, y) \ W" "(x, y) \ V" for x y - proof - - have "j(x, y) \ {0..1}" - using j that by blast - moreover have "g(x, y) \ S" - using \V' \ V = {}\ \W - U \ V'\ \g ` U \ S\ that by fastforce - ultimately show ?thesis - using hS by blast - qed - with \a \ S\ \g ` U \ S\ - show "(\x. if x \ W - V then a else h (j x, g x)) ` W \ S" - by auto - next - show "\x\T. (if x \ W - V then a else h (j x, g x)) = f x" - using \T \ V\ by (auto simp: j0 j1 gf) - qed - qed - then show ?lhs - by (simp add: AR_eq_absolute_extensor) -qed - - -lemma ANR_retract_of_ANR: - fixes S :: "'a::euclidean_space set" - assumes "ANR T" "S retract_of T" - shows "ANR S" -using assms -apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def) -apply (clarsimp elim!: all_forward) -apply (erule impCE, metis subset_trans) -apply (clarsimp elim!: ex_forward) -apply (rule_tac x="r \ g" in exI) -by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans) - -lemma AR_retract_of_AR: - fixes S :: "'a::euclidean_space set" - shows "\AR T; S retract_of T\ \ AR S" -using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce - -lemma ENR_retract_of_ENR: - "\ENR T; S retract_of T\ \ ENR S" -by (meson ENR_def retract_of_trans) - -lemma retract_of_UNIV: - fixes S :: "'a::euclidean_space set" - shows "S retract_of UNIV \ AR S \ closed S" -by (metis AR_ANR AR_imp_retract ENR_def ENR_imp_ANR closed_UNIV closed_closedin contractible_UNIV empty_not_UNIV open_UNIV retract_of_closed retract_of_contractible retract_of_empty(1) subtopology_UNIV) - -lemma compact_AR: - fixes S :: "'a::euclidean_space set" - shows "compact S \ AR S \ compact S \ S retract_of UNIV" -using compact_imp_closed retract_of_UNIV by blast - -subsection\More properties of ARs, ANRs and ENRs\ - -lemma not_AR_empty [simp]: "~ AR({})" - by (auto simp: AR_def) - -lemma ENR_empty [simp]: "ENR {}" - by (simp add: ENR_def) - -lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)" - by (simp add: ENR_imp_ANR) - -lemma convex_imp_AR: - fixes S :: "'a::euclidean_space set" - shows "\convex S; S \ {}\ \ AR S" -apply (rule absolute_extensor_imp_AR) -apply (rule Dugundji, assumption+) -by blast - -lemma convex_imp_ANR: - fixes S :: "'a::euclidean_space set" - shows "convex S \ ANR S" -using ANR_empty AR_imp_ANR convex_imp_AR by blast - -lemma ENR_convex_closed: - fixes S :: "'a::euclidean_space set" - shows "\closed S; convex S\ \ ENR S" -using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast - -lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)" - using retract_of_UNIV by auto - -lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)" - by (simp add: AR_imp_ANR) - -lemma ENR_UNIV [simp]:"ENR UNIV" - using ENR_def by blast - -lemma AR_singleton: - fixes a :: "'a::euclidean_space" - shows "AR {a}" - using retract_of_UNIV by blast - -lemma ANR_singleton: - fixes a :: "'a::euclidean_space" - shows "ANR {a}" - by (simp add: AR_imp_ANR AR_singleton) - -lemma ENR_singleton: "ENR {a}" - using ENR_def by blast - -subsection\ARs closed under union\ - -lemma AR_closed_Un_local_aux: - fixes U :: "'a::euclidean_space set" - assumes "closedin (subtopology euclidean U) S" - "closedin (subtopology euclidean U) T" - "AR S" "AR T" "AR(S \ T)" - shows "(S \ T) retract_of U" -proof - - have "S \ T \ {}" - using assms AR_def by fastforce - have "S \ U" "T \ U" - using assms by (auto simp: closedin_imp_subset) - define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" - define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" - define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" - have US': "closedin (subtopology euclidean U) S'" - using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] - by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) - have UT': "closedin (subtopology euclidean U) T'" - using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] - by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) - have "S \ S'" - using S'_def \S \ U\ setdist_sing_in_set by fastforce - have "T \ T'" - using T'_def \T \ U\ setdist_sing_in_set by fastforce - have "S \ T \ W" "W \ U" - using \S \ U\ by (auto simp: W_def setdist_sing_in_set) - have "(S \ T) retract_of W" - apply (rule AR_imp_absolute_retract [OF \AR(S \ T)\]) - apply (simp add: homeomorphic_refl) - apply (rule closedin_subset_trans [of U]) - apply (simp_all add: assms closedin_Int \S \ T \ W\ \W \ U\) - done - then obtain r0 - where "S \ T \ W" and contr0: "continuous_on W r0" - and "r0 ` W \ S \ T" - and r0 [simp]: "\x. x \ S \ T \ r0 x = x" - by (auto simp: retract_of_def retraction_def) - have ST: "x \ W \ x \ S \ x \ T" for x - using setdist_eq_0_closedin \S \ T \ {}\ assms - by (force simp: W_def setdist_sing_in_set) - have "S' \ T' = W" - by (auto simp: S'_def T'_def W_def) - then have cloUW: "closedin (subtopology euclidean U) W" - using closedin_Int US' UT' by blast - define r where "r \ \x. if x \ W then r0 x else x" - have "r ` (W \ S) \ S" "r ` (W \ T) \ T" - using \r0 ` W \ S \ T\ r_def by auto - have contr: "continuous_on (W \ (S \ T)) r" - unfolding r_def - proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) - show "closedin (subtopology euclidean (W \ (S \ T))) W" - using \S \ U\ \T \ U\ \W \ U\ \closedin (subtopology euclidean U) W\ closedin_subset_trans by fastforce - show "closedin (subtopology euclidean (W \ (S \ T))) (S \ T)" - by (meson \S \ U\ \T \ U\ \W \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) - show "\x. x \ W \ x \ W \ x \ S \ T \ x \ W \ r0 x = x" - by (auto simp: ST) - qed - have cloUWS: "closedin (subtopology euclidean U) (W \ S)" - by (simp add: cloUW assms closedin_Un) - obtain g where contg: "continuous_on U g" - and "g ` U \ S" and geqr: "\x. x \ W \ S \ g x = r x" - apply (rule AR_imp_absolute_extensor [OF \AR S\ _ _ cloUWS]) - apply (rule continuous_on_subset [OF contr]) - using \r ` (W \ S) \ S\ apply auto - done - have cloUWT: "closedin (subtopology euclidean U) (W \ T)" - by (simp add: cloUW assms closedin_Un) - obtain h where conth: "continuous_on U h" - and "h ` U \ T" and heqr: "\x. x \ W \ T \ h x = r x" - apply (rule AR_imp_absolute_extensor [OF \AR T\ _ _ cloUWT]) - apply (rule continuous_on_subset [OF contr]) - using \r ` (W \ T) \ T\ apply auto - done - have "U = S' \ T'" - by (force simp: S'_def T'_def) - then have cont: "continuous_on U (\x. if x \ S' then g x else h x)" - apply (rule ssubst) - apply (rule continuous_on_cases_local) - using US' UT' \S' \ T' = W\ \U = S' \ T'\ - contg conth continuous_on_subset geqr heqr apply auto - done - have UST: "(\x. if x \ S' then g x else h x) ` U \ S \ T" - using \g ` U \ S\ \h ` U \ T\ by auto - show ?thesis - apply (simp add: retract_of_def retraction_def \S \ U\ \T \ U\) - apply (rule_tac x="\x. if x \ S' then g x else h x" in exI) - apply (intro conjI cont UST) - by (metis IntI ST Un_iff \S \ S'\ \S' \ T' = W\ \T \ T'\ subsetD geqr heqr r0 r_def) -qed - - -proposition AR_closed_Un_local: - fixes S :: "'a::euclidean_space set" - assumes STS: "closedin (subtopology euclidean (S \ T)) S" - and STT: "closedin (subtopology euclidean (S \ T)) T" - and "AR S" "AR T" "AR(S \ T)" - shows "AR(S \ T)" -proof - - have "C retract_of U" - if hom: "S \ T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" - for U and C :: "('a * real) set" - proof - - obtain f g where hom: "homeomorphism (S \ T) C f g" - using hom by (force simp: homeomorphic_def) - have US: "closedin (subtopology euclidean U) (C \ g -` S)" - apply (rule closedin_trans [OF _ UC]) - apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) - using hom homeomorphism_def apply blast - apply (metis hom homeomorphism_def set_eq_subset) - done - have UT: "closedin (subtopology euclidean U) (C \ g -` T)" - apply (rule closedin_trans [OF _ UC]) - apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) - using hom homeomorphism_def apply blast - apply (metis hom homeomorphism_def set_eq_subset) - done - have ARS: "AR (C \ g -` S)" - apply (rule AR_homeomorphic_AR [OF \AR S\]) - apply (simp add: homeomorphic_def) - apply (rule_tac x=g in exI) - apply (rule_tac x=f in exI) - using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) - apply (rule_tac x="f x" in image_eqI, auto) - done - have ART: "AR (C \ g -` T)" - apply (rule AR_homeomorphic_AR [OF \AR T\]) - apply (simp add: homeomorphic_def) - apply (rule_tac x=g in exI) - apply (rule_tac x=f in exI) - using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) - apply (rule_tac x="f x" in image_eqI, auto) - done - have ARI: "AR ((C \ g -` S) \ (C \ g -` T))" - apply (rule AR_homeomorphic_AR [OF \AR (S \ T)\]) - apply (simp add: homeomorphic_def) - apply (rule_tac x=g in exI) - apply (rule_tac x=f in exI) - using hom - apply (auto simp: homeomorphism_def elim!: continuous_on_subset) - apply (rule_tac x="f x" in image_eqI, auto) - done - have "C = (C \ g -` S) \ (C \ g -` T)" - using hom by (auto simp: homeomorphism_def) - then show ?thesis - by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI]) - qed - then show ?thesis - by (force simp: AR_def) -qed - -corollary AR_closed_Un: - fixes S :: "'a::euclidean_space set" - shows "\closed S; closed T; AR S; AR T; AR (S \ T)\ \ AR (S \ T)" -by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV) - -subsection\ANRs closed under union\ - -lemma ANR_closed_Un_local_aux: - fixes U :: "'a::euclidean_space set" - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" - and "ANR S" "ANR T" "ANR(S \ T)" - obtains V where "openin (subtopology euclidean U) V" "(S \ T) retract_of V" -proof (cases "S = {} \ T = {}") - case True with assms that show ?thesis - by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb) -next - case False - then have [simp]: "S \ {}" "T \ {}" by auto - have "S \ U" "T \ U" - using assms by (auto simp: closedin_imp_subset) - define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" - define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" - define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" - have cloUS': "closedin (subtopology euclidean U) S'" - using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] - by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) - have cloUT': "closedin (subtopology euclidean U) T'" - using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] - by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) - have "S \ S'" - using S'_def \S \ U\ setdist_sing_in_set by fastforce - have "T \ T'" - using T'_def \T \ U\ setdist_sing_in_set by fastforce - have "S' \ T' = U" - by (auto simp: S'_def T'_def) - have "W \ S'" - by (simp add: Collect_mono S'_def W_def) - have "W \ T'" - by (simp add: Collect_mono T'_def W_def) - have ST_W: "S \ T \ W" and "W \ U" - using \S \ U\ by (force simp: W_def setdist_sing_in_set)+ - have "S' \ T' = W" - by (auto simp: S'_def T'_def W_def) - then have cloUW: "closedin (subtopology euclidean U) W" - using closedin_Int cloUS' cloUT' by blast - obtain W' W0 where "openin (subtopology euclidean W) W'" - and cloWW0: "closedin (subtopology euclidean W) W0" - and "S \ T \ W'" "W' \ W0" - and ret: "(S \ T) retract_of W0" - apply (rule ANR_imp_closed_neighbourhood_retract [OF \ANR(S \ T)\]) - apply (rule closedin_subset_trans [of U, OF _ ST_W \W \ U\]) - apply (blast intro: assms)+ - done - then obtain U0 where opeUU0: "openin (subtopology euclidean U) U0" - and U0: "S \ T \ U0" "U0 \ W \ W0" - unfolding openin_open using \W \ U\ by blast - have "W0 \ U" - using \W \ U\ cloWW0 closedin_subset by fastforce - obtain r0 - where "S \ T \ W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \ S \ T" - and r0 [simp]: "\x. x \ S \ T \ r0 x = x" - using ret by (force simp: retract_of_def retraction_def) - have ST: "x \ W \ x \ S \ x \ T" for x - using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin) - define r where "r \ \x. if x \ W0 then r0 x else x" - have "r ` (W0 \ S) \ S" "r ` (W0 \ T) \ T" - using \r0 ` W0 \ S \ T\ r_def by auto - have contr: "continuous_on (W0 \ (S \ T)) r" - unfolding r_def - proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) - show "closedin (subtopology euclidean (W0 \ (S \ T))) W0" - apply (rule closedin_subset_trans [of U]) - using cloWW0 cloUW closedin_trans \W0 \ U\ \S \ U\ \T \ U\ apply blast+ - done - show "closedin (subtopology euclidean (W0 \ (S \ T))) (S \ T)" - by (meson \S \ U\ \T \ U\ \W0 \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) - show "\x. x \ W0 \ x \ W0 \ x \ S \ T \ x \ W0 \ r0 x = x" - using ST cloWW0 closedin_subset by fastforce - qed - have cloS'WS: "closedin (subtopology euclidean S') (W0 \ S)" - by (meson closedin_subset_trans US cloUS' \S \ S'\ \W \ S'\ cloUW cloWW0 - closedin_Un closedin_imp_subset closedin_trans) - obtain W1 g where "W0 \ S \ W1" and contg: "continuous_on W1 g" - and opeSW1: "openin (subtopology euclidean S') W1" - and "g ` W1 \ S" and geqr: "\x. x \ W0 \ S \ g x = r x" - apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ _ \r ` (W0 \ S) \ S\ cloS'WS]) - apply (rule continuous_on_subset [OF contr], blast+) - done - have cloT'WT: "closedin (subtopology euclidean T') (W0 \ T)" - by (meson closedin_subset_trans UT cloUT' \T \ T'\ \W \ T'\ cloUW cloWW0 - closedin_Un closedin_imp_subset closedin_trans) - obtain W2 h where "W0 \ T \ W2" and conth: "continuous_on W2 h" - and opeSW2: "openin (subtopology euclidean T') W2" - and "h ` W2 \ T" and heqr: "\x. x \ W0 \ T \ h x = r x" - apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ _ \r ` (W0 \ T) \ T\ cloT'WT]) - apply (rule continuous_on_subset [OF contr], blast+) - done - have "S' \ T' = W" - by (force simp: S'_def T'_def W_def) - obtain O1 O2 where "open O1" "W1 = S' \ O1" "open O2" "W2 = T' \ O2" - using opeSW1 opeSW2 by (force simp: openin_open) - show ?thesis - proof - have eq: "W1 - (W - U0) \ (W2 - (W - U0)) = - ((U - T') \ O1 \ (U - S') \ O2 \ U \ O1 \ O2) - (W - U0)" - using \U0 \ W \ W0\ \W0 \ S \ W1\ \W0 \ T \ W2\ - by (auto simp: \S' \ T' = U\ [symmetric] \S' \ T' = W\ [symmetric] \W1 = S' \ O1\ \W2 = T' \ O2\) - show "openin (subtopology euclidean U) (W1 - (W - U0) \ (W2 - (W - U0)))" - apply (subst eq) - apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \open O1\ \open O2\, simp_all) - done - have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \ (W2 - (W - U0)))) (W1 - (W - U0))" - using cloUS' apply (simp add: closedin_closed) - apply (erule ex_forward) - using U0 \W0 \ S \ W1\ - apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) - done - have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \ (W2 - (W - U0)))) (W2 - (W - U0))" - using cloUT' apply (simp add: closedin_closed) - apply (erule ex_forward) - using U0 \W0 \ T \ W2\ - apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) - done - have *: "\x\S \ T. (if x \ S' then g x else h x) = x" - using ST \S' \ T' = W\ cloT'WT closedin_subset geqr heqr - apply (auto simp: r_def, fastforce) - using \S \ S'\ \T \ T'\ \W0 \ S \ W1\ \W1 = S' \ O1\ by auto - have "\r. continuous_on (W1 - (W - U0) \ (W2 - (W - U0))) r \ - r ` (W1 - (W - U0) \ (W2 - (W - U0))) \ S \ T \ - (\x\S \ T. r x = x)" - apply (rule_tac x = "\x. if x \ S' then g x else h x" in exI) - apply (intro conjI *) - apply (rule continuous_on_cases_local - [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]]) - using \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = W\ - \g ` W1 \ S\ \h ` W2 \ T\ apply auto - using \U0 \ W \ W0\ \W0 \ S \ W1\ apply (fastforce simp add: geqr heqr)+ - done - then show "S \ T retract_of W1 - (W - U0) \ (W2 - (W - U0))" - using \W0 \ S \ W1\ \W0 \ T \ W2\ ST opeUU0 U0 - by (auto simp: retract_of_def retraction_def) - qed -qed - - -proposition ANR_closed_Un_local: - fixes S :: "'a::euclidean_space set" - assumes STS: "closedin (subtopology euclidean (S \ T)) S" - and STT: "closedin (subtopology euclidean (S \ T)) T" - and "ANR S" "ANR T" "ANR(S \ T)" - shows "ANR(S \ T)" -proof - - have "\T. openin (subtopology euclidean U) T \ C retract_of T" - if hom: "S \ T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" - for U and C :: "('a * real) set" - proof - - obtain f g where hom: "homeomorphism (S \ T) C f g" - using hom by (force simp: homeomorphic_def) - have US: "closedin (subtopology euclidean U) (C \ g -` S)" - apply (rule closedin_trans [OF _ UC]) - apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) - using hom [unfolded homeomorphism_def] apply blast - apply (metis hom homeomorphism_def set_eq_subset) - done - have UT: "closedin (subtopology euclidean U) (C \ g -` T)" - apply (rule closedin_trans [OF _ UC]) - apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) - using hom [unfolded homeomorphism_def] apply blast - apply (metis hom homeomorphism_def set_eq_subset) - done - have ANRS: "ANR (C \ g -` S)" - apply (rule ANR_homeomorphic_ANR [OF \ANR S\]) - apply (simp add: homeomorphic_def) - apply (rule_tac x=g in exI) - apply (rule_tac x=f in exI) - using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) - apply (rule_tac x="f x" in image_eqI, auto) - done - have ANRT: "ANR (C \ g -` T)" - apply (rule ANR_homeomorphic_ANR [OF \ANR T\]) - apply (simp add: homeomorphic_def) - apply (rule_tac x=g in exI) - apply (rule_tac x=f in exI) - using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) - apply (rule_tac x="f x" in image_eqI, auto) - done - have ANRI: "ANR ((C \ g -` S) \ (C \ g -` T))" - apply (rule ANR_homeomorphic_ANR [OF \ANR (S \ T)\]) - apply (simp add: homeomorphic_def) - apply (rule_tac x=g in exI) - apply (rule_tac x=f in exI) - using hom - apply (auto simp: homeomorphism_def elim!: continuous_on_subset) - apply (rule_tac x="f x" in image_eqI, auto) - done - have "C = (C \ g -` S) \ (C \ g -` T)" - using hom by (auto simp: homeomorphism_def) - then show ?thesis - by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI]) - qed - then show ?thesis - by (auto simp: ANR_def) -qed - -corollary ANR_closed_Un: - fixes S :: "'a::euclidean_space set" - shows "\closed S; closed T; ANR S; ANR T; ANR (S \ T)\ \ ANR (S \ T)" -by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int) - -lemma ANR_openin: - fixes S :: "'a::euclidean_space set" - assumes "ANR T" and opeTS: "openin (subtopology euclidean T) S" - shows "ANR S" -proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) - fix f :: "'a \ real \ 'a" and U C - assume contf: "continuous_on C f" and fim: "f ` C \ S" - and cloUC: "closedin (subtopology euclidean U) C" - have "f ` C \ T" - using fim opeTS openin_imp_subset by blast - obtain W g where "C \ W" - and UW: "openin (subtopology euclidean U) W" - and contg: "continuous_on W g" - and gim: "g ` W \ T" - and geq: "\x. x \ C \ g x = f x" - apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf \f ` C \ T\ cloUC]) - using fim by auto - show "\V g. C \ V \ openin (subtopology euclidean U) V \ continuous_on V g \ g ` V \ S \ (\x\C. g x = f x)" - proof (intro exI conjI) - show "C \ W \ g -` S" - using \C \ W\ fim geq by blast - show "openin (subtopology euclidean U) (W \ g -` S)" - by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) - show "continuous_on (W \ g -` S) g" - by (blast intro: continuous_on_subset [OF contg]) - show "g ` (W \ g -` S) \ S" - using gim by blast - show "\x\C. g x = f x" - using geq by blast - qed -qed - -lemma ENR_openin: - fixes S :: "'a::euclidean_space set" - assumes "ENR T" and opeTS: "openin (subtopology euclidean T) S" - shows "ENR S" - using assms apply (simp add: ENR_ANR) - using ANR_openin locally_open_subset by blast - -lemma ANR_neighborhood_retract: - fixes S :: "'a::euclidean_space set" - assumes "ANR U" "S retract_of T" "openin (subtopology euclidean U) T" - shows "ANR S" - using ANR_openin ANR_retract_of_ANR assms by blast - -lemma ENR_neighborhood_retract: - fixes S :: "'a::euclidean_space set" - assumes "ENR U" "S retract_of T" "openin (subtopology euclidean U) T" - shows "ENR S" - using ENR_openin ENR_retract_of_ENR assms by blast - -lemma ANR_rel_interior: - fixes S :: "'a::euclidean_space set" - shows "ANR S \ ANR(rel_interior S)" - by (blast intro: ANR_openin openin_set_rel_interior) - -lemma ANR_delete: - fixes S :: "'a::euclidean_space set" - shows "ANR S \ ANR(S - {a})" - by (blast intro: ANR_openin openin_delete openin_subtopology_self) - -lemma ENR_rel_interior: - fixes S :: "'a::euclidean_space set" - shows "ENR S \ ENR(rel_interior S)" - by (blast intro: ENR_openin openin_set_rel_interior) - -lemma ENR_delete: - fixes S :: "'a::euclidean_space set" - shows "ENR S \ ENR(S - {a})" - by (blast intro: ENR_openin openin_delete openin_subtopology_self) - -lemma open_imp_ENR: "open S \ ENR S" - using ENR_def by blast - -lemma open_imp_ANR: - fixes S :: "'a::euclidean_space set" - shows "open S \ ANR S" - by (simp add: ENR_imp_ANR open_imp_ENR) - -lemma ANR_ball [iff]: - fixes a :: "'a::euclidean_space" - shows "ANR(ball a r)" - by (simp add: convex_imp_ANR) - -lemma ENR_ball [iff]: "ENR(ball a r)" - by (simp add: open_imp_ENR) - -lemma AR_ball [simp]: - fixes a :: "'a::euclidean_space" - shows "AR(ball a r) \ 0 < r" - by (auto simp: AR_ANR convex_imp_contractible) - -lemma ANR_cball [iff]: - fixes a :: "'a::euclidean_space" - shows "ANR(cball a r)" - by (simp add: convex_imp_ANR) - -lemma ENR_cball: - fixes a :: "'a::euclidean_space" - shows "ENR(cball a r)" - using ENR_convex_closed by blast - -lemma AR_cball [simp]: - fixes a :: "'a::euclidean_space" - shows "AR(cball a r) \ 0 \ r" - by (auto simp: AR_ANR convex_imp_contractible) - -lemma ANR_box [iff]: - fixes a :: "'a::euclidean_space" - shows "ANR(cbox a b)" "ANR(box a b)" - by (auto simp: convex_imp_ANR open_imp_ANR) - -lemma ENR_box [iff]: - fixes a :: "'a::euclidean_space" - shows "ENR(cbox a b)" "ENR(box a b)" -apply (simp add: ENR_convex_closed closed_cbox) -by (simp add: open_box open_imp_ENR) - -lemma AR_box [simp]: - "AR(cbox a b) \ cbox a b \ {}" "AR(box a b) \ box a b \ {}" - by (auto simp: AR_ANR convex_imp_contractible) - -lemma ANR_interior: - fixes S :: "'a::euclidean_space set" - shows "ANR(interior S)" - by (simp add: open_imp_ANR) - -lemma ENR_interior: - fixes S :: "'a::euclidean_space set" - shows "ENR(interior S)" - by (simp add: open_imp_ENR) - -lemma AR_imp_contractible: - fixes S :: "'a::euclidean_space set" - shows "AR S \ contractible S" - by (simp add: AR_ANR) - -lemma ENR_imp_locally_compact: - fixes S :: "'a::euclidean_space set" - shows "ENR S \ locally compact S" - by (simp add: ENR_ANR) - -lemma ANR_imp_locally_path_connected: - fixes S :: "'a::euclidean_space set" - assumes "ANR S" - shows "locally path_connected S" -proof - - obtain U and T :: "('a \ real) set" - where "convex U" "U \ {}" - and UT: "closedin (subtopology euclidean U) T" - and "S homeomorphic T" - apply (rule homeomorphic_closedin_convex [of S]) - using aff_dim_le_DIM [of S] apply auto - done - then have "locally path_connected T" - by (meson ANR_imp_absolute_neighbourhood_retract - assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) - then have S: "locally path_connected S" - if "openin (subtopology euclidean U) V" "T retract_of V" "U \ {}" for V - using \S homeomorphic T\ homeomorphic_locally homeomorphic_path_connectedness by blast - show ?thesis - using assms - apply (clarsimp simp: ANR_def) - apply (drule_tac x=U in spec) - apply (drule_tac x=T in spec) - using \S homeomorphic T\ \U \ {}\ UT apply (blast intro: S) - done -qed - -lemma ANR_imp_locally_connected: - fixes S :: "'a::euclidean_space set" - assumes "ANR S" - shows "locally connected S" -using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto - -lemma AR_imp_locally_path_connected: - fixes S :: "'a::euclidean_space set" - assumes "AR S" - shows "locally path_connected S" -by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms) - -lemma AR_imp_locally_connected: - fixes S :: "'a::euclidean_space set" - assumes "AR S" - shows "locally connected S" -using ANR_imp_locally_connected AR_ANR assms by blast - -lemma ENR_imp_locally_path_connected: - fixes S :: "'a::euclidean_space set" - assumes "ENR S" - shows "locally path_connected S" -by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms) - -lemma ENR_imp_locally_connected: - fixes S :: "'a::euclidean_space set" - assumes "ENR S" - shows "locally connected S" -using ANR_imp_locally_connected ENR_ANR assms by blast - -lemma ANR_Times: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "ANR S" "ANR T" shows "ANR(S \ T)" -proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) - fix f :: " ('a \ 'b) \ real \ 'a \ 'b" and U C - assume "continuous_on C f" and fim: "f ` C \ S \ T" - and cloUC: "closedin (subtopology euclidean U) C" - have contf1: "continuous_on C (fst \ f)" - by (simp add: \continuous_on C f\ continuous_on_fst) - obtain W1 g where "C \ W1" - and UW1: "openin (subtopology euclidean U) W1" - and contg: "continuous_on W1 g" - and gim: "g ` W1 \ S" - and geq: "\x. x \ C \ g x = (fst \ f) x" - apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ contf1 _ cloUC]) - using fim apply auto - done - have contf2: "continuous_on C (snd \ f)" - by (simp add: \continuous_on C f\ continuous_on_snd) - obtain W2 h where "C \ W2" - and UW2: "openin (subtopology euclidean U) W2" - and conth: "continuous_on W2 h" - and him: "h ` W2 \ T" - and heq: "\x. x \ C \ h x = (snd \ f) x" - apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf2 _ cloUC]) - using fim apply auto - done - show "\V g. C \ V \ - openin (subtopology euclidean U) V \ - continuous_on V g \ g ` V \ S \ T \ (\x\C. g x = f x)" - proof (intro exI conjI) - show "C \ W1 \ W2" - by (simp add: \C \ W1\ \C \ W2\) - show "openin (subtopology euclidean U) (W1 \ W2)" - by (simp add: UW1 UW2 openin_Int) - show "continuous_on (W1 \ W2) (\x. (g x, h x))" - by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) - show "(\x. (g x, h x)) ` (W1 \ W2) \ S \ T" - using gim him by blast - show "(\x\C. (g x, h x) = f x)" - using geq heq by auto - qed -qed - -lemma AR_Times: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "AR S" "AR T" shows "AR(S \ T)" -using assms by (simp add: AR_ANR ANR_Times contractible_Times) +subsubsection \We continue with ANRs and ENRs\ lemma ENR_rel_frontier_convex: fixes S :: "'a::euclidean_space set" @@ -4065,7 +4071,7 @@ shows "ANR S" by (metis ANR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) -proposition ANR_finite_Union_convex_closed: +lemma ANR_finite_Union_convex_closed: fixes \ :: "'a::euclidean_space set set" assumes \: "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" shows "ANR(\\)" @@ -4149,7 +4155,7 @@ shows "ANR c" by (metis ANR_connected_component_ANR assms componentsE) -subsection\Original ANR material, now for ENRs\ +subsubsection\Original ANR material, now for ENRs\ lemma ENR_bounded: fixes S :: "'a::euclidean_space set" @@ -4258,7 +4264,7 @@ SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);; *) -subsection\Finally, spheres are ANRs and ENRs\ +subsubsection\Finally, spheres are ANRs and ENRs\ lemma absolute_retract_homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and U :: "'b::euclidean_space set" @@ -4291,7 +4297,7 @@ shows "sphere a r retract_of (- {a})" by (simp add: assms sphere_retract_of_punctured_universe_gen) -proposition ENR_sphere: +lemma ENR_sphere: fixes a :: "'a::euclidean_space" shows "ENR(sphere a r)" proof (cases "0 < r") @@ -4307,13 +4313,12 @@ by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial) qed -corollary ANR_sphere: +corollary%unimportant ANR_sphere: fixes a :: "'a::euclidean_space" shows "ANR(sphere a r)" by (simp add: ENR_imp_ANR ENR_sphere) - -subsection\Spheres are connected, etc\ +subsubsection\Spheres are connected, etc\ lemma locally_path_connected_sphere_gen: fixes S :: "'a::euclidean_space set" @@ -4352,8 +4357,7 @@ shows "locally connected(sphere a r)" using ANR_imp_locally_connected ANR_sphere by blast - -subsection\Borsuk homotopy extension theorem\ +subsubsection\Borsuk homotopy extension theorem\ text\It's only this late so we can use the concept of retraction, saying that the domain sets or range set are ENRs.\ @@ -4491,7 +4495,7 @@ qed -corollary nullhomotopic_into_ANR_extension: +corollary%unimportant nullhomotopic_into_ANR_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" @@ -4527,7 +4531,7 @@ done qed -corollary nullhomotopic_into_rel_frontier_extension: +corollary%unimportant nullhomotopic_into_rel_frontier_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" @@ -4538,7 +4542,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 nullhomotopic_into_sphere_extension: +corollary%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" @@ -4560,7 +4564,7 @@ done qed -proposition Borsuk_map_essential_bounded_component: +proposition%unimportant Borsuk_map_essential_bounded_component: fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S" shows "bounded (connected_component_set (- S) a) \ @@ -4700,8 +4704,7 @@ by (simp add: Borsuk_maps_homotopic_in_path_component) qed - -subsection\More extension theorems\ +subsubsection\More extension theorems\ lemma extension_from_clopen: assumes ope: "openin (subtopology euclidean S) T" @@ -4819,7 +4822,7 @@ qed qed -proposition homotopic_neighbourhood_extension: +proposition%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" @@ -4884,7 +4887,7 @@ qed text\ Homotopy on a union of closed-open sets.\ -proposition homotopic_on_clopen_Union: +proposition%unimportant homotopic_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes "\S. S \ \ \ closedin (subtopology euclidean (\\)) S" and "\S. S \ \ \ openin (subtopology euclidean (\\)) S" @@ -4957,7 +4960,7 @@ qed qed -proposition homotopic_on_components_eq: +lemma homotopic_on_components_eq: fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set" assumes S: "locally connected S \ compact S" and "ANR T" shows "homotopic_with (\x. True) S T f g \ @@ -5080,8 +5083,7 @@ qed qed - -subsection\The complement of a set and path-connectedness\ +subsubsection\The complement of a set and path-connectedness\ text\Complement in dimension N > 1 of set homeomorphic to any interval in any dimension is (path-)connected. This naively generalizes the argument diff -r cedf3480fdad -r 75129a73aca3 src/HOL/Analysis/Euclidean_Space.thy --- a/src/HOL/Analysis/Euclidean_Space.thy Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/Analysis/Euclidean_Space.thy Thu Jul 12 11:23:46 2018 +0200 @@ -12,7 +12,7 @@ subsection \Type class of Euclidean spaces\ -class%important euclidean_space = real_inner + +class euclidean_space = real_inner + fixes Basis :: "'a set" assumes nonempty_Basis [simp]: "Basis \ {}" assumes finite_Basis [simp]: "finite Basis" @@ -224,7 +224,7 @@ subsubsection%unimportant \Type @{typ real}\ -instantiation%important real :: euclidean_space +instantiation real :: euclidean_space begin definition @@ -240,7 +240,7 @@ subsubsection%unimportant \Type @{typ complex}\ -instantiation%important complex :: euclidean_space +instantiation complex :: euclidean_space begin definition Basis_complex_def: "Basis = {1, \}" @@ -261,7 +261,7 @@ subsubsection%unimportant \Type @{typ "'a \ 'b"}\ -instantiation%important prod :: (euclidean_space, euclidean_space) euclidean_space +instantiation prod :: (euclidean_space, euclidean_space) euclidean_space begin definition diff -r cedf3480fdad -r 75129a73aca3 src/HOL/Analysis/Inner_Product.thy --- a/src/HOL/Analysis/Inner_Product.thy Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/Analysis/Inner_Product.thy Thu Jul 12 11:23:46 2018 +0200 @@ -274,7 +274,7 @@ subsection \Class instances\ -instantiation%important real :: real_inner +instantiation real :: real_inner begin definition inner_real_def [simp]: "inner = ( * )" @@ -303,7 +303,7 @@ and real_inner_1_right[simp]: "inner x 1 = x" by simp_all -instantiation%important complex :: real_inner +instantiation complex :: real_inner begin definition inner_complex_def: diff -r cedf3480fdad -r 75129a73aca3 src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/Analysis/Measure_Space.thy Thu Jul 12 11:23:46 2018 +0200 @@ -2834,7 +2834,7 @@ of the lexicographical order are point-wise ordered. \ -instantiation%important measure :: (type) order_bot +instantiation measure :: (type) order_bot begin inductive less_eq_measure :: "'a measure \ 'a measure \ bool" where @@ -3028,7 +3028,7 @@ (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def sigma_sets_superset_generator sigma_sets_le_sets_iff) -instantiation%important measure :: (type) semilattice_sup +instantiation measure :: (type) semilattice_sup begin definition%important sup_measure :: "'a measure \ 'a measure \ 'a measure" @@ -3159,7 +3159,7 @@ by (simp add: A(3)) qed -instantiation%important measure :: (type) complete_lattice +instantiation measure :: (type) complete_lattice begin interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure" diff -r cedf3480fdad -r 75129a73aca3 src/HOL/Analysis/Product_Vector.thy --- a/src/HOL/Analysis/Product_Vector.thy Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/Analysis/Product_Vector.thy Thu Jul 12 11:23:46 2018 +0200 @@ -58,7 +58,7 @@ subsection \Product is a real vector space\ -instantiation%important prod :: (real_vector, real_vector) real_vector +instantiation prod :: (real_vector, real_vector) real_vector begin definition scaleR_prod_def: @@ -113,7 +113,7 @@ (* TODO: Product of uniform spaces and compatibility with metric_spaces! *) -instantiation%important prod :: (metric_space, metric_space) dist +instantiation prod :: (metric_space, metric_space) dist begin definition%important dist_prod_def[code del]: @@ -135,7 +135,7 @@ declare uniformity_Abort[where 'a="'a :: metric_space \ 'b :: metric_space", code] -instantiation%important prod :: (metric_space, metric_space) metric_space +instantiation prod :: (metric_space, metric_space) metric_space begin lemma dist_Pair_Pair: "dist (a, b) (c, d) = sqrt ((dist a c)\<^sup>2 + (dist b d)\<^sup>2)" @@ -270,7 +270,7 @@ subsection \Product is a normed vector space\ -instantiation%important prod :: (real_normed_vector, real_normed_vector) real_normed_vector +instantiation prod :: (real_normed_vector, real_normed_vector) real_normed_vector begin definition norm_prod_def[code del]: @@ -398,7 +398,7 @@ subsection \Product is an inner product space\ -instantiation%important prod :: (real_inner, real_inner) real_inner +instantiation prod :: (real_inner, real_inner) real_inner begin definition inner_prod_def: diff -r cedf3480fdad -r 75129a73aca3 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Jul 12 11:23:46 2018 +0200 @@ -603,7 +603,7 @@ text \Textbooks define Polish spaces as completely metrizable. We assume the topology to be complete for a given metric.\ -class%important polish_space = complete_space + second_countable_topology +class polish_space = complete_space + second_countable_topology subsection \General notion of a topology as a value\ @@ -4573,7 +4573,7 @@ Heine-Borel property if every closed and bounded subset is compact. \ -class%important heine_borel = metric_space + +class heine_borel = metric_space + assumes bounded_imp_convergent_subsequence: "bounded (range f) \ \l r. strict_mono (r::nat\nat) \ ((f \ r) \ l) sequentially" diff -r cedf3480fdad -r 75129a73aca3 src/HOL/ROOT --- a/src/HOL/ROOT Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/ROOT Thu Jul 12 11:23:46 2018 +0200 @@ -58,7 +58,7 @@ document_files "root.bib" "root.tex" session "HOL-Analysis" (main timing) in Analysis = HOL + - options [document_tags = "theorem%important,corollary%important,proposition%important,%unimportant", + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] sessions "HOL-Library"