# HG changeset patch # User nipkow # Date 1531478435 -7200 # Node ID 27432da24236844c78f4e723e7399ec8fd438ab0 # Parent 79abf4201e8dd94448eb73644dae0cc000f8a9f8 unit_cube = cbox 0 One diff -r 79abf4201e8d -r 27432da24236 src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Thu Jul 12 17:23:38 2018 +0100 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Fri Jul 13 12:40:35 2018 +0200 @@ -18,40 +18,6 @@ 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\ @@ -3126,44 +3092,44 @@ subsection \Brouwer's fixed point theorem\ -text \We start proving Brouwer's fixed point theorem for unit cubes.\ +text \We start proving Brouwer's fixed point theorem for the unit cube = \cbox 0 One\.\ lemma brouwer_cube: fixes f :: "'a::euclidean_space \ 'a" - assumes "continuous_on unit_cube f" - and "f ` unit_cube \ unit_cube" - shows "\x\unit_cube. f x = x" + assumes "continuous_on (cbox 0 One) f" + and "f ` cbox 0 One \ cbox 0 One" + shows "\x\cbox 0 One. f x = x" proof (rule ccontr) define n where "n = DIM('a)" have n: "1 \ n" "0 < n" "n \ 0" unfolding n_def by (auto simp: Suc_le_eq DIM_positive) assume "\ ?thesis" - then have *: "\ (\x\unit_cube. f x - x = 0)" + then have *: "\ (\x\cbox 0 One. f x - x = 0)" by auto obtain d where - d: "d > 0" "\x. x \ unit_cube \ d \ norm (f x - x)" - apply (rule brouwer_compactness_lemma[OF compact_unit_cube _ *]) + d: "d > 0" "\x. x \ cbox 0 One \ d \ norm (f x - x)" + apply (rule brouwer_compactness_lemma[OF compact_cbox _ *]) apply (rule continuous_intros assms)+ apply blast done - have *: "\x. x \ unit_cube \ f x \ unit_cube" - "\x. x \ (unit_cube::'a set) \ (\i\Basis. True \ 0 \ x \ i \ x \ i \ 1)" + have *: "\x. x \ cbox 0 One \ f x \ cbox 0 One" + "\x. x \ (cbox 0 One::'a set) \ (\i\Basis. True \ 0 \ x \ i \ x \ i \ 1)" using assms(2)[unfolded image_subset_iff Ball_def] - unfolding mem_unit_cube + unfolding cbox_def by auto obtain label :: "'a \ 'a \ nat" where label [rule_format]: "\x. \i\Basis. label x i \ 1" - "\x. \i\Basis. x \ unit_cube \ x \ i = 0 \ label x i = 0" - "\x. \i\Basis. x \ unit_cube \ x \ i = 1 \ label x i = 1" - "\x. \i\Basis. x \ unit_cube \ label x i = 0 \ x \ i \ f x \ i" - "\x. \i\Basis. x \ unit_cube \ label x i = 1 \ f x \ i \ x \ i" + "\x. \i\Basis. x \ cbox 0 One \ x \ i = 0 \ label x i = 0" + "\x. \i\Basis. x \ cbox 0 One \ x \ i = 1 \ label x i = 1" + "\x. \i\Basis. x \ cbox 0 One \ label x i = 0 \ x \ i \ f x \ i" + "\x. \i\Basis. x \ cbox 0 One \ label x i = 1 \ f x \ i \ x \ i" using kuhn_labelling_lemma[OF *] by auto note label = this [rule_format] - have lem1: "\x\unit_cube. \y\unit_cube. \i\Basis. label x i \ label y i \ + have lem1: "\x\cbox 0 One. \y\cbox 0 One. \i\Basis. label x i \ label y i \ \f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" proof safe fix x y :: 'a - assume x: "x \ unit_cube" and y: "y \ unit_cube" + assume x: "x \ cbox 0 One" and y: "y \ cbox 0 One" fix i assume i: "label x i \ label y i" "i \ Basis" have *: "\x y fx fy :: real. x \ fx \ fy \ y \ fx \ x \ y \ fy \ @@ -3188,18 +3154,18 @@ finally show "\f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" unfolding inner_simps . qed - have "\e>0. \x\unit_cube. \y\unit_cube. \z\unit_cube. \i\Basis. + have "\e>0. \x\cbox 0 One. \y\cbox 0 One. \z\cbox 0 One. \i\Basis. norm (x - z) < e \ norm (y - z) < e \ label x i \ label y i \ \(f(z) - z)\i\ < d / (real n)" proof - have d': "d / real n / 8 > 0" using d(1) by (simp add: n_def DIM_positive) - have *: "uniformly_continuous_on unit_cube f" - by (rule compact_uniformly_continuous[OF assms(1) compact_unit_cube]) + have *: "uniformly_continuous_on (cbox 0 One) f" + by (rule compact_uniformly_continuous[OF assms(1) compact_cbox]) obtain e where e: "e > 0" - "\x x'. x \ unit_cube \ - x' \ unit_cube \ + "\x x'. x \ cbox 0 One \ + x' \ cbox 0 One \ norm (x' - x) < e \ norm (f x' - f x) < d / real n / 8" using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] @@ -3211,7 +3177,7 @@ using d' e by auto fix x y z i assume as: - "x \ unit_cube" "y \ unit_cube" "z \ unit_cube" + "x \ cbox 0 One" "y \ cbox 0 One" "z \ cbox 0 One" "norm (x - z) < min (e / 2) (d / real n / 8)" "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \ label y i" @@ -3249,9 +3215,9 @@ then obtain e where e: "e > 0" - "\x y z i. x \ unit_cube \ - y \ unit_cube \ - z \ unit_cube \ + "\x y z i. x \ cbox 0 One \ + y \ cbox 0 One \ + z \ cbox 0 One \ i \ Basis \ norm (x - z) < e \ norm (y - z) < e \ label x i \ label y i \ \(f z - z) \ i\ < d / real n" @@ -3290,9 +3256,9 @@ using label(1)[OF b''] by auto { fix x :: "nat \ nat" and i assume "\i p" "i < n" "x i = p \ x i = 0" - then have "(\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ (unit_cube::'a set)" + then have "(\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ (cbox 0 One::'a set)" using b'_Basis - by (auto simp: mem_unit_cube inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) } + by (auto simp: cbox_def inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) } note cube = this have q2: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0)" @@ -3314,8 +3280,8 @@ have "\i\Basis. q (b' i) \ {0..p}" using q(1) b' by (auto intro: less_imp_le simp: bij_betw_def) - then have "z \ unit_cube" - unfolding z_def mem_unit_cube + then have "z \ cbox 0 One" + unfolding z_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) then have d_fz_z: "d \ norm (f z - z)" @@ -3353,18 +3319,18 @@ using q(1)[rule_format,OF b'_im] apply (auto simp: Suc_le_eq) done - then have "r' \ unit_cube" - unfolding r'_def mem_unit_cube + then have "r' \ cbox 0 One" + unfolding r'_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) define s' :: 'a where "s' = (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ s (b' i) \ p" using b'_im q(1) rs(2) by fastforce - then have "s' \ unit_cube" - unfolding s'_def mem_unit_cube + then have "s' \ cbox 0 One" + unfolding s'_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) - have "z \ unit_cube" - unfolding z_def mem_unit_cube + have "z \ cbox 0 One" + unfolding z_def cbox_def using b'_Basis q(1)[rule_format,OF b'_im] \p > 0\ by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) { @@ -3394,7 +3360,7 @@ then have "\(f z - z) \ i\ < d / real n" using rs(3) i unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' - by (intro e(2)[OF \r'\unit_cube\ \s'\unit_cube\ \z\unit_cube\]) auto + by (intro e(2)[OF \r'\cbox 0 One\ \s'\cbox 0 One\ \z\cbox 0 One\]) auto then show False using i by auto qed @@ -3410,7 +3376,7 @@ and "f ` S \ S" obtains x where "x \ S" and "f x = x" proof - - let ?U = "unit_cube :: 'a set" + let ?U = "cbox 0 One :: 'a set" have "\Basis /\<^sub>R 2 \ interior ?U" proof (rule interiorI) let ?I = "(\i\Basis. {x::'a. 0 < x \ i} \ {x. x \ i < 1})" @@ -3418,12 +3384,12 @@ by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner continuous_on_const continuous_on_id) show "\Basis /\<^sub>R 2 \ ?I" by simp - show "?I \ unit_cube" - unfolding unit_cube_def by force + show "?I \ cbox 0 One" + unfolding cbox_def by force qed then have *: "interior ?U \ {}" by fast have *: "?U homeomorphic S" - using homeomorphic_convex_compact[OF convex_unit_cube compact_unit_cube * assms(2,1,3)] . + using homeomorphic_convex_compact[OF convex_box(1) compact_cbox * assms(2,1,3)] . have "\f. continuous_on ?U f \ f ` ?U \ ?U \ (\x\?U. f x = x)" using brouwer_cube by auto