diff --git a/src/HOL/Analysis/Brouwer_Fixpoint.thy b/src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy @@ -1,2602 +1,2602 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP *) (* At the moment this is just Brouwer's fixpoint theorem. The proof is from *) (* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518 *) (* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf". *) (* *) (* The script below is quite messy, but at least we avoid formalizing any *) (* topological machinery; we don't even use barycentric subdivision; this is *) (* the big advantage of Kuhn's proof over the usual Sperner's lemma one. *) (* *) (* (c) Copyright, John Harrison 1998-2008 *) section \Brouwer's Fixed Point Theorem\ theory Brouwer_Fixpoint imports Homeomorphism Derivative begin subsection \Retractions\ 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_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 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_canon (\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_canon (\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_canon (\x. True) U S f (\x. c)" and "continuous_on U f" "f ` U \ T" obtains c where "homotopic_with_canon (\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_openin_vimage_iff: "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" if retraction: "retraction S T r" and "U \ T" using retraction apply (rule retractionE) apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) using \U \ T\ apply (auto 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 homotopic_into_retract: "\f ` S \ T; g ` S \ T; T retract_of U; homotopic_with_canon (\x. True) S U f g\ \ homotopic_with_canon (\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: idempotent_imp_retraction intro!: retraction_openin_vimage_iff elim!: locally_connected_quotient_image retract_ofE) 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: idempotent_imp_retraction intro!: retraction_openin_vimage_iff elim!: locally_path_connected_quotient_image retract_ofE) text \A few simple lemmas about deformation retracts\ lemma deformation_retract_imp_homotopy_eqv: fixes S :: "'a::euclidean_space set" assumes "homotopic_with_canon (\x. True) S S id r" and r: "retraction S T r" shows "S homotopy_eqv T" proof - have "homotopic_with_canon (\x. True) S S (id \ r) id" by (simp add: assms(1) homotopic_with_symD) moreover have "homotopic_with_canon (\x. True) T T (r \ id) id" using r unfolding retraction_def by (metis eq_id_iff homotopic_with_id2 topspace_euclidean_subtopology) ultimately show ?thesis unfolding homotopy_equivalent_space_def by (metis (no_types, lifting) continuous_map_subtopology_eu continuous_on_id' id_def image_id r retraction_def) qed lemma deformation_retract: fixes S :: "'a::euclidean_space set" shows "(\r. homotopic_with_canon (\x. True) S S id r \ retraction S T r) \ T retract_of S \ (\f. homotopic_with_canon (\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_with_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_canon (\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_canon (\x. True) S S id (\x. a)" using assms by (auto simp: contractible_def 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 \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)" shows "f a = g a" proof - have "f ` (A - {a}) = g ` (A - {a})" by (intro image_cong) (simp_all add: eq) then have "B - {f a} = B - {g a}" using f g a by (auto simp: bij_betw_def inj_on_image_set_diff set_eq_iff) moreover have "f a \ B" "g a \ B" using f g a by (auto simp: bij_betw_def) ultimately show ?thesis by auto qed lemma swap_image: "Fun.swap i j f ` A = (if i \ A then (if j \ A then f ` A else f ` ((A - {i}) \ {j})) else (if j \ A then f ` ((A - {j}) \ {i}) else f ` A))" by (auto simp: swap_def cong: image_cong_simp) lemmas swap_apply1 = swap_apply(1) lemmas swap_apply2 = swap_apply(2) lemma pointwise_minimal_pointwise_maximal: fixes s :: "(nat \ nat) set" assumes "finite s" and "s \ {}" and "\x\s. \y\s. x \ y \ y \ x" shows "\a\s. \x\s. a \ x" and "\a\s. \x\s. x \ a" using assms proof (induct s rule: finite_ne_induct) case (insert b s) assume *: "\x\insert b s. \y\insert b s. x \ y \ y \ x" then obtain u l where "l \ s" "\b\s. l \ b" "u \ s" "\b\s. b \ u" using insert by auto with * show "\a\insert b s. \x\insert b s. a \ x" "\a\insert b s. \x\insert b s. x \ a" using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+ qed auto lemma kuhn_labelling_lemma: fixes P Q :: "'a::euclidean_space \ bool" assumes "\x. P x \ P (f x)" and "\x. P x \ (\i\Basis. Q i \ 0 \ x\i \ x\i \ 1)" shows "\l. (\x.\i\Basis. l x i \ (1::nat)) \ (\x.\i\Basis. P x \ Q i \ (x\i = 0) \ (l x i = 0)) \ (\x.\i\Basis. P x \ Q i \ (x\i = 1) \ (l x i = 1)) \ (\x.\i\Basis. P x \ Q i \ (l x i = 0) \ x\i \ f x\i) \ (\x.\i\Basis. P x \ Q i \ (l x i = 1) \ f x\i \ x\i)" proof - { fix x i let ?R = "\y. (P x \ Q i \ x \ i = 0 \ y = (0::nat)) \ (P x \ Q i \ x \ i = 1 \ y = 1) \ (P x \ Q i \ y = 0 \ x \ i \ f x \ i) \ (P x \ Q i \ y = 1 \ f x \ i \ x \ i)" { assume "P x" "Q i" "i \ Basis" with assms have "0 \ f x \ i \ f x \ i \ 1" by auto } then have "i \ Basis \ ?R 0 \ ?R 1" by auto } then show ?thesis unfolding all_conj_distrib[symmetric] Ball_def (* FIXME: shouldn't this work by metis? *) by (subst choice_iff[symmetric])+ blast qed subsubsection \The key "counting" observation, somewhat abstracted\ lemma kuhn_counting_lemma: fixes bnd compo compo' face S F defines "nF s == card {f\F. face f s \ compo' f}" assumes [simp, intro]: "finite F" \ \faces\ and [simp, intro]: "finite S" \ \simplices\ and "\f. f \ F \ bnd f \ card {s\S. face f s} = 1" and "\f. f \ F \ \ bnd f \ card {s\S. face f s} = 2" and "\s. s \ S \ compo s \ nF s = 1" and "\s. s \ S \ \ compo s \ nF s = 0 \ nF s = 2" and "odd (card {f\F. compo' f \ bnd f})" shows "odd (card {s\S. compo s})" proof - have "(\s | s \ S \ \ compo s. nF s) + (\s | s \ S \ compo s. nF s) = (\s\S. nF s)" by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong) also have "\ = (\s\S. card {f \ {f\F. compo' f \ bnd f}. face f s}) + (\s\S. card {f \ {f\F. compo' f \ \ bnd f}. face f s})" unfolding sum.distrib[symmetric] by (subst card_Un_disjoint[symmetric]) (auto simp: nF_def intro!: sum.cong arg_cong[where f=card]) also have "\ = 1 * card {f\F. compo' f \ bnd f} + 2 * card {f\F. compo' f \ \ bnd f}" using assms(4,5) by (fastforce intro!: arg_cong2[where f="(+)"] sum_multicount) finally have "odd ((\s | s \ S \ \ compo s. nF s) + card {s\S. compo s})" using assms(6,8) by simp moreover have "(\s | s \ S \ \ compo s. nF s) = (\s | s \ S \ \ compo s \ nF s = 0. nF s) + (\s | s \ S \ \ compo s \ nF s = 2. nF s)" using assms(7) by (subst sum.union_disjoint[symmetric]) (fastforce intro!: sum.cong)+ ultimately show ?thesis by auto qed subsubsection \The odd/even result for faces of complete vertices, generalized\ lemma kuhn_complete_lemma: assumes [simp]: "finite simplices" and face: "\f s. face f s \ (\a\s. f = s - {a})" and card_s[simp]: "\s. s \ simplices \ card s = n + 2" and rl_bd: "\s. s \ simplices \ rl ` s \ {..Suc n}" and bnd: "\f s. s \ simplices \ face f s \ bnd f \ card {s\simplices. face f s} = 1" and nbnd: "\f s. s \ simplices \ face f s \ \ bnd f \ card {s\simplices. face f s} = 2" and odd_card: "odd (card {f. (\s\simplices. face f s) \ rl ` f = {..n} \ bnd f})" shows "odd (card {s\simplices. (rl ` s = {..Suc n})})" proof (rule kuhn_counting_lemma) have finite_s[simp]: "\s. s \ simplices \ finite s" by (metis add_is_0 zero_neq_numeral card.infinite assms(3)) let ?F = "{f. \s\simplices. face f s}" have F_eq: "?F = (\s\simplices. \a\s. {s - {a}})" by (auto simp: face) show "finite ?F" using \finite simplices\ unfolding F_eq by auto show "card {s \ simplices. face f s} = 1" if "f \ ?F" "bnd f" for f using bnd that by auto show "card {s \ simplices. face f s} = 2" if "f \ ?F" "\ bnd f" for f using nbnd that by auto show "odd (card {f \ {f. \s\simplices. face f s}. rl ` f = {..n} \ bnd f})" using odd_card by simp fix s assume s[simp]: "s \ simplices" let ?S = "{f \ {f. \s\simplices. face f s}. face f s \ rl ` f = {..n}}" have "?S = (\a. s - {a}) ` {a\s. rl ` (s - {a}) = {..n}}" using s by (fastforce simp: face) then have card_S: "card ?S = card {a\s. rl ` (s - {a}) = {..n}}" by (auto intro!: card_image inj_onI) { assume rl: "rl ` s = {..Suc n}" then have inj_rl: "inj_on rl s" by (intro eq_card_imp_inj_on) auto moreover obtain a where "rl a = Suc n" "a \ s" by (metis atMost_iff image_iff le_Suc_eq rl) ultimately have n: "{..n} = rl ` (s - {a})" by (auto simp: inj_on_image_set_diff rl) have "{a\s. rl ` (s - {a}) = {..n}} = {a}" using inj_rl \a \ s\ by (auto simp: n inj_on_image_eq_iff[OF inj_rl]) then show "card ?S = 1" unfolding card_S by simp } { assume rl: "rl ` s \ {..Suc n}" show "card ?S = 0 \ card ?S = 2" proof cases assume *: "{..n} \ rl ` s" with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}" by (auto simp: atMost_Suc subset_insert_iff split: if_split_asm) then have "\ inj_on rl s" by (intro pigeonhole) simp then obtain a b where ab: "a \ s" "b \ s" "rl a = rl b" "a \ b" by (auto simp: inj_on_def) then have eq: "rl ` (s - {a}) = rl ` s" by auto with ab have inj: "inj_on rl (s - {a})" by (intro eq_card_imp_inj_on) (auto simp: rl_s card_Diff_singleton_if) { fix x assume "x \ s" "x \ {a, b}" then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})" by (auto simp: eq inj_on_image_set_diff[OF inj]) also have "\ = rl ` (s - {x})" using ab \x \ {a, b}\ by auto also assume "\ = rl ` s" finally have False using \x\s\ by auto } moreover { fix x assume "x \ {a, b}" with ab have "x \ s \ rl ` (s - {x}) = rl ` s" by (simp add: set_eq_iff image_iff Bex_def) metis } ultimately have "{a\s. rl ` (s - {a}) = {..n}} = {a, b}" unfolding rl_s[symmetric] by fastforce with \a \ b\ show "card ?S = 0 \ card ?S = 2" unfolding card_S by simp next assume "\ {..n} \ rl ` s" then have "\x. rl ` (s - {x}) \ {..n}" by auto then show "card ?S = 0 \ card ?S = 2" unfolding card_S by simp qed } qed fact locale kuhn_simplex = fixes p n and base upd and s :: "(nat \ nat) set" assumes base: "base \ {..< n} \ {..< p}" assumes base_out: "\i. n \ i \ base i = p" assumes upd: "bij_betw upd {..< n} {..< n}" assumes s_pre: "s = (\i j. if j \ upd`{..< i} then Suc (base j) else base j) ` {.. n}" begin definition "enum i j = (if j \ upd`{..< i} then Suc (base j) else base j)" lemma s_eq: "s = enum ` {.. n}" unfolding s_pre enum_def[abs_def] .. lemma upd_space: "i < n \ upd i < n" using upd by (auto dest!: bij_betwE) lemma s_space: "s \ {..< n} \ {.. p}" proof - { fix i assume "i \ n" then have "enum i \ {..< n} \ {.. p}" proof (induct i) case 0 then show ?case using base by (auto simp: Pi_iff less_imp_le enum_def) next case (Suc i) with base show ?case by (auto simp: Pi_iff Suc_le_eq less_imp_le enum_def intro: upd_space) qed } then show ?thesis by (auto simp: s_eq) qed lemma inj_upd: "inj_on upd {..< n}" using upd by (simp add: bij_betw_def) lemma inj_enum: "inj_on enum {.. n}" proof - { fix x y :: nat assume "x \ y" "x \ n" "y \ n" with upd have "upd ` {..< x} \ upd ` {..< y}" by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def) then have "enum x \ enum y" by (auto simp: enum_def fun_eq_iff) } then show ?thesis by (auto simp: inj_on_def) qed lemma enum_0: "enum 0 = base" by (simp add: enum_def[abs_def]) lemma base_in_s: "base \ s" unfolding s_eq by (subst enum_0[symmetric]) auto lemma enum_in: "i \ n \ enum i \ s" unfolding s_eq by auto lemma one_step: assumes a: "a \ s" "j < n" assumes *: "\a'. a' \ s \ a' \ a \ a' j = p'" shows "a j \ p'" proof assume "a j = p'" with * a have "\a'. a' \ s \ a' j = p'" by auto then have "\i. i \ n \ enum i j = p'" unfolding s_eq by auto from this[of 0] this[of n] have "j \ upd ` {..< n}" by (auto simp: enum_def fun_eq_iff split: if_split_asm) with upd \j < n\ show False by (auto simp: bij_betw_def) qed lemma upd_inj: "i < n \ j < n \ upd i = upd j \ i = j" using upd by (auto simp: bij_betw_def inj_on_eq_iff) lemma upd_surj: "upd ` {..< n} = {..< n}" using upd by (auto simp: bij_betw_def) lemma in_upd_image: "A \ {..< n} \ i < n \ upd i \ upd ` A \ i \ A" using inj_on_image_mem_iff[of upd "{..< n}"] upd by (auto simp: bij_betw_def) lemma enum_inj: "i \ n \ j \ n \ enum i = enum j \ i = j" using inj_enum by (auto simp: inj_on_eq_iff) lemma in_enum_image: "A \ {.. n} \ i \ n \ enum i \ enum ` A \ i \ A" using inj_on_image_mem_iff[OF inj_enum] by auto lemma enum_mono: "i \ n \ j \ n \ enum i \ enum j \ i \ j" by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric]) lemma enum_strict_mono: "i \ n \ j \ n \ enum i < enum j \ i < j" using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less) lemma chain: "a \ s \ b \ s \ a \ b \ b \ a" by (auto simp: s_eq enum_mono) lemma less: "a \ s \ b \ s \ a i < b i \ a < b" using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric]) lemma enum_0_bot: "a \ s \ a = enum 0 \ (\a'\s. a \ a')" unfolding s_eq by (auto simp: enum_mono Ball_def) lemma enum_n_top: "a \ s \ a = enum n \ (\a'\s. a' \ a)" unfolding s_eq by (auto simp: enum_mono Ball_def) lemma enum_Suc: "i < n \ enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))" by (auto simp: fun_eq_iff enum_def upd_inj) lemma enum_eq_p: "i \ n \ n \ j \ enum i j = p" by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric]) lemma out_eq_p: "a \ s \ n \ j \ a j = p" unfolding s_eq by (auto simp: enum_eq_p) lemma s_le_p: "a \ s \ a j \ p" using out_eq_p[of a j] s_space by (cases "j < n") auto lemma le_Suc_base: "a \ s \ a j \ Suc (base j)" unfolding s_eq by (auto simp: enum_def) lemma base_le: "a \ s \ base j \ a j" unfolding s_eq by (auto simp: enum_def) lemma enum_le_p: "i \ n \ j < n \ enum i j \ p" using enum_in[of i] s_space by auto lemma enum_less: "a \ s \ i < n \ enum i < a \ enum (Suc i) \ a" unfolding s_eq by (auto simp: enum_strict_mono enum_mono) lemma ksimplex_0: "n = 0 \ s = {(\x. p)}" using s_eq enum_def base_out by auto lemma replace_0: assumes "j < n" "a \ s" and p: "\x\s - {a}. x j = 0" and "x \ s" shows "x \ a" proof cases assume "x \ a" have "a j \ 0" using assms by (intro one_step[where a=a]) auto with less[OF \x\s\ \a\s\, of j] p[rule_format, of x] \x \ s\ \x \ a\ show ?thesis by auto qed simp lemma replace_1: assumes "j < n" "a \ s" and p: "\x\s - {a}. x j = p" and "x \ s" shows "a \ x" proof cases assume "x \ a" have "a j \ p" using assms by (intro one_step[where a=a]) auto with enum_le_p[of _ j] \j < n\ \a\s\ have "a j < p" by (auto simp: less_le s_eq) with less[OF \a\s\ \x\s\, of j] p[rule_format, of x] \x \ s\ \x \ a\ show ?thesis by auto qed simp end locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t for p n b_s u_s s b_t u_t t begin lemma enum_eq: assumes l: "i \ l" "l \ j" and "j + d \ n" assumes eq: "s.enum ` {i .. j} = t.enum ` {i + d .. j + d}" shows "s.enum l = t.enum (l + d)" using l proof (induct l rule: dec_induct) case base then have s: "s.enum i \ t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \ s.enum ` {i .. j}" using eq by auto from t \i \ j\ \j + d \ n\ have "s.enum i \ t.enum (i + d)" by (auto simp: s.enum_mono) moreover from s \i \ j\ \j + d \ n\ have "t.enum (i + d) \ s.enum i" by (auto simp: t.enum_mono) ultimately show ?case by auto next case (step l) moreover from step.prems \j + d \ n\ have "s.enum l < s.enum (Suc l)" "t.enum (l + d) < t.enum (Suc l + d)" by (simp_all add: s.enum_strict_mono t.enum_strict_mono) moreover have "s.enum (Suc l) \ t.enum ` {i + d .. j + d}" "t.enum (Suc l + d) \ s.enum ` {i .. j}" using step \j + d \ n\ eq by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" using \j + d \ n\ by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) (auto intro!: s.enum_in t.enum_in) then show ?case by simp qed lemma ksimplex_eq_bot: assumes a: "a \ s" "\a'. a' \ s \ a \ a'" assumes b: "b \ t" "\b'. b' \ t \ b \ b'" assumes eq: "s - {a} = t - {b}" shows "s = t" proof cases assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp next assume "n \ 0" have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)" "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)" using \n \ 0\ by (simp_all add: s.enum_Suc t.enum_Suc) moreover have e0: "a = s.enum 0" "b = t.enum 0" using a b by (simp_all add: s.enum_0_bot t.enum_0_bot) moreover { fix j assume "0 < j" "j \ n" moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}" unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum j = t.enum j" using enum_eq[of "1" j n 0] eq by auto } note enum_eq = this then have "s.enum (Suc 0) = t.enum (Suc 0)" using \n \ 0\ by auto moreover { fix j assume "Suc j < n" with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"] have "u_s (Suc j) = u_t (Suc j)" using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"] by (auto simp: fun_eq_iff split: if_split_asm) } then have "\j. 0 < j \ j < n \ u_s j = u_t j" by (auto simp: gr0_conv_Suc) with \n \ 0\ have "u_t 0 = u_s 0" by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto ultimately have "a = b" by simp with assms show "s = t" by auto qed lemma ksimplex_eq_top: assumes a: "a \ s" "\a'. a' \ s \ a' \ a" assumes b: "b \ t" "\b'. b' \ t \ b' \ b" assumes eq: "s - {a} = t - {b}" shows "s = t" proof (cases n) assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp next case (Suc n') have "s.enum n = (s.enum n') (u_s n' := Suc (s.enum n' (u_s n')))" "t.enum n = (t.enum n') (u_t n' := Suc (t.enum n' (u_t n')))" using Suc by (simp_all add: s.enum_Suc t.enum_Suc) moreover have en: "a = s.enum n" "b = t.enum n" using a b by (simp_all add: s.enum_n_top t.enum_n_top) moreover { fix j assume "j < n" moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}" unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc) ultimately have "s.enum j = t.enum j" using enum_eq[of "0" j n' 0] eq Suc by auto } note enum_eq = this then have "s.enum n' = t.enum n'" using Suc by auto moreover { fix j assume "j < n'" with enum_eq[of j] enum_eq[of "Suc j"] have "u_s j = u_t j" using s.enum_Suc[of j] t.enum_Suc[of j] by (auto simp: Suc fun_eq_iff split: if_split_asm) } then have "\j. j < n' \ u_s j = u_t j" by (auto simp: gr0_conv_Suc) then have "u_t n' = u_s n'" by (intro bij_betw_singleton_eq[OF t.upd s.upd, of n']) (auto simp: Suc) ultimately have "a = b" by simp with assms show "s = t" by auto qed end inductive ksimplex for p n :: nat where ksimplex: "kuhn_simplex p n base upd s \ ksimplex p n s" lemma finite_ksimplexes: "finite {s. ksimplex p n s}" proof (rule finite_subset) { fix a s assume "ksimplex p n s" "a \ s" then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p n b u s . from s_space \a \ s\ out_eq_p[OF \a \ s\] have "a \ (\f x. if n \ x then p else f x) ` ({..< n} \\<^sub>E {.. p})" by (auto simp: image_iff subset_eq Pi_iff split: if_split_asm intro!: bexI[of _ "restrict a {..< n}"]) } then show "{s. ksimplex p n s} \ Pow ((\f x. if n \ x then p else f x) ` ({..< n} \\<^sub>E {.. p}))" by auto qed (simp add: finite_PiE) lemma ksimplex_card: assumes "ksimplex p n s" shows "card s = Suc n" using assms proof cases case (ksimplex u b) then interpret kuhn_simplex p n u b s . show ?thesis by (simp add: card_image s_eq inj_enum) qed lemma simplex_top_face: assumes "0 < p" "\x\s'. x n = p" shows "ksimplex p n s' \ (\s a. ksimplex p (Suc n) s \ a \ s \ s' = s - {a})" using assms proof safe fix s a assume "ksimplex p (Suc n) s" and a: "a \ s" and na: "\x\s - {a}. x n = p" then show "ksimplex p n (s - {a})" proof cases case (ksimplex base upd) then interpret kuhn_simplex p "Suc n" base upd "s" . have "a n < p" using one_step[of a n p] na \a\s\ s_space by (auto simp: less_le) then have "a = enum 0" using \a \ s\ na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n]) then have s_eq: "s - {a} = enum ` Suc ` {.. n}" using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident in_enum_image subset_eq) then have "enum 1 \ s - {a}" by auto then have "upd 0 = n" using \a n < p\ \a = enum 0\ na[rule_format, of "enum 1"] by (auto simp: fun_eq_iff enum_Suc split: if_split_asm) then have "bij_betw upd (Suc ` {..< n}) {..< n}" using upd by (subst notIn_Un_bij_betw3[where b=0]) (auto simp: lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) then have "bij_betw (upd\Suc) {..enum 1 \ s - {a}\] \a = enum 0\ by (auto simp: \upd 0 = n\) show ?thesis proof (rule ksimplex.intros, standard) show "bij_betw (upd\Suc) {..< n} {..< n}" by fact show "base(n := p) \ {.. {..i. n\i \ (base(n := p)) i = p" using base base_out by (auto simp: Pi_iff) have "\i. Suc ` {..< i} = {..< Suc i} - {0}" by (auto simp: image_iff Ball_def) arith then have upd_Suc: "\i. i \ n \ (upd\Suc) ` {..< i} = upd ` {..< Suc i} - {n}" using \upd 0 = n\ upd_inj by (auto simp add: image_iff less_Suc_eq_0_disj) have n_in_upd: "\i. n \ upd ` {..< Suc i}" using \upd 0 = n\ by auto define f' where "f' i j = (if j \ (upd\Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j { fix x i assume i [arith]: "i \ n" with upd_Suc have "(upd \ Suc) ` {..a n < p\ \a = enum 0\ \upd 0 = n\ \a n = p - 1\ have "enum (Suc i) x = f' i x" by (auto simp add: f'_def enum_def) } then show "s - {a} = f' ` {.. n}" unfolding s_eq image_comp by (intro image_cong) auto qed qed next assume "ksimplex p n s'" and *: "\x\s'. x n = p" then show "\s a. ksimplex p (Suc n) s \ a \ s \ s' = s - {a}" proof cases case (ksimplex base upd) then interpret kuhn_simplex p n base upd s' . define b where "b = base (n := p - 1)" define u where "u i = (case i of 0 \ n | Suc i \ upd i)" for i have "ksimplex p (Suc n) (s' \ {b})" proof (rule ksimplex.intros, standard) show "b \ {.. {..0 < p\ unfolding lessThan_Suc b_def by (auto simp: PiE_iff) show "\i. Suc n \ i \ b i = p" using base_out by (auto simp: b_def) have "bij_betw u (Suc ` {..< n} \ {0}) ({.. {u 0})" using upd by (intro notIn_Un_bij_betw) (auto simp: u_def bij_betw_def image_comp comp_def inj_on_def) then show "bij_betw u {.. u`{..< i} then Suc (b j) else b j)" for i j have u_eq: "\i. i \ n \ u ` {..< Suc i} = upd ` {..< i} \ { n }" by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith { fix x have "x \ n \ n \ upd ` {.. {0})" unfolding atMost_Suc_eq_insert_0 by simp also have "\ = (f' \ Suc) ` {.. n} \ {b}" by (auto simp: f'_def) also have "(f' \ Suc) ` {.. n} = s'" using \0 < p\ base_out[of n] unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd) finally show "s' \ {b} = f' ` {.. Suc n}" .. qed moreover have "b \ s'" using * \0 < p\ by (auto simp: b_def) ultimately show ?thesis by auto qed qed lemma ksimplex_replace_0: assumes s: "ksimplex p n s" and a: "a \ s" assumes j: "j < n" and p: "\x\s - {a}. x j = 0" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 1" using s proof cases case (ksimplex b_s u_s) { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t by intro_locales fact+ assume b: "b \ t" "t - {b} = s - {a}" with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t" by (intro ksimplex_eq_top[of a b]) auto } then have "{s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = {s}" using s \a \ s\ by auto then show ?thesis by simp qed lemma ksimplex_replace_1: assumes s: "ksimplex p n s" and a: "a \ s" assumes j: "j < n" and p: "\x\s - {a}. x j = p" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 1" using s proof cases case (ksimplex b_s u_s) { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t by intro_locales fact+ assume b: "b \ t" "t - {b} = s - {a}" with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t" by (intro ksimplex_eq_bot[of a b]) auto } then have "{s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = {s}" using s \a \ s\ by auto then show ?thesis by simp qed lemma ksimplex_replace_2: assumes s: "ksimplex p n s" and "a \ s" and "n \ 0" and lb: "\jx\s - {a}. x j \ 0" and ub: "\jx\s - {a}. x j \ p" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 2" using s proof cases case (ksimplex base upd) then interpret kuhn_simplex p n base upd s . from \a \ s\ obtain i where "i \ n" "a = enum i" unfolding s_eq by auto from \i \ n\ have "i = 0 \ i = n \ (0 < i \ i < n)" by linarith then have "\!s'. s' \ s \ ksimplex p n s' \ (\b\s'. s - {a} = s'- {b})" proof (elim disjE conjE) assume "i = 0" define rot where [abs_def]: "rot i = (if i + 1 = n then 0 else i + 1)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" by (auto simp: bij_betw_def inj_on_def image_iff Ball_def rot_def) arith+ from rot upd have "bij_betw ?upd {.. ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \ rot" "f' ` {.. n}" proof from \a = enum i\ ub \n \ 0\ \i = 0\ obtain i' where "i' \ n" "enum i' \ enum 0" "enum i' (upd 0) \ p" unfolding s_eq by (auto intro: upd_space simp: enum_inj) then have "enum 1 \ enum i'" "enum i' (upd 0) < p" using enum_le_p[of i' "upd 0"] by (auto simp: enum_inj enum_mono upd_space) then have "enum 1 (upd 0) < p" by (auto simp: le_fun_def intro: le_less_trans) then show "enum (Suc 0) \ {.. {..n \ 0\ by (auto simp: enum_0 enum_Suc PiE_iff extensional_def upd_space) { fix i assume "n \ i" then show "enum (Suc 0) i = p" using \n \ 0\ by (auto simp: enum_eq_p) } show "bij_betw ?upd {..n \ 0\ show ?thesis by (simp only: f'_def enum_def fun_eq_iff image_comp [symmetric]) (auto simp add: upd_inj) qed then have "enum ` Suc ` {..< n} = f' ` {..< n}" by (force simp: enum_inj) also have "Suc ` {..< n} = {.. n} - {0}" by (auto simp: image_iff Ball_def) arith also have "{..< n} = {.. n} - {n}" by auto finally have eq: "s - {a} = f' ` {.. n} - {f' n}" unfolding s_eq \a = enum i\ \i = 0\ by (simp add: inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f']) have "enum 0 < f' 0" using \n \ 0\ by (simp add: enum_strict_mono f'_eq_enum) also have "\ < f' n" using \n \ 0\ b.enum_strict_mono[of 0 n] unfolding b_enum by simp finally have "a \ f' n" using \a = enum i\ \i = 0\ by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b u where "kuhn_simplex p n b u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b u t . { fix x assume "x \ s" "x \ a" then have "x (upd 0) = enum (Suc 0) (upd 0)" by (auto simp: \a = enum i\ \i = 0\ s_eq enum_def enum_inj) } then have eq_upd0: "\x\t-{c}. x (upd 0) = enum (Suc 0) (upd 0)" unfolding eq_sma[symmetric] by auto then have "c (upd 0) \ enum (Suc 0) (upd 0)" using \n \ 0\ by (intro t.one_step[OF \c\t\ ]) (auto simp: upd_space) then have "c (upd 0) < enum (Suc 0) (upd 0) \ c (upd 0) > enum (Suc 0) (upd 0)" by auto then have "t = s \ t = f' ` {..n}" proof (elim disjE conjE) assume *: "c (upd 0) < enum (Suc 0) (upd 0)" interpret st: kuhn_simplex_pair p n base upd s b u t .. { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "c \ x" by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } note top = this have "s = t" using \a = enum i\ \i = 0\ \c \ t\ by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma]) (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) then show ?thesis by simp next assume *: "c (upd 0) > enum (Suc 0) (upd 0)" interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \ rot" "f' ` {.. n}" b u t .. have eq: "f' ` {..n} - {f' n} = t - {c}" using eq_sma eq by simp { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "x \ c" by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } note top = this have "f' ` {..n} = t" using \a = enum i\ \i = 0\ \c \ t\ by (intro st.ksimplex_eq_top[OF _ _ _ _ eq]) (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top) then show ?thesis by simp qed } with ks_f' eq \a \ f' n\ \n \ 0\ show ?thesis apply (intro ex1I[of _ "f' ` {.. n}"]) apply auto [] apply metis done next assume "i = n" from \n \ 0\ obtain n' where n': "n = Suc n'" by (cases n) auto define rot where "rot i = (case i of 0 \ n' | Suc i \ i)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" by (auto simp: bij_betw_def inj_on_def image_iff Bex_def rot_def n' split: nat.splits) arith from rot upd have "bij_betw ?upd {.. ?upd`{..< i} then Suc (b j) else b j)" for i j interpret b: kuhn_simplex p n b "upd \ rot" "f' ` {.. n}" proof { fix i assume "n \ i" then show "b i = p" using base_out[of i] upd_space[of n'] by (auto simp: b_def n') } show "b \ {.. {..n \ 0\ upd_space[of n'] by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n') show "bij_betw ?upd {..n \ 0\ by auto { from \a = enum i\ \n \ 0\ \i = n\ lb upd_space[of n'] obtain i' where "i' \ n" "enum i' \ enum n" "0 < enum i' (upd n')" unfolding s_eq by (auto simp: enum_inj n') moreover have "enum i' (upd n') = base (upd n')" unfolding enum_def using \i' \ n\ \enum i' \ enum n\ by (auto simp: n' upd_inj enum_inj) ultimately have "0 < base (upd n')" by auto } then have benum1: "b.enum (Suc 0) = base" unfolding b.enum_Suc[OF \0] b.enum_0 by (auto simp: b_def rot_def) have [simp]: "\j. Suc j < n \ rot ` {..< Suc j} = {n'} \ {..< j}" by (auto simp: rot_def image_iff Ball_def split: nat.splits) have rot_simps: "\j. rot (Suc j) = j" "rot 0 = n'" by (simp_all add: rot_def) { fix j assume j: "Suc j \ n" then have "b.enum (Suc j) = enum j" by (induct j) (auto simp: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) } note b_enum_eq_enum = this then have "enum ` {..< n} = b.enum ` Suc ` {..< n}" by (auto simp: image_comp intro!: image_cong) also have "Suc ` {..< n} = {.. n} - {0}" by (auto simp: image_iff Ball_def) arith also have "{..< n} = {.. n} - {n}" by auto finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}" unfolding s_eq \a = enum i\ \i = n\ using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"] inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"] by (simp add: comp_def) have "b.enum 0 \ b.enum n" by (simp add: b.enum_mono) also have "b.enum n < enum n" using \n \ 0\ by (simp add: enum_strict_mono b_enum_eq_enum n') finally have "a \ b.enum 0" using \a = enum i\ \i = n\ by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b' u where "kuhn_simplex p n b' u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b' u t . { fix x assume "x \ s" "x \ a" then have "x (upd n') = enum n' (upd n')" by (auto simp: \a = enum i\ n' \i = n\ s_eq enum_def enum_inj in_upd_image) } then have eq_upd0: "\x\t-{c}. x (upd n') = enum n' (upd n')" unfolding eq_sma[symmetric] by auto then have "c (upd n') \ enum n' (upd n')" using \n \ 0\ by (intro t.one_step[OF \c\t\ ]) (auto simp: n' upd_space[unfolded n']) then have "c (upd n') < enum n' (upd n') \ c (upd n') > enum n' (upd n')" by auto then have "t = s \ t = b.enum ` {..n}" proof (elim disjE conjE) assume *: "c (upd n') > enum n' (upd n')" interpret st: kuhn_simplex_pair p n base upd s b' u t .. { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "x \ c" by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } note top = this have "s = t" using \a = enum i\ \i = n\ \c \ t\ by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma]) (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) then show ?thesis by simp next assume *: "c (upd n') < enum n' (upd n')" interpret st: kuhn_simplex_pair p n b "upd \ rot" "f' ` {.. n}" b' u t .. have eq: "f' ` {..n} - {b.enum 0} = t - {c}" using eq_sma eq f' by simp { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "c \ x" by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } note bot = this have "f' ` {..n} = t" using \a = enum i\ \i = n\ \c \ t\ by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq]) (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot) with f' show ?thesis by simp qed } with ks_f' eq \a \ b.enum 0\ \n \ 0\ show ?thesis apply (intro ex1I[of _ "b.enum ` {.. n}"]) apply auto [] apply metis done next assume i: "0 < i" "i < n" define i' where "i' = i - 1" with i have "Suc i' < n" by simp with i have Suc_i': "Suc i' = i" by (simp add: i'_def) let ?upd = "Fun.swap i' i upd" from i upd have "bij_betw ?upd {..< n} {..< n}" by (subst bij_betw_swap_iff) (auto simp: i'_def) define f' where [abs_def]: "f' i j = (if j \ ?upd`{..< i} then Suc (base j) else base j)" for i j interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}" proof show "base \ {.. {.. i" then show "base i = p" by (rule base_out) } show "bij_betw ?upd {.. {..n}" using i by auto { fix j assume "j \ n" moreover have "j < i \ i = j \ i < j" by arith moreover note i ultimately have "enum j = b.enum j \ j \ i" unfolding enum_def[abs_def] b.enum_def[abs_def] by (auto simp: fun_eq_iff swap_image i'_def in_upd_image inj_on_image_set_diff[OF inj_upd]) } note enum_eq_benum = this then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})" by (intro image_cong) auto then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}" unfolding s_eq \a = enum i\ using inj_on_image_set_diff[OF inj_enum Diff_subset \{i} \ {..n}\] inj_on_image_set_diff[OF b.inj_enum Diff_subset \{i} \ {..n}\] by (simp add: comp_def) have "a \ b.enum i" using \a = enum i\ enum_eq_benum i by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b' u where "kuhn_simplex p n b' u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b' u t . have "enum i' \ s - {a}" "enum (i + 1) \ s - {a}" using \a = enum i\ i enum_in by (auto simp: enum_inj i'_def) then obtain l k where l: "t.enum l = enum i'" "l \ n" "t.enum l \ c" and k: "t.enum k = enum (i + 1)" "k \ n" "t.enum k \ c" unfolding eq_sma by (auto simp: t.s_eq) with i have "t.enum l < t.enum k" by (simp add: enum_strict_mono i'_def) with \l \ n\ \k \ n\ have "l < k" by (simp add: t.enum_strict_mono) { assume "Suc l = k" have "enum (Suc (Suc i')) = t.enum (Suc l)" using i by (simp add: k \Suc l = k\ i'_def) then have False using \l < k\ \k \ n\ \Suc i' < n\ by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: if_split_asm) (metis Suc_lessD n_not_Suc_n upd_inj) } with \l < k\ have "Suc l < k" by arith have c_eq: "c = t.enum (Suc l)" proof (rule ccontr) assume "c \ t.enum (Suc l)" then have "t.enum (Suc l) \ s - {a}" using \l < k\ \k \ n\ by (simp add: t.s_eq eq_sma) then obtain j where "t.enum (Suc l) = enum j" "j \ n" "enum j \ enum i" unfolding s_eq \a = enum i\ by auto with i have "t.enum (Suc l) \ t.enum l \ t.enum k \ t.enum (Suc l)" by (auto simp: i'_def enum_mono enum_inj l k) with \Suc l < k\ \k \ n\ show False by (simp add: t.enum_mono) qed { have "t.enum (Suc (Suc l)) \ s - {a}" unfolding eq_sma c_eq t.s_eq using \Suc l < k\ \k \ n\ by (auto simp: t.enum_inj) then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \ n" "j \ i" by (auto simp: s_eq \a = enum i\) moreover have "enum i' < t.enum (Suc (Suc l))" unfolding l(1)[symmetric] using \Suc l < k\ \k \ n\ by (auto simp: t.enum_strict_mono) ultimately have "i' < j" using i by (simp add: enum_strict_mono i'_def) with \j \ i\ \j \ n\ have "t.enum k \ t.enum (Suc (Suc l))" unfolding i'_def by (simp add: enum_mono k eq) then have "k \ Suc (Suc l)" using \k \ n\ \Suc l < k\ by (simp add: t.enum_mono) } with \Suc l < k\ have "Suc (Suc l) = k" by simp then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))" using i by (simp add: k i'_def) also have "\ = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" using \Suc l < k\ \k \ n\ by (simp add: t.enum_Suc l t.upd_inj) finally have "(u l = upd i' \ u (Suc l) = upd (Suc i')) \ (u l = upd (Suc i') \ u (Suc l) = upd i')" using \Suc i' < n\ by (auto simp: enum_Suc fun_eq_iff split: if_split_asm) then have "t = s \ t = b.enum ` {..n}" proof (elim disjE conjE) assume u: "u l = upd i'" have "c = t.enum (Suc l)" unfolding c_eq .. also have "t.enum (Suc l) = enum (Suc i')" using u \l < k\ \k \ n\ \Suc i' < n\ by (simp add: enum_Suc t.enum_Suc l) also have "\ = a" using \a = enum i\ i by (simp add: i'_def) finally show ?thesis using eq_sma \a \ s\ \c \ t\ by auto next assume u: "u l = upd (Suc i')" define B where "B = b.enum ` {..n}" have "b.enum i' = enum i'" using enum_eq_benum[of i'] i by (auto simp: i'_def gr0_conv_Suc) have "c = t.enum (Suc l)" unfolding c_eq .. also have "t.enum (Suc l) = b.enum (Suc i')" using u \l < k\ \k \ n\ \Suc i' < n\ by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \b.enum i' = enum i'\) (simp add: Suc_i') also have "\ = b.enum i" using i by (simp add: i'_def) finally have "c = b.enum i" . then have "t - {c} = B - {c}" "c \ B" unfolding eq_sma[symmetric] eq B_def using i by auto with \c \ t\ have "t = B" by auto then show ?thesis by (simp add: B_def) qed } with ks_f' eq \a \ b.enum i\ \n \ 0\ \i \ n\ show ?thesis apply (intro ex1I[of _ "b.enum ` {.. n}"]) apply auto [] apply metis done qed then show ?thesis using s \a \ s\ by (simp add: card_2_iff' Ex1_def) metis qed text \Hence another step towards concreteness.\ lemma kuhn_simplex_lemma: assumes "\s. ksimplex p (Suc n) s \ rl ` s \ {.. Suc n}" and "odd (card {f. \s a. ksimplex p (Suc n) s \ a \ s \ (f = s - {a}) \ rl ` f = {..n} \ ((\j\n. \x\f. x j = 0) \ (\j\n. \x\f. x j = p))})" shows "odd (card {s. ksimplex p (Suc n) s \ rl ` s = {..Suc n}})" proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq, where bnd="\f. (\j\{..n}. \x\f. x j = 0) \ (\j\{..n}. \x\f. x j = p)"], safe del: notI) have *: "\x y. x = y \ odd (card x) \ odd (card y)" by auto show "odd (card {f. (\s\{s. ksimplex p (Suc n) s}. \a\s. f = s - {a}) \ rl ` f = {..n} \ ((\j\{..n}. \x\f. x j = 0) \ (\j\{..n}. \x\f. x j = p))})" apply (rule *[OF _ assms(2)]) apply (auto simp: atLeast0AtMost) done next fix s assume s: "ksimplex p (Suc n) s" then show "card s = n + 2" by (simp add: ksimplex_card) fix a assume a: "a \ s" then show "rl a \ Suc n" using assms(1) s by (auto simp: subset_eq) let ?S = "{t. ksimplex p (Suc n) t \ (\b\t. s - {a} = t - {b})}" { fix j assume j: "j \ n" "\x\s - {a}. x j = 0" with s a show "card ?S = 1" using ksimplex_replace_0[of p "n + 1" s a j] by (subst eq_commute) simp } { fix j assume j: "j \ n" "\x\s - {a}. x j = p" with s a show "card ?S = 1" using ksimplex_replace_1[of p "n + 1" s a j] by (subst eq_commute) simp } { assume "card ?S \ 2" "\ (\j\{..n}. \x\s - {a}. x j = p)" with s a show "\j\{..n}. \x\s - {a}. x j = 0" using ksimplex_replace_2[of p "n + 1" s a] by (subst (asm) eq_commute) auto } qed subsubsection \Reduced labelling\ definition reduced :: "nat \ (nat \ nat) \ nat" where "reduced n x = (LEAST k. k = n \ x k \ 0)" lemma reduced_labelling: shows "reduced n x \ n" and "\i x (reduced n x) \ 0" proof - show "reduced n x \ n" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto show "\i x (reduced n x) \ 0" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ qed lemma reduced_labelling_unique: "r \ n \ \i r = n \ x r \ 0 \ reduced n x = r" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) (metis le_less not_le)+ lemma reduced_labelling_zero: "j < n \ x j = 0 \ reduced n x \ j" using reduced_labelling[of n x] by auto lemma reduce_labelling_zero[simp]: "reduced 0 x = 0" by (rule reduced_labelling_unique) auto lemma reduced_labelling_nonzero: "j < n \ x j \ 0 \ reduced n x \ j" using reduced_labelling[of n x] by (elim allE[where x=j]) auto lemma reduced_labelling_Suc: "reduced (Suc n) x \ Suc n \ reduced (Suc n) x = reduced n x" using reduced_labelling[of "Suc n" x] by (intro reduced_labelling_unique[symmetric]) auto lemma complete_face_top: assumes "\x\f. \j\n. x j = 0 \ lab x j = 0" and "\x\f. \j\n. x j = p \ lab x j = 1" and eq: "(reduced (Suc n) \ lab) ` f = {..n}" shows "((\j\n. \x\f. x j = 0) \ (\j\n. \x\f. x j = p)) \ (\x\f. x n = p)" proof (safe del: disjCI) fix x j assume j: "j \ n" "\x\f. x j = 0" { fix x assume "x \ f" with assms j have "reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_zero) auto } moreover have "j \ (reduced (Suc n) \ lab) ` f" using j eq by auto ultimately show "x n = p" by force next fix x j assume j: "j \ n" "\x\f. x j = p" and x: "x \ f" have "j = n" proof (rule ccontr) assume "\ ?thesis" { fix x assume "x \ f" with assms j have "reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_nonzero) auto then have "reduced (Suc n) (lab x) \ n" using \j \ n\ \j \ n\ by simp } moreover have "n \ (reduced (Suc n) \ lab) ` f" using eq by auto ultimately show False by force qed moreover have "j \ (reduced (Suc n) \ lab) ` f" using j eq by auto ultimately show "x n = p" using j x by auto qed auto text \Hence we get just about the nice induction.\ lemma kuhn_induction: assumes "0 < p" and lab_0: "\x. \j\n. (\j. x j \ p) \ x j = 0 \ lab x j = 0" and lab_1: "\x. \j\n. (\j. x j \ p) \ x j = p \ lab x j = 1" and odd: "odd (card {s. ksimplex p n s \ (reduced n\lab) ` s = {..n}})" shows "odd (card {s. ksimplex p (Suc n) s \ (reduced (Suc n)\lab) ` s = {..Suc n}})" proof - let ?rl = "reduced (Suc n) \ lab" and ?ext = "\f v. \j\n. \x\f. x j = v" let ?ext = "\s. (\j\n. \x\s. x j = 0) \ (\j\n. \x\s. x j = p)" have "\s. ksimplex p (Suc n) s \ ?rl ` s \ {..Suc n}" by (simp add: reduced_labelling subset_eq) moreover have "{s. ksimplex p n s \ (reduced n \ lab) ` s = {..n}} = {f. \s a. ksimplex p (Suc n) s \ a \ s \ f = s - {a} \ ?rl ` f = {..n} \ ?ext f}" proof (intro set_eqI, safe del: disjCI equalityI disjE) fix s assume s: "ksimplex p n s" and rl: "(reduced n \ lab) ` s = {..n}" from s obtain u b where "kuhn_simplex p n u b s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p n u b s . have all_eq_p: "\x\s. x n = p" by (auto simp: out_eq_p) moreover { fix x assume "x \ s" with lab_1[rule_format, of n x] all_eq_p s_le_p[of x] have "?rl x \ n" by (auto intro!: reduced_labelling_nonzero) then have "?rl x = reduced n (lab x)" by (auto intro!: reduced_labelling_Suc) } then have "?rl ` s = {..n}" using rl by (simp cong: image_cong) moreover obtain t a where "ksimplex p (Suc n) t" "a \ t" "s = t - {a}" using s unfolding simplex_top_face[OF \0 < p\ all_eq_p] by auto ultimately show "\t a. ksimplex p (Suc n) t \ a \ t \ s = t - {a} \ ?rl ` s = {..n} \ ?ext s" by auto next fix x s a assume s: "ksimplex p (Suc n) s" and rl: "?rl ` (s - {a}) = {.. n}" and a: "a \ s" and "?ext (s - {a})" from s obtain u b where "kuhn_simplex p (Suc n) u b s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p "Suc n" u b s . have all_eq_p: "\x\s. x (Suc n) = p" by (auto simp: out_eq_p) { fix x assume "x \ s - {a}" then have "?rl x \ ?rl ` (s - {a})" by auto then have "?rl x \ n" unfolding rl by auto then have "?rl x = reduced n (lab x)" by (auto intro!: reduced_labelling_Suc) } then show rl': "(reduced n\lab) ` (s - {a}) = {..n}" unfolding rl[symmetric] by (intro image_cong) auto from \?ext (s - {a})\ have all_eq_p: "\x\s - {a}. x n = p" proof (elim disjE exE conjE) fix j assume "j \ n" "\x\s - {a}. x j = 0" with lab_0[rule_format, of j] all_eq_p s_le_p have "\x. x \ s - {a} \ reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_zero) auto moreover have "j \ ?rl ` (s - {a})" using \j \ n\ unfolding rl by auto ultimately show ?thesis by force next fix j assume "j \ n" and eq_p: "\x\s - {a}. x j = p" show ?thesis proof cases assume "j = n" with eq_p show ?thesis by simp next assume "j \ n" { fix x assume x: "x \ s - {a}" have "reduced n (lab x) \ j" proof (rule reduced_labelling_nonzero) show "lab x j \ 0" using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \j \ n\ by auto show "j < n" using \j \ n\ \j \ n\ by simp qed then have "reduced n (lab x) \ n" using \j \ n\ \j \ n\ by simp } moreover have "n \ (reduced n\lab) ` (s - {a})" unfolding rl' by auto ultimately show ?thesis by force qed qed show "ksimplex p n (s - {a})" unfolding simplex_top_face[OF \0 < p\ all_eq_p] using s a by auto qed ultimately show ?thesis using assms by (intro kuhn_simplex_lemma) auto qed text \And so we get the final combinatorial result.\ lemma ksimplex_0: "ksimplex p 0 s \ s = {(\x. p)}" proof assume "ksimplex p 0 s" then show "s = {(\x. p)}" by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases) next assume s: "s = {(\x. p)}" show "ksimplex p 0 s" proof (intro ksimplex, unfold_locales) show "(\_. p) \ {..<0::nat} \ {..x j. (\j. x j \ p) \ j < n \ x j = 0 \ lab x j = 0" and "\x j. (\j. x j \ p) \ j < n \ x j = p \ lab x j = 1" shows "odd (card {s. ksimplex p n s \ (reduced n\lab) ` s = {..n}})" (is "odd (card (?M n))") using assms proof (induct n) case 0 then show ?case by (simp add: ksimplex_0 cong: conj_cong) next case (Suc n) then have "odd (card (?M n))" by force with Suc show ?case using kuhn_induction[of p n] by (auto simp: comp_def) qed lemma kuhn_lemma: fixes n p :: nat assumes "0 < p" and "\x. (\i p) \ (\i label x i = 1)" and "\x. (\i p) \ (\i label x i = 0)" and "\x. (\i p) \ (\i label x i = 1)" obtains q where "\iir s. (\j r j \ r j \ q j + 1) \ (\j s j \ s j \ q j + 1) \ label r i \ label s i" proof - let ?rl = "reduced n \ label" let ?A = "{s. ksimplex p n s \ ?rl ` s = {..n}}" have "odd (card ?A)" using assms by (intro kuhn_combinatorial[of p n label]) auto then have "?A \ {}" by (rule odd_card_imp_not_empty) then obtain s b u where "kuhn_simplex p n b u s" and rl: "?rl ` s = {..n}" by (auto elim: ksimplex.cases) interpret kuhn_simplex p n b u s by fact show ?thesis proof (intro that[of b] allI impI) fix i assume "i < n" then show "b i < p" using base by auto next fix i assume "i < n" then have "i \ {.. n}" "Suc i \ {.. n}" by auto then obtain u v where u: "u \ s" "Suc i = ?rl u" and v: "v \ s" "i = ?rl v" unfolding rl[symmetric] by blast have "label u i \ label v i" using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"] u(2)[symmetric] v(2)[symmetric] \i < n\ by auto moreover have "b j \ u j" "u j \ b j + 1" "b j \ v j" "v j \ b j + 1" if "j < n" for j using that base_le[OF \u\s\] le_Suc_base[OF \u\s\] base_le[OF \v\s\] le_Suc_base[OF \v\s\] by auto ultimately show "\r s. (\j r j \ r j \ b j + 1) \ (\j s j \ s j \ b j + 1) \ label r i \ label s i" by blast qed qed subsubsection \Main result for the unit cube\ lemma kuhn_labelling_lemma': assumes "(\x::nat\real. P x \ P (f x))" and "\x. P x \ (\i::nat. Q i \ 0 \ x i \ x i \ 1)" shows "\l. (\x i. l x i \ (1::nat)) \ (\x i. P x \ Q i \ x i = 0 \ l x i = 0) \ (\x i. P x \ Q i \ x i = 1 \ l x i = 1) \ (\x i. P x \ Q i \ l x i = 0 \ x i \ f x i) \ (\x i. P x \ Q i \ l x i = 1 \ f x i \ x i)" proof - have and_forall_thm: "\P Q. (\x. P x) \ (\x. Q x) \ (\x. P x \ Q x)" by auto have *: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 \ x \ 1 \ x \ y \ x \ 0 \ y \ x" by auto show ?thesis unfolding and_forall_thm apply (subst choice_iff[symmetric])+ apply rule apply rule proof - fix x x' let ?R = "\y::nat. (P x \ Q x' \ x x' = 0 \ y = 0) \ (P x \ Q x' \ x x' = 1 \ y = 1) \ (P x \ Q x' \ y = 0 \ x x' \ (f x) x') \ (P x \ Q x' \ y = 1 \ (f x) x' \ x x')" have "0 \ f x x' \ f x x' \ 1" if "P x" "Q x'" using assms(2)[rule_format,of "f x" x'] that apply (drule_tac assms(1)[rule_format]) apply auto done then have "?R 0 \ ?R 1" by auto then show "\y\1. ?R y" by auto qed qed subsection \Brouwer's fixed point theorem\ 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 (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) assume "\ ?thesis" then have *: "\ (\x\cbox 0 One. f x - x = 0)" by auto obtain d where 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 \ 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 cbox_def by auto obtain label :: "'a \ 'a \ nat" where label [rule_format]: "\x. \i\Basis. label x i \ 1" "\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\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 \ 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 \ \fx - x\ \ \fy - fx\ + \y - x\" by auto have "\(f x - x) \ i\ \ \(f y - f x)\i\ + \(y - x)\i\" proof (cases "label x i = 0") case True then have fxy: "\ f y \ i \ y \ i \ f x \ i \ x \ i" by (metis True i label(1) label(5) le_antisym less_one not_le_imp_less y) show ?thesis unfolding inner_simps by (rule *) (auto simp: True i label x y fxy) next case False then show ?thesis using label [OF \i \ Basis\] i(1) x y apply (auto simp: inner_diff_left le_Suc_eq) by (metis "*") qed also have "\ \ norm (f y - f x) + norm (y - x)" by (simp add: add_mono i(2) norm_bound_Basis_le) finally show "\f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" unfolding inner_simps . qed 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) 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 \ 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'] unfolding dist_norm by blast show ?thesis proof (intro exI conjI ballI impI) show "0 < min (e / 2) (d / real n / 8)" using d' e by auto fix x y z i assume as: "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" assume i: "i \ Basis" have *: "\z fz x fx n1 n2 n3 n4 d4 d :: real. \fx - x\ \ n1 + n2 \ \fx - fz\ \ n3 \ \x - z\ \ n4 \ n1 < d4 \ n2 < 2 * d4 \ n3 < d4 \ n4 < d4 \ (8 * d4 = d) \ \fz - z\ < d" by auto show "\(f z - z) \ i\ < d / real n" unfolding inner_simps proof (rule *) show "\f x \ i - x \ i\ \ norm (f y -f x) + norm (y - x)" using as(1) as(2) as(6) i lem1 by blast show "norm (f x - f z) < d / real n / 8" using d' e as by auto show "\f x \ i - f z \ i\ \ norm (f x - f z)" "\x \ i - z \ i\ \ norm (x - z)" unfolding inner_diff_left[symmetric] by (rule Basis_le_norm[OF i])+ have tria: "norm (y - x) \ norm (y - z) + norm (x - z)" using dist_triangle[of y x z, unfolded dist_norm] unfolding norm_minus_commute by auto also have "\ < e / 2 + e / 2" using as(4) as(5) by auto finally show "norm (f y - f x) < d / real n / 8" using as(1) as(2) e(2) by auto have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" using as(4) as(5) by auto with tria show "norm (y - x) < 2 * (d / real n / 8)" by auto qed (use as in auto) qed qed then obtain e where e: "e > 0" "\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" by blast obtain p :: nat where p: "1 + real n / e \ real p" using real_arch_simple .. have "1 + real n / e > 0" using e(1) n by (simp add: add_pos_pos) then have "p > 0" using p by auto obtain b :: "nat \ 'a" where b: "bij_betw b {..< n} Basis" by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) define b' where "b' = inv_into {..< n} b" then have b': "bij_betw b' Basis {..< n}" using bij_betw_inv_into[OF b] by auto then have b'_Basis: "\i. i \ Basis \ b' i \ {..< n}" unfolding bij_betw_def by (auto simp: set_eq_iff) have bb'[simp]:"\i. i \ Basis \ b (b' i) = i" unfolding b'_def using b by (auto simp: f_inv_into_f bij_betw_def) have b'b[simp]:"\i. i < n \ b' (b i) = i" unfolding b'_def using b by (auto simp: inv_into_f_eq bij_betw_def) have *: "\x :: nat. x = 0 \ x = 1 \ x \ 1" by auto have b'': "\j. j < n \ b j \ Basis" using b unfolding bij_betw_def by auto have q1: "0 < p" "\x. (\i p) \ (\ii\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0 \ (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" unfolding * using \p > 0\ \n > 0\ 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) \ (cbox 0 One::'a set)" using b'_Basis by (auto simp: cbox_def 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)" unfolding o_def using cube \p > 0\ by (intro allI impI label(2)) (auto simp: b'') have q3: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" using cube \p > 0\ unfolding o_def by (intro allI impI label(3)) (auto simp: b'') obtain q where q: "\iir s. (\j r j \ r j \ q j + 1) \ (\j s j \ s j \ q j + 1) \ (label (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i) \ b) i \ (label (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i) \ b) i" by (rule kuhn_lemma[OF q1 q2 q3]) define z :: 'a where "z = (\i\Basis. (real (q (b' i)) / real p) *\<^sub>R i)" have "\i\Basis. d / real n \ \(f z - z)\i\" proof (rule ccontr) 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 \ 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)" by (rule d) assume "\ ?thesis" then have as: "\i\Basis. \f z \ i - z \ i\ < d / real n" using \n > 0\ by (auto simp: not_le inner_diff) have "norm (f z - z) \ (\i\Basis. \f z \ i - z \ i\)" unfolding inner_diff_left[symmetric] by (rule norm_le_l1) also have "\ < (\(i::'a) \ Basis. d / real n)" by (meson as finite_Basis nonempty_Basis sum_strict_mono) also have "\ = d" using DIM_positive[where 'a='a] by (auto simp: n_def) finally show False using d_fz_z by auto qed then obtain i where i: "i \ Basis" "d / real n \ \(f z - z) \ i\" .. have *: "b' i < n" using i and b'[unfolded bij_betw_def] by auto obtain r s where rs: "\j. j < n \ q j \ r j \ r j \ q j + 1" "\j. j < n \ q j \ s j \ s j \ q j + 1" "(label (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i) \ b) (b' i) \ (label (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i) \ b) (b' i)" using q(2)[rule_format,OF *] by blast have b'_im: "\i. i \ Basis \ b' i < n" using b' unfolding bij_betw_def by auto define r' ::'a where "r' = (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ r (b' i) \ p" apply (rule order_trans) apply (rule rs(1)[OF b'_im,THEN conjunct2]) using q(1)[rule_format,OF b'_im] apply (auto simp: Suc_le_eq) done 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' \ 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 \ 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) { have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" by (rule sum_mono) (use rs(1)[OF b'_im] in force) also have "\ < e * real p" using p \e > 0\ \p > 0\ by (auto simp: field_simps n_def) finally have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) < e * real p" . } moreover { have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" by (rule sum_mono) (use rs(2)[OF b'_im] in force) also have "\ < e * real p" using p \e > 0\ \p > 0\ by (auto simp: field_simps n_def) finally have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) < e * real p" . } ultimately have "norm (r' - z) < e" and "norm (s' - z) < e" unfolding r'_def s'_def z_def using \p > 0\ apply (rule_tac[!] le_less_trans[OF norm_le_l1]) apply (auto simp: field_simps sum_divide_distrib[symmetric] inner_diff_left) done 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'\cbox 0 One\ \s'\cbox 0 One\ \z\cbox 0 One\]) auto then show False using i by auto qed text \Next step is to prove it for nonempty interiors.\ lemma brouwer_weak: fixes f :: "'a::euclidean_space \ 'a" assumes "compact S" and "convex S" and "interior S \ {}" and "continuous_on S f" and "f ` S \ S" obtains x where "x \ S" and "f x = x" proof - 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})" show "open ?I" by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner) show "\Basis /\<^sub>R 2 \ ?I" by simp 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_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 then show ?thesis unfolding homeomorphic_fixpoint_property[OF *] using assms by (auto intro: that) qed text \Then the particular case for closed balls.\ lemma brouwer_ball: fixes f :: "'a::euclidean_space \ 'a" assumes "e > 0" and "continuous_on (cball a e) f" and "f ` cball a e \ cball a e" obtains x where "x \ cball a e" and "f x = x" using brouwer_weak[OF compact_cball convex_cball, of a e f] unfolding interior_cball ball_eq_empty using assms by auto 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" and fim: "f ` S \ S" obtains x where "x \ S" and "f x = x" proof - have "\e>0. S \ cball 0 e" using compact_imp_bounded[OF \compact S\] unfolding bounded_pos by auto then obtain e where e: "e > 0" "S \ cball 0 e" by blast have "\x\ cball 0 e. (f \ closest_point S) x = x" proof (rule_tac brouwer_ball[OF e(1)]) show "continuous_on (cball 0 e) (f \ closest_point S)" apply (rule continuous_on_compose) using S compact_eq_bounded_closed continuous_on_closest_point apply blast by (meson S contf closest_point_in_set compact_imp_closed continuous_on_subset image_subsetI) show "(f \ closest_point S) ` cball 0 e \ cball 0 e" by clarsimp (metis S fim closest_point_exists(1) compact_eq_bounded_closed e(2) image_subset_iff mem_cball_0 subsetCE) qed (use assms in auto) then obtain x where x: "x \ cball 0 e" "(f \ closest_point S) x = x" .. have "x \ S" by (metis closest_point_in_set comp_apply compact_imp_closed fim image_eqI S(1) S(3) subset_iff x(2)) then have *: "closest_point S x = x" by (rule closest_point_self) show thesis proof show "closest_point S x \ S" by (simp add: "*" \x \ S\) show "f (closest_point S x) = closest_point S x" using "*" x(2) by auto qed qed subsection \Applications\ text \So we get the no-retraction theorem.\ corollary no_retraction_cball: fixes a :: "'a::euclidean_space" assumes "e > 0" shows "\ (frontier (cball a e) retract_of (cball a e))" proof assume *: "frontier (cball a e) retract_of (cball a e)" have **: "\xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" using scaleR_left_distrib[of 1 1 a] by auto obtain x where x: "x \ {x. norm (a - x) = e}" "2 *\<^sub>R a - x = x" proof (rule retract_fixpoint_property[OF *, of "\x. scaleR 2 a - x"]) show "continuous_on (frontier (cball a e)) ((-) (2 *\<^sub>R a))" by (intro continuous_intros) show "(-) (2 *\<^sub>R a) ` frontier (cball a e) \ frontier (cball a e)" by clarsimp (metis "**" dist_norm norm_minus_cancel) qed (auto simp: dist_norm intro: brouwer_ball[OF assms]) then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" by (auto simp: algebra_simps) then have "a = x" unfolding scaleR_left_distrib[symmetric] by auto then show False using x assms by auto qed corollary contractible_sphere: fixes a :: "'a::euclidean_space" shows "contractible(sphere a r) \ r \ 0" proof (cases "0 < r") case True then show ?thesis unfolding contractible_def nullhomotopic_from_sphere_extension using no_retraction_cball [OF True, of a] by (auto simp: retract_of_def retraction_def) next case False then show ?thesis unfolding contractible_def nullhomotopic_from_sphere_extension using less_eq_real_def by auto qed corollary connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by auto next case equal then show ?thesis by auto next case greater show ?thesis proof assume L: ?lhs have "False" if 1: "DIM('a) = 1" proof - obtain x y where xy: "sphere a r = {x,y}" "x \ y" using sphere_1D_doubleton [OF 1 greater] by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist) then have "finite (sphere a r)" by auto with L \r > 0\ xy show "False" using connected_finite_iff_sing by auto qed with greater show ?rhs by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq) next assume ?rhs then show ?lhs using connected_sphere greater by auto qed qed corollary path_connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "path_connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connected_sphere_eq path_connected_imp_connected by blast next assume R: ?rhs then show ?lhs by (auto simp: contractible_imp_path_connected contractible_sphere path_connected_sphere) qed proposition frontier_subset_retraction: fixes S :: "'a::euclidean_space set" assumes "bounded S" and fros: "frontier S \ T" and contf: "continuous_on (closure S) f" and fim: "f ` S \ T" and fid: "\x. x \ T \ f x = x" shows "S \ T" proof (rule ccontr) assume "\ S \ T" then obtain a where "a \ S" "a \ T" by blast define g where "g \ \z. if z \ closure S then f z else z" have "continuous_on (closure S \ closure(-S)) g" unfolding g_def apply (rule continuous_on_cases) using fros fid frontier_closures by (auto simp: contf) moreover have "closure S \ closure(- S) = UNIV" using closure_Un by fastforce ultimately have contg: "continuous_on UNIV g" by metis obtain B where "0 < B" and B: "closure S \ ball a B" using \bounded S\ bounded_subset_ballD by blast have notga: "g x \ a" for x unfolding g_def using fros fim \a \ T\ apply (auto simp: frontier_def) using fid interior_subset apply fastforce by (simp add: \a \ S\ closure_def) define h where "h \ (\y. a + (B / norm(y - a)) *\<^sub>R (y - a)) \ g" have "\ (frontier (cball a B) retract_of (cball a B))" by (metis no_retraction_cball \0 < B\) then have "\k. \ retraction (cball a B) (frontier (cball a B)) k" by (simp add: retract_of_def) moreover have "retraction (cball a B) (frontier (cball a B)) h" unfolding retraction_def proof (intro conjI ballI) show "frontier (cball a B) \ cball a B" by force show "continuous_on (cball a B) h" unfolding h_def by (intro continuous_intros) (use contg continuous_on_subset notga in auto) show "h ` cball a B \ frontier (cball a B)" using \0 < B\ by (auto simp: h_def notga dist_norm) show "\x. x \ frontier (cball a B) \ h x = x" apply (auto simp: h_def algebra_simps) apply (simp add: vector_add_divide_simps notga) by (metis (no_types, hide_lams) B add.commute dist_commute dist_norm g_def mem_ball not_less_iff_gr_or_eq subset_eq) qed ultimately show False by simp qed 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" and relS: "rel_frontier S \ T" and affS: "T \ affine hull S" obtains r where "homotopic_with_canon (\x. True) (T - {a}) (T - {a}) id r" "retraction (T - {a}) (rel_frontier S) r" proof - have "\d. 0 < d \ (a + d *\<^sub>R l) \ rel_frontier S \ (\e. 0 \ e \ e < d \ (a + e *\<^sub>R l) \ rel_interior S)" if "(a + l) \ affine hull S" "l \ 0" for l apply (rule ray_to_rel_frontier [OF \bounded S\ arelS]) apply (rule that)+ by metis then obtain dd where dd1: "\l. \(a + l) \ affine hull S; l \ 0\ \ 0 < dd l \ (a + dd l *\<^sub>R l) \ rel_frontier S" and dd2: "\l e. \(a + l) \ affine hull S; e < dd l; 0 \ e; l \ 0\ \ (a + e *\<^sub>R l) \ rel_interior S" by metis+ have aaffS: "a \ affine hull S" by (meson arelS subsetD hull_inc rel_interior_subset) have "((\z. z - a) ` (affine hull S - {a})) = ((\z. z - a) ` (affine hull S)) - {0}" by auto moreover have "continuous_on (((\z. z - a) ` (affine hull S)) - {0}) (\x. dd x *\<^sub>R x)" proof (rule continuous_on_compact_surface_projection) show "compact (rel_frontier ((\z. z - a) ` S))" by (simp add: \bounded S\ bounded_translation_minus compact_rel_frontier_bounded) have releq: "rel_frontier ((\z. z - a) ` S) = (\z. z - a) ` rel_frontier S" using rel_frontier_translation [of "-a"] add.commute by simp also have "\ \ (\z. z - a) ` (affine hull S) - {0}" using rel_frontier_affine_hull arelS rel_frontier_def by fastforce finally show "rel_frontier ((\z. z - a) ` S) \ (\z. z - a) ` (affine hull S) - {0}" . show "cone ((\z. z - a) ` (affine hull S))" by (rule subspace_imp_cone) (use aaffS in \simp add: subspace_affine image_comp o_def affine_translation_aux [of a]\) show "(0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)) \ (dd x = k)" if x: "x \ (\z. z - a) ` (affine hull S) - {0}" for k x proof show "dd x = k \ 0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)" using dd1 [of x] that image_iff by (fastforce simp add: releq) next assume k: "0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)" have False if "dd x < k" proof - have "k \ 0" "a + k *\<^sub>R x \ closure S" using k closure_translation [of "-a"] by (auto simp: rel_frontier_def cong: image_cong_simp) then have segsub: "open_segment a (a + k *\<^sub>R x) \ rel_interior S" by (metis rel_interior_closure_convex_segment [OF \convex S\ arelS]) have "x \ 0" and xaffS: "a + x \ affine hull S" using x by auto then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \ rel_frontier S" using dd1 by auto moreover have "a + dd x *\<^sub>R x \ open_segment a (a + k *\<^sub>R x)" using k \x \ 0\ \0 < dd x\ apply (simp add: in_segment) apply (rule_tac x = "dd x / k" in exI) apply (simp add: field_simps that) apply (simp add: vector_add_divide_simps algebra_simps) done ultimately show ?thesis using segsub by (auto simp: rel_frontier_def) qed moreover have False if "k < dd x" using x k that rel_frontier_def by (fastforce simp: algebra_simps releq dest!: dd2) ultimately show "dd x = k" by fastforce qed qed ultimately have *: "continuous_on ((\z. z - a) ` (affine hull S - {a})) (\x. dd x *\<^sub>R x)" by auto have "continuous_on (affine hull S - {a}) ((\x. a + dd x *\<^sub>R x) \ (\z. z - a))" by (intro * continuous_intros continuous_on_compose) with affS have contdd: "continuous_on (T - {a}) ((\x. a + dd x *\<^sub>R x) \ (\z. z - a))" by (blast intro: continuous_on_subset) show ?thesis proof show "homotopic_with_canon (\x. True) (T - {a}) (T - {a}) id (\x. a + dd (x - a) *\<^sub>R (x - a))" proof (rule homotopic_with_linear) show "continuous_on (T - {a}) id" by (intro continuous_intros continuous_on_compose) show "continuous_on (T - {a}) (\x. a + dd (x - a) *\<^sub>R (x - a))" using contdd by (simp add: o_def) show "closed_segment (id x) (a + dd (x - a) *\<^sub>R (x - a)) \ T - {a}" if "x \ T - {a}" for x proof (clarsimp simp: in_segment, intro conjI) fix u::real assume u: "0 \ u" "u \ 1" have "a + dd (x - a) *\<^sub>R (x - a) \ T" by (metis DiffD1 DiffD2 add.commute add.right_neutral affS dd1 diff_add_cancel relS singletonI subsetCE that) then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \ T" using convexD [OF \convex T\] that u by simp have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \ (1 - u + u * d) *\<^sub>R (x - a) = 0" for d by (auto simp: algebra_simps) have "x \ T" "x \ a" using that by auto then have axa: "a + (x - a) \ affine hull S" by (metis (no_types) add.commute affS diff_add_cancel rev_subsetD) then have "\ dd (x - a) \ 0 \ a + dd (x - a) *\<^sub>R (x - a) \ rel_frontier S" using \x \ a\ dd1 by fastforce with \x \ a\ show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \ a" apply (auto simp: iff) using less_eq_real_def mult_le_0_iff not_less u by fastforce qed qed show "retraction (T - {a}) (rel_frontier S) (\x. a + dd (x - a) *\<^sub>R (x - a))" proof (simp add: retraction_def, intro conjI ballI) show "rel_frontier S \ T - {a}" using arelS relS rel_frontier_def by fastforce show "continuous_on (T - {a}) (\x. a + dd (x - a) *\<^sub>R (x - a))" using contdd by (simp add: o_def) show "(\x. a + dd (x - a) *\<^sub>R (x - a)) ` (T - {a}) \ rel_frontier S" apply (auto simp: rel_frontier_def) apply (metis Diff_subset add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def subset_iff) by (metis DiffE add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def rev_subsetD) show "a + dd (x - a) *\<^sub>R (x - a) = x" if x: "x \ rel_frontier S" for x proof - have "x \ a" using that arelS by (auto simp: rel_frontier_def) have False if "dd (x - a) < 1" proof - have "x \ closure S" using x by (auto simp: rel_frontier_def) then have segsub: "open_segment a x \ rel_interior S" by (metis rel_interior_closure_convex_segment [OF \convex S\ arelS]) have xaffS: "x \ affine hull S" using affS relS x by auto then have "0 < dd (x - a)" and inS: "a + dd (x - a) *\<^sub>R (x - a) \ rel_frontier S" using dd1 by (auto simp: \x \ a\) moreover have "a + dd (x - a) *\<^sub>R (x - a) \ open_segment a x" using \x \ a\ \0 < dd (x - a)\ apply (simp add: in_segment) apply (rule_tac x = "dd (x - a)" in exI) apply (simp add: algebra_simps that) done ultimately show ?thesis using segsub by (auto simp: rel_frontier_def) qed moreover have False if "1 < dd (x - a)" using x that dd2 [of "x - a" 1] \x \ a\ closure_affine_hull by (auto simp: rel_frontier_def) ultimately have "dd (x - a) = 1" \ \similar to another proof above\ by fastforce with that show ?thesis by (simp add: rel_frontier_def) qed qed qed qed corollary rel_frontier_retract_of_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" "a \ rel_interior S" shows "rel_frontier S retract_of (affine hull S - {a})" apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S "affine hull S" a]) apply (auto simp: affine_imp_convex rel_frontier_affine_hull retract_of_def assms) done corollary rel_boundary_retract_of_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" "a \ rel_interior S" shows "(S - rel_interior S) retract_of (affine hull S - {a})" by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def rel_frontier_retract_of_punctured_affine_hull) lemma homotopy_eqv_rel_frontier_punctured_convex: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ rel_interior S" "convex T" "rel_frontier S \ T" "T \ affine hull S" shows "(rel_frontier S) homotopy_eqv (T - {a})" apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S T]) using assms apply auto using deformation_retract_imp_homotopy_eqv homotopy_equivalent_space_sym by blast lemma homotopy_eqv_rel_frontier_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ rel_interior S" shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})" apply (rule homotopy_eqv_rel_frontier_punctured_convex) using assms rel_frontier_affine_hull by force+ lemma path_connected_sphere_gen: assumes "convex S" "bounded S" "aff_dim S \ 1" shows "path_connected(rel_frontier S)" proof (cases "rel_interior S = {}") case True then show ?thesis by (simp add: \convex S\ convex_imp_path_connected rel_frontier_def) next case False then show ?thesis by (metis aff_dim_affine_hull affine_affine_hull affine_imp_convex all_not_in_conv assms path_connected_punctured_convex rel_frontier_retract_of_punctured_affine_hull retract_of_path_connected) qed lemma connected_sphere_gen: assumes "convex S" "bounded S" "aff_dim S \ 1" shows "connected(rel_frontier S)" by (simp add: assms path_connected_imp_connected path_connected_sphere_gen) 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))" by (rule continuous_intros | force)+ lemma Borsuk_map_into_sphere: "(\x. inverse(norm (x - a)) *\<^sub>R (x - a)) ` s \ sphere 0 1 \ (a \ s)" by auto (metis eq_iff_diff_eq_0 left_inverse norm_eq_zero) lemma Borsuk_maps_homotopic_in_path_component: assumes "path_component (- s) a b" shows "homotopic_with_canon (\x. True) s (sphere 0 1) (\x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\x. inverse(norm(x - b)) *\<^sub>R (x - b))" proof - obtain g where "path g" "path_image g \ -s" "pathstart g = a" "pathfinish g = b" using assms by (auto simp: path_component_def) then show ?thesis apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def) apply (rule_tac x = "\z. inverse(norm(snd z - (g \ fst)z)) *\<^sub>R (snd z - (g \ fst)z)" in exI) apply (intro conjI continuous_intros) apply (rule continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+ done qed lemma non_extensible_Borsuk_map: fixes a :: "'a :: euclidean_space" assumes "compact s" and cin: "c \ components(- s)" and boc: "bounded c" and "a \ c" shows "\ (\g. continuous_on (s \ c) g \ g ` (s \ c) \ sphere 0 1 \ (\x \ s. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))" proof - have "closed s" using assms by (simp add: compact_imp_closed) have "c \ -s" using assms by (simp add: in_components_subset) with \a \ c\ have "a \ s" by blast then have ceq: "c = connected_component_set (- s) a" by (metis \a \ c\ cin components_iff connected_component_eq) then have "bounded (s \ connected_component_set (- s) a)" using \compact s\ boc compact_imp_bounded by auto with bounded_subset_ballD obtain r where "0 < r" and r: "(s \ connected_component_set (- s) a) \ ball a r" by blast { fix g assume "continuous_on (s \ c) g" "g ` (s \ c) \ sphere 0 1" and [simp]: "\x. x \ s \ g x = (x - a) /\<^sub>R norm (x - a)" then have [simp]: "\x. x \ s \ c \ norm (g x) = 1" by force have cb_eq: "cball a r = (s \ connected_component_set (- s) a) \ (cball a r - connected_component_set (- s) a)" using ball_subset_cball [of a r] r by auto have cont1: "continuous_on (s \ connected_component_set (- s) a) (\x. a + r *\<^sub>R g x)" apply (rule continuous_intros)+ using \continuous_on (s \ c) g\ ceq by blast have cont2: "continuous_on (cball a r - connected_component_set (- s) a) (\x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" by (rule continuous_intros | force simp: \a \ s\)+ have 1: "continuous_on (cball a r) (\x. if connected_component (- s) a x then a + r *\<^sub>R g x else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" apply (subst cb_eq) apply (rule continuous_on_cases [OF _ _ cont1 cont2]) using ceq cin apply (auto intro: closed_Un_complement_component simp: \closed s\ open_Compl open_connected_component) done have 2: "(\x. a + r *\<^sub>R g x) ` (cball a r \ connected_component_set (- s) a) \ sphere a r " using \0 < r\ by (force simp: dist_norm ceq) have "retraction (cball a r) (sphere a r) (\x. if x \ connected_component_set (- s) a then a + r *\<^sub>R g x else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" using \0 < r\ apply (simp add: retraction_def dist_norm 1 2, safe) apply (force simp: dist_norm abs_if mult_less_0_iff divide_simps \a \ s\) using r by (auto simp: dist_norm norm_minus_commute) then have False using no_retraction_cball [OF \0 < r\, of a, unfolded retract_of_def, simplified, rule_format, of "\x. if x \ connected_component_set (- s) a then a + r *\<^sub>R g x else a + r *\<^sub>R inverse(norm(x - a)) *\<^sub>R (x - a)"] by blast } then show ?thesis by blast qed subsubsection \Proving surjectivity via Brouwer fixpoint theorem\ lemma brouwer_surjective: fixes f :: "'n::euclidean_space \ 'n" assumes "compact T" and "convex T" and "T \ {}" and "continuous_on T f" and "\x y. \x\S; y\T\ \ x + (y - f y) \ T" and "x \ S" shows "\y\T. f y = x" proof - have *: "\x y. f y = x \ x + (y - f y) = y" by (auto simp add: algebra_simps) show ?thesis unfolding * apply (rule brouwer[OF assms(1-3), of "\y. x + (y - f y)"]) apply (intro continuous_intros) using assms apply auto done qed lemma brouwer_surjective_cball: fixes f :: "'n::euclidean_space \ 'n" assumes "continuous_on (cball a e) f" and "e > 0" and "x \ S" and "\x y. \x\S; y\cball a e\ \ x + (y - f y) \ cball a e" shows "\y\cball a e. f y = x" apply (rule brouwer_surjective) apply (rule compact_cball convex_cball)+ unfolding cball_eq_empty using assms apply auto done subsubsection \Inverse function theorem\ text \See Sussmann: "Multidifferential calculus", Theorem 2.1.1\ lemma sussmann_open_mapping: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "open S" and contf: "continuous_on S f" and "x \ S" and derf: "(f has_derivative f') (at x)" and "bounded_linear g'" "f' \ g' = id" and "T \ S" and x: "x \ interior T" shows "f x \ interior (f ` T)" proof - interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto interpret g': bounded_linear g' using assms by auto obtain B where B: "0 < B" "\x. norm (g' x) \ norm x * B" using bounded_linear.pos_bounded[OF assms(5)] by blast hence *: "1 / (2 * B) > 0" by auto obtain e0 where e0: "0 < e0" "\y. norm (y - x) < e0 \ norm (f y - f x - f' (y - x)) \ 1 / (2 * B) * norm (y - x)" using derf unfolding has_derivative_at_alt using * by blast obtain e1 where e1: "0 < e1" "cball x e1 \ T" using mem_interior_cball x by blast have *: "0 < e0 / B" "0 < e1 / B" using e0 e1 B by auto obtain e where e: "0 < e" "e < e0 / B" "e < e1 / B" using field_lbound_gt_zero[OF *] by blast have lem: "\y\cball (f x) e. f (x + g' (y - f x)) = z" if "z\cball (f x) (e / 2)" for z proof (rule brouwer_surjective_cball) have z: "z \ S" if as: "y \cball (f x) e" "z = x + (g' y - g' (f x))" for y z proof- have "dist x z = norm (g' (f x) - g' y)" unfolding as(2) and dist_norm by auto also have "\ \ norm (f x - y) * B" by (metis B(2) g'.diff) also have "\ \ e * B" - by (metis B(1) dist_norm mem_cball real_mult_le_cancel_iff1 that(1)) + by (metis B(1) dist_norm mem_cball mult_le_cancel_iff1 that(1)) also have "\ \ e1" using B(1) e(3) pos_less_divide_eq by fastforce finally have "z \ cball x e1" by force then show "z \ S" using e1 assms(7) by auto qed show "continuous_on (cball (f x) e) (\y. f (x + g' (y - f x)))" unfolding g'.diff proof (rule continuous_on_compose2 [OF _ _ order_refl, of _ _ f]) show "continuous_on ((\y. x + (g' y - g' (f x))) ` cball (f x) e) f" by (rule continuous_on_subset[OF contf]) (use z in blast) show "continuous_on (cball (f x) e) (\y. x + (g' y - g' (f x)))" by (intro continuous_intros linear_continuous_on[OF \bounded_linear g'\]) qed next fix y z assume y: "y \ cball (f x) (e / 2)" and z: "z \ cball (f x) e" have "norm (g' (z - f x)) \ norm (z - f x) * B" using B by auto also have "\ \ e * B" - by (metis B(1) z dist_norm mem_cball norm_minus_commute real_mult_le_cancel_iff1) + by (metis B(1) z dist_norm mem_cball norm_minus_commute mult_le_cancel_iff1) also have "\ < e0" using B(1) e(2) pos_less_divide_eq by blast finally have *: "norm (x + g' (z - f x) - x) < e0" by auto have **: "f x + f' (x + g' (z - f x) - x) = z" using assms(6)[unfolded o_def id_def,THEN cong] by auto have "norm (f x - (y + (z - f (x + g' (z - f x))))) \ norm (f (x + g' (z - f x)) - z) + norm (f x - y)" using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by (auto simp add: algebra_simps) also have "\ \ 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0(2)[rule_format, OF *] by (simp only: algebra_simps **) auto also have "\ \ 1 / (B * 2) * norm (g' (z - f x)) + e/2" using y by (auto simp: dist_norm) also have "\ \ 1 / (B * 2) * B * norm (z - f x) + e/2" using * B by (auto simp add: field_simps) also have "\ \ 1 / 2 * norm (z - f x) + e/2" by auto also have "\ \ e/2 + e/2" using B(1) \norm (z - f x) * B \ e * B\ by auto finally show "y + (z - f (x + g' (z - f x))) \ cball (f x) e" by (auto simp: dist_norm) qed (use e that in auto) show ?thesis unfolding mem_interior proof (intro exI conjI subsetI) fix y assume "y \ ball (f x) (e / 2)" then have *: "y \ cball (f x) (e / 2)" by auto obtain z where z: "z \ cball (f x) e" "f (x + g' (z - f x)) = y" using lem * by blast then have "norm (g' (z - f x)) \ norm (z - f x) * B" using B by (auto simp add: field_simps) also have "\ \ e * B" - by (metis B(1) dist_norm mem_cball norm_minus_commute real_mult_le_cancel_iff1 z(1)) + by (metis B(1) dist_norm mem_cball norm_minus_commute mult_le_cancel_iff1 z(1)) also have "\ \ e1" using e B unfolding less_divide_eq by auto finally have "x + g'(z - f x) \ T" by (metis add_diff_cancel diff_diff_add dist_norm e1(2) mem_cball norm_minus_commute subset_eq) then show "y \ f ` T" using z by auto qed (use e in auto) qed text \Hence the following eccentric variant of the inverse function theorem. This has no continuity assumptions, but we do need the inverse function. We could put \f' \ g = I\ but this happens to fit with the minimal linear algebra theory I've set up so far.\ lemma has_derivative_inverse_strong: fixes f :: "'n::euclidean_space \ 'n" assumes "open S" and "x \ S" and contf: "continuous_on S f" and gf: "\x. x \ S \ g (f x) = x" and derf: "(f has_derivative f') (at x)" and id: "f' \ g' = id" shows "(g has_derivative g') (at (f x))" proof - have linf: "bounded_linear f'" using derf unfolding has_derivative_def by auto then have ling: "bounded_linear g'" unfolding linear_conv_bounded_linear[symmetric] using id right_inverse_linear by blast moreover have "g' \ f' = id" using id linf ling unfolding linear_conv_bounded_linear[symmetric] using linear_inverse_left by auto moreover have *: "\T. \T \ S; x \ interior T\ \ f x \ interior (f ` T)" apply (rule sussmann_open_mapping) apply (rule assms ling)+ apply auto done have "continuous (at (f x)) g" unfolding continuous_at Lim_at proof (rule, rule) fix e :: real assume "e > 0" then have "f x \ interior (f ` (ball x e \ S))" using *[rule_format,of "ball x e \ S"] \x \ S\ by (auto simp add: interior_open[OF open_ball] interior_open[OF assms(1)]) then obtain d where d: "0 < d" "ball (f x) d \ f ` (ball x e \ S)" unfolding mem_interior by blast show "\d>0. \y. 0 < dist y (f x) \ dist y (f x) < d \ dist (g y) (g (f x)) < e" proof (intro exI allI impI conjI) fix y assume "0 < dist y (f x) \ dist y (f x) < d" then have "g y \ g ` f ` (ball x e \ S)" by (metis d(2) dist_commute mem_ball rev_image_eqI subset_iff) then show "dist (g y) (g (f x)) < e" using gf[OF \x \ S\] by (simp add: assms(4) dist_commute image_iff) qed (use d in auto) qed moreover have "f x \ interior (f ` S)" apply (rule sussmann_open_mapping) apply (rule assms ling)+ using interior_open[OF assms(1)] and \x \ S\ apply auto done moreover have "f (g y) = y" if "y \ interior (f ` S)" for y by (metis gf imageE interiorE subsetD that) ultimately show ?thesis using assms by (metis has_derivative_inverse_basic_x open_interior) qed text \A rewrite based on the other domain.\ lemma has_derivative_inverse_strong_x: fixes f :: "'a::euclidean_space \ 'a" assumes "open S" and "g y \ S" and "continuous_on S f" and "\x. x \ S \ g (f x) = x" and "(f has_derivative f') (at (g y))" and "f' \ g' = id" and "f (g y) = y" shows "(g has_derivative g') (at y)" using has_derivative_inverse_strong[OF assms(1-6)] unfolding assms(7) by simp text \On a region.\ theorem has_derivative_inverse_on: fixes f :: "'n::euclidean_space \ 'n" assumes "open S" and derf: "\x. x \ S \ (f has_derivative f'(x)) (at x)" and "\x. x \ S \ g (f x) = x" and "f' x \ g' x = id" and "x \ S" shows "(g has_derivative g'(x)) (at (f x))" proof (rule has_derivative_inverse_strong[where g'="g' x" and f=f]) show "continuous_on S f" unfolding continuous_on_eq_continuous_at[OF \open S\] using derf has_derivative_continuous by blast qed (use assms in auto) end diff --git a/src/HOL/Analysis/Change_Of_Vars.thy b/src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy +++ b/src/HOL/Analysis/Change_Of_Vars.thy @@ -1,3478 +1,3478 @@ (* Title: HOL/Analysis/Change_Of_Vars.thy Authors: LC Paulson, based on material from HOL Light *) section\Change of Variables Theorems\ theory Change_Of_Vars imports Vitali_Covering_Theorem Determinants begin subsection \Measurable Shear and Stretch\ proposition fixes a :: "real^'n" assumes "m \ n" and ab_ne: "cbox a b \ {}" and an: "0 \ a$n" shows measurable_shear_interval: "(\x. \ i. if i = m then x$m + x$n else x$i) ` (cbox a b) \ lmeasurable" (is "?f ` _ \ _") and measure_shear_interval: "measure lebesgue ((\x. \ i. if i = m then x$m + x$n else x$i) ` cbox a b) = measure lebesgue (cbox a b)" (is "?Q") proof - have lin: "linear ?f" by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) show fab: "?f ` cbox a b \ lmeasurable" by (simp add: lin measurable_linear_image_interval) let ?c = "\ i. if i = m then b$m + b$n else b$i" let ?mn = "axis m 1 - axis n (1::real)" have eq1: "measure lebesgue (cbox a ?c) = measure lebesgue (?f ` cbox a b) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a$m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m})" proof (rule measure_Un3_negligible) show "cbox a ?c \ {x. ?mn \ x \ a$m} \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "negligible {x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ a $ m}) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible ((?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ a $ m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ b$m}) \ {x. ?mn \ x = b$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m})) \ {x. ?mn \ x = b$m}" using \m \ n\ ab_ne apply (auto simp: algebra_simps mem_box_cart inner_axis') apply (drule_tac x=m in spec)+ apply simp done ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) show "?f ` cbox a b \ cbox a ?c \ {x. ?mn \ x \ a $ m} \ cbox a ?c \ {x. ?mn \ x \ b$m} = cbox a ?c" (is "?lhs = _") proof show "?lhs \ cbox a ?c" by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans) show "cbox a ?c \ ?lhs" apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \m \ n\]) apply (auto simp: mem_box_cart split: if_split_asm) done qed qed (fact fab) let ?d = "\ i. if i = m then a $ m - b $ m else 0" have eq2: "measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a $ m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m}) = measure lebesgue (cbox a (\ i. if i = m then a $ m + b $ n else b $ i))" proof (rule measure_translate_add[of "cbox a ?c \ {x. ?mn \ x \ a$m}" "cbox a ?c \ {x. ?mn \ x \ b$m}" "(\ i. if i = m then a$m - b$m else 0)" "cbox a (\ i. if i = m then a$m + b$n else b$i)"]) show "(cbox a ?c \ {x. ?mn \ x \ a$m}) \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "\x. \x $ n + a $ m \ x $ m\ \ x \ (+) (\ i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \ x $ m}" using \m \ n\ by (rule_tac x="x - (\ i. if i = m then a$m - b$m else 0)" in image_eqI) (simp_all add: mem_box_cart) then have imeq: "(+) ?d ` {x. b $ m \ ?mn \ x} = {x. a $ m \ ?mn \ x}" using \m \ n\ by (auto simp: mem_box_cart inner_axis' algebra_simps) have "\x. \0 \ a $ n; x $ n + a $ m \ x $ m; \i. i \ m \ a $ i \ x $ i \ x $ i \ b $ i\ \ a $ m \ x $ m" using \m \ n\ by force then have "(+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {x. a $ m \ ?mn \ x}" using an ab_ne apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq) apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib) by (metis (full_types) add_mono mult_2_right) then show "cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i)" (is "?lhs = ?rhs") using an \m \ n\ apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force) apply (drule_tac x=n in spec)+ by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1)) have "negligible{x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x})) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}))" by (rule negligible_subset) qed have ac_ne: "cbox a ?c \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have ax_ne: "cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\ i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)" by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove if_distrib [of "\u. u - z" for z] prod.remove) show ?Q using eq1 eq2 eq3 by (simp add: algebra_simps) qed proposition fixes S :: "(real^'n) set" assumes "S \ lmeasurable" shows measurable_stretch: "((\x. \ k. m k * x$k) ` S) \ lmeasurable" (is "?f ` S \ _") and measure_stretch: "measure lebesgue ((\x. \ k. m k * x$k) ` S) = \prod m UNIV\ * measure lebesgue S" (is "?MEQ") proof - have "(?f ` S) \ lmeasurable \ ?MEQ" proof (cases "\k. m k \ 0") case True have m0: "0 < \prod m UNIV\" using True by simp have "(indicat_real (?f ` S) has_integral \prod m UNIV\ * measure lebesgue S) UNIV" proof (clarsimp simp add: has_integral_alt [where i=UNIV]) fix e :: "real" assume "e > 0" have "(indicat_real S has_integral (measure lebesgue S)) UNIV" using assms lmeasurable_iff_has_integral by blast then obtain B where "B>0" and B: "\a b. ball 0 B \ cbox a b \ \z. (indicat_real S has_integral z) (cbox a b) \ \z - measure lebesgue S\ < e / \prod m UNIV\" by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0 m0 \e > 0\) show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (indicat_real (?f ` S) has_integral z) (cbox a b) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" proof (intro exI conjI allI) let ?C = "Max (range (\k. \m k\)) * B" show "?C > 0" using True \B > 0\ by (simp add: Max_gr_iff) show "ball 0 ?C \ cbox u v \ (\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" for u v proof assume uv: "ball 0 ?C \ cbox u v" with \?C > 0\ have cbox_ne: "cbox u v \ {}" using centre_in_ball by blast let ?\ = "\k. u$k / m k" let ?\ = "\k. v$k / m k" have invm0: "\k. inverse (m k) \ 0" using True by auto have "ball 0 B \ (\x. \ k. x $ k / m k) ` ball 0 ?C" proof clarsimp fix x :: "real^'n" assume x: "norm x < B" have [simp]: "\Max (range (\k. \m k\))\ = Max (range (\k. \m k\))" by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI) have "norm (\ k. m k * x $ k) \ norm (Max (range (\k. \m k\)) *\<^sub>R x)" by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono) also have "\ < ?C" - using x by simp (metis \B > 0\ \?C > 0\ mult.commute real_mult_less_iff1 zero_less_mult_pos) + using x \0 < (MAX k. \m k\) * B\ \0 < B\ zero_less_mult_pos2 by fastforce finally have "norm (\ k. m k * x $ k) < ?C" . then show "x \ (\x. \ k. x $ k / m k) ` ball 0 ?C" using stretch_Galois [of "inverse \ m"] True by (auto simp: image_iff field_simps) qed then have Bsub: "ball 0 B \ cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))" using cbox_ne uv image_stretch_interval_cart [of "inverse \ m" u v, symmetric] by (force simp: field_simps) obtain z where zint: "(indicat_real S has_integral z) (cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" and zless: "\z - measure lebesgue S\ < e / \prod m UNIV\" using B [OF Bsub] by blast have ind: "indicat_real (?f ` S) = (\x. indicator S (\ k. x$k / m k))" using True stretch_Galois [of m] by (force simp: indicator_def) show "\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e" proof (simp add: ind, intro conjI exI) have "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) ((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" using True has_integral_stretch_cart [OF zint, of "inverse \ m"] by (simp add: field_simps prod_dividef) moreover have "((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))) = cbox u v" using True image_stretch_interval_cart [of "inverse \ m" u v, symmetric] image_stretch_interval_cart [of "\k. 1" u v, symmetric] \cbox u v \ {}\ by (simp add: field_simps image_comp o_def) ultimately show "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) (cbox u v)" by simp have "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ = \prod m UNIV\ * \z - measure lebesgue S\" by (metis (no_types, hide_lams) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib') also have "\ < e" using zless True by (simp add: field_simps) finally show "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ < e" . qed qed qed qed then show ?thesis by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable) next case False then obtain k where "m k = 0" and prm: "prod m UNIV = 0" by auto have nfS: "negligible (?f ` S)" by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \m k = 0\ in auto) then have "(?f ` S) \ lmeasurable" by (simp add: negligible_iff_measure) with nfS show ?thesis by (simp add: prm negligible_iff_measure0) qed then show "(?f ` S) \ lmeasurable" ?MEQ by metis+ qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "linear f" "S \ lmeasurable" shows measurable_linear_image: "(f ` S) \ lmeasurable" and measure_linear_image: "measure lebesgue (f ` S) = \det (matrix f)\ * measure lebesgue S" (is "?Q f S") proof - have "\S \ lmeasurable. (f ` S) \ lmeasurable \ ?Q f S" proof (rule induct_linear_elementary [OF \linear f\]; intro ballI) fix f g and S :: "(real,'n) vec set" assume "linear f" and "linear g" and f [rule_format]: "\S \ lmeasurable. f ` S \ lmeasurable \ ?Q f S" and g [rule_format]: "\S \ lmeasurable. g ` S \ lmeasurable \ ?Q g S" and S: "S \ lmeasurable" then have gS: "g ` S \ lmeasurable" by blast show "(f \ g) ` S \ lmeasurable \ ?Q (f \ g) S" using f [OF gS] g [OF S] matrix_compose [OF \linear g\ \linear f\] by (simp add: o_def image_comp abs_mult det_mul) next fix f :: "real^'n::_ \ real^'n::_" and i and S :: "(real^'n::_) set" assume "linear f" and 0: "\x. f x $ i = 0" and "S \ lmeasurable" then have "\ inj f" by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component) have detf: "det (matrix f) = 0" using \\ inj f\ det_nz_iff_inj[OF \linear f\] by blast show "f ` S \ lmeasurable \ ?Q f S" proof show "f ` S \ lmeasurable" using lmeasurable_iff_indicator_has_integral \linear f\ \\ inj f\ negligible_UNIV negligible_linear_singular_image by blast have "measure lebesgue (f ` S) = 0" by (meson \\ inj f\ \linear f\ negligible_imp_measure0 negligible_linear_singular_image) also have "\ = \det (matrix f)\ * measure lebesgue S" by (simp add: detf) finally show "?Q f S" . qed next fix c and S :: "(real^'n::_) set" assume "S \ lmeasurable" show "(\a. \ i. c i * a $ i) ` S \ lmeasurable \ ?Q (\a. \ i. c i * a $ i) S" proof show "(\a. \ i. c i * a $ i) ` S \ lmeasurable" by (simp add: \S \ lmeasurable\ measurable_stretch) show "?Q (\a. \ i. c i * a $ i) S" by (simp add: measure_stretch [OF \S \ lmeasurable\, of c] axis_def matrix_def det_diagonal) qed next fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i" have lin: "linear ?h" by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def) have meq: "measure lebesgue ((\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i) ` cbox a b) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have him: "?h ` (cbox a b) \ {}" by blast have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)" by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis swap_id_eq)+ show ?thesis using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\i. (b - a)$i", symmetric] by (simp add: eq content_cbox_cart False) qed have "(\ i j. if Fun.swap m n id i = j then 1 else 0) = (\ i j. if j = Fun.swap m n id i then 1 else (0::real))" by (auto intro!: Cart_lambda_cong) then have "matrix ?h = transpose(\ i j. mat 1 $ i $ Fun.swap m n id j)" by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def) then have 1: "\det (matrix ?h)\ = 1" by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult) show "?h ` S \ lmeasurable \ ?Q ?h S" proof show "?h ` S \ lmeasurable" "?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force+ qed next fix m n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. if i = m then v $ m + v $ n else v $ i" have lin: "linear ?h" by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff) consider "m < n" | " n < m" using \m \ n\ less_linear by blast then have 1: "det(matrix ?h) = 1" proof cases assume "m < n" have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j < i\ axis_def \m < n\ by auto with \m < n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) next assume "n < m" have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j > i\ axis_def \m > n\ by auto with \m > n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) qed have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have ne: "(+) (\ i. if i = n then - a $ n else 0) ` cbox a b \ {}" by auto let ?v = "\ i. if i = n then - a $ n else 0" have "?h ` cbox a b = (+) (\ i. if i = m \ i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)" using \m \ n\ unfolding image_comp o_def by (force simp: vec_eq_iff) then have "measure lebesgue (?h ` (cbox a b)) = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` (+) ?v ` cbox a b)" by (rule ssubst) (rule measure_translation) also have "\ = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))" by (metis (no_types, lifting) cbox_translation) also have "\ = measure lebesgue ((+) (\ i. if i = n then - a $ n else 0) ` cbox a b)" apply (subst measure_shear_interval) using \m \ n\ ne apply auto apply (simp add: cbox_translation) by (metis cbox_borel cbox_translation measure_completion sets_lborel) also have "\ = measure lebesgue (cbox a b)" by (rule measure_translation) finally show ?thesis . qed show "?h ` S \ lmeasurable \ ?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force qed with assms show "(f ` S) \ lmeasurable" "?Q f S" by metis+ qed lemma fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes f: "orthogonal_transformation f" and S: "S \ lmeasurable" shows measurable_orthogonal_image: "f ` S \ lmeasurable" and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S" proof - have "linear f" by (simp add: f orthogonal_transformation_linear) then show "f ` S \ lmeasurable" by (metis S measurable_linear_image) show "measure lebesgue (f ` S) = measure lebesgue S" by (simp add: measure_linear_image \linear f\ S f) qed proposition measure_semicontinuous_with_hausdist_explicit: assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0" obtains d where "d > 0" "\T. \T \ lmeasurable; \y. y \ T \ \x. x \ S \ dist x y < d\ \ measure lebesgue T < measure lebesgue S + e" proof (cases "S = {}") case True with that \e > 0\ show ?thesis by force next case False then have frS: "frontier S \ {}" using \bounded S\ frontier_eq_empty not_bounded_UNIV by blast have "S \ lmeasurable" by (simp add: \bounded S\ measurable_Jordan neg) have null: "(frontier S) \ null_sets lebesgue" by (metis neg negligible_iff_null_sets) have "frontier S \ lmeasurable" and mS0: "measure lebesgue (frontier S) = 0" using neg negligible_imp_measurable negligible_iff_measure by blast+ with \e > 0\ sets_lebesgue_outer_open obtain U where "open U" and U: "frontier S \ U" "U - frontier S \ lmeasurable" "emeasure lebesgue (U - frontier S) < e" by (metis fmeasurableD) with null have "U \ lmeasurable" by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel) have "measure lebesgue (U - frontier S) = measure lebesgue U" using mS0 by (simp add: \U \ lmeasurable\ fmeasurableD measure_Diff_null_set null) with U have mU: "measure lebesgue U < e" by (simp add: emeasure_eq_measure2 ennreal_less_iff) show ?thesis proof have "U \ UNIV" using \U \ lmeasurable\ by auto then have "- U \ {}" by blast with \open U\ \frontier S \ U\ show "setdist (frontier S) (- U) > 0" by (auto simp: \bounded S\ open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS) fix T assume "T \ lmeasurable" and T: "\t. t \ T \ \y. y \ S \ dist y t < setdist (frontier S) (- U)" then have "measure lebesgue T - measure lebesgue S \ measure lebesgue (T - S)" by (simp add: \S \ lmeasurable\ measure_diff_le_measure_setdiff) also have "\ \ measure lebesgue U" proof - have "T - S \ U" proof clarify fix x assume "x \ T" and "x \ S" then obtain y where "y \ S" and y: "dist y x < setdist (frontier S) (- U)" using T by blast have "closed_segment x y \ frontier S \ {}" using connected_Int_frontier \x \ S\ \y \ S\ by blast then obtain z where z: "z \ closed_segment x y" "z \ frontier S" by auto with y have "dist z x < setdist(frontier S) (- U)" by (auto simp: dist_commute dest!: dist_in_closed_segment) with z have False if "x \ -U" using setdist_le_dist [OF \z \ frontier S\ that] by auto then show "x \ U" by blast qed then show ?thesis by (simp add: \S \ lmeasurable\ \T \ lmeasurable\ \U \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Diff) qed finally have "measure lebesgue T - measure lebesgue S \ measure lebesgue U" . with mU show "measure lebesgue T < measure lebesgue S + e" by linarith qed qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" and bounded: "\x. x \ S \ \det (matrix (f' x))\ \ B" shows measurable_bounded_differentiable_image: "f ` S \ lmeasurable" and measure_bounded_differentiable_image: "measure lebesgue (f ` S) \ B * measure lebesgue S" (is "?M") proof - have "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ B * measure lebesgue S" proof (cases "B < 0") case True then have "S = {}" by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI) then show ?thesis by auto next case False then have "B \ 0" by arith let ?\ = "measure lebesgue" have f_diff: "f differentiable_on S" using deriv by (auto simp: differentiable_on_def differentiable_def) have eps: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * ?\ S" (is "?ME") if "e > 0" for e proof - have eps_d: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * (?\ S + d)" (is "?MD") if "d > 0" for d proof - obtain T where T: "open T" "S \ T" and TS: "(T-S) \ lmeasurable" and "emeasure lebesgue (T-S) < ennreal d" using S \d > 0\ sets_lebesgue_outer_open by blast then have "?\ (T-S) < d" by (metis emeasure_eq_measure2 ennreal_leI not_less) with S T TS have "T \ lmeasurable" and Tless: "?\ T < ?\ S + d" by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D) have "\r. 0 < r \ r < d \ ball x r \ T \ f ` (S \ ball x r) \ lmeasurable \ ?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" if "x \ S" "d > 0" for x d proof - have lin: "linear (f' x)" and lim0: "((\y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \ 0) (at x within S)" using deriv \x \ S\ by (auto simp: has_derivative_within bounded_linear.linear field_simps) have bo: "bounded (f' x ` ball 0 1)" by (simp add: bounded_linear_image linear_linear lin) have neg: "negligible (frontier (f' x ` ball 0 1))" using deriv has_derivative_linear \x \ S\ by (auto intro!: negligible_convex_frontier [OF convex_linear_image]) let ?unit_vol = "content (ball (0 :: real ^ 'n :: {finite, wellorder}) 1)" have 0: "0 < e * ?unit_vol" using \e > 0\ by (simp add: content_ball_pos) obtain k where "k > 0" and k: "\U. \U \ lmeasurable; \y. y \ U \ \z. z \ f' x ` ball 0 1 \ dist z y < k\ \ ?\ U < ?\ (f' x ` ball 0 1) + e * ?unit_vol" using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast obtain l where "l > 0" and l: "ball x l \ T" using \x \ S\ \open T\ \S \ T\ openE by blast obtain \ where "0 < \" and \: "\y. \y \ S; y \ x; dist y x < \\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x) < k" using lim0 \k > 0\ by (simp add: Lim_within) (auto simp add: field_simps) define r where "r \ min (min l (\/2)) (min 1 (d/2))" show ?thesis proof (intro exI conjI) show "r > 0" "r < d" using \l > 0\ \\ > 0\ \d > 0\ by (auto simp: r_def) have "r \ l" by (auto simp: r_def) with l show "ball x r \ T" by auto have ex_lessK: "\x' \ ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k" if "y \ S" and "dist x y < r" for y proof (cases "y = x") case True with lin linear_0 \k > 0\ that show ?thesis by (rule_tac x=0 in bexI) (auto simp: linear_0) next case False then show ?thesis proof (rule_tac x="(y - x) /\<^sub>R r" in bexI) have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r" by (simp add: lin linear_scale) then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)" by (simp add: dist_norm) also have "\ = norm (f' x (y - x) - (f y - f x)) / r" using \r > 0\ by (simp add: divide_simps scale_right_diff_distrib [symmetric]) also have "\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x)" using that \r > 0\ False by (simp add: field_split_simps dist_norm norm_minus_commute mult_right_mono) also have "\ < k" using that \0 < \\ by (simp add: dist_commute r_def \ [OF \y \ S\ False]) finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" . show "(y - x) /\<^sub>R r \ ball 0 1" using that \r > 0\ by (simp add: dist_norm divide_simps norm_minus_commute) qed qed let ?rfs = "(\x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \ ball x r)" have rfs_mble: "?rfs \ lmeasurable" proof (rule bounded_set_imp_lmeasurable) have "f differentiable_on S \ ball x r" using f_diff by (auto simp: fmeasurableD differentiable_on_subset) with S show "?rfs \ sets lebesgue" by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue) let ?B = "(\(x, y). x + y) ` (f' x ` ball 0 1 \ ball 0 k)" have "bounded ?B" by (simp add: bounded_plus [OF bo]) moreover have "?rfs \ ?B" apply (auto simp: dist_norm image_iff dest!: ex_lessK) by (metis (no_types, hide_lams) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball) ultimately show "bounded (?rfs)" by (rule bounded_subset) qed then have "(\x. r *\<^sub>R x) ` ?rfs \ lmeasurable" by (simp add: measurable_linear_image) with \r > 0\ have "(+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) then have "(+) (f x) ` (+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" using measurable_translation by blast then show fsb: "f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) have "?\ (f ` (S \ ball x r)) = ?\ (?rfs) * r ^ CARD('n)" using \r > 0\ fsb by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp) also have "\ \ (\det (matrix (f' x))\ * ?unit_vol + e * ?unit_vol) * r ^ CARD('n)" proof - have "?\ (?rfs) < ?\ (f' x ` ball 0 1) + e * ?unit_vol" using rfs_mble by (force intro: k dest!: ex_lessK) then have "?\ (?rfs) < \det (matrix (f' x))\ * ?unit_vol + e * ?unit_vol" by (simp add: lin measure_linear_image [of "f' x"]) with \r > 0\ show ?thesis by auto qed also have "\ \ (B + e) * ?\ (ball x r)" using bounded [OF \x \ S\] \r > 0\ by (simp add: algebra_simps content_ball_conv_unit_ball[of r] content_ball_pos) finally show "?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" . qed qed then obtain r where r0d: "\x d. \x \ S; d > 0\ \ 0 < r x d \ r x d < d" and rT: "\x d. \x \ S; d > 0\ \ ball x (r x d) \ T" and r: "\x d. \x \ S; d > 0\ \ (f ` (S \ ball x (r x d))) \ lmeasurable \ ?\ (f ` (S \ ball x (r x d))) \ (B + e) * ?\ (ball x (r x d))" by metis obtain C where "countable C" and Csub: "C \ {(x,r x t) |x t. x \ S \ 0 < t}" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible(S - (\i \ C. ball (fst i) (snd i)))" apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \ S \ 0 < t}" fst snd]) apply auto by (metis dist_eq_0_iff r0d) let ?UB = "(\(x,s) \ C. ball x s)" have eq: "f ` (S \ ?UB) = (\(x,s) \ C. f ` (S \ ball x s))" by auto have mle: "?\ (\(x,s) \ K. f ` (S \ ball x s)) \ (B + e) * (?\ S + d)" (is "?l \ ?r") if "K \ C" and "finite K" for K proof - have gt0: "b > 0" if "(a, b) \ K" for a b using Csub that \K \ C\ r0d by auto have inj: "inj_on (\(x, y). ball x y) K" by (force simp: inj_on_def ball_eq_ball_iff dest: gt0) have disjnt: "disjoint ((\(x, y). ball x y) ` K)" using pwC that apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff) by (metis subsetD fst_conv snd_conv) have "?l \ (\i\K. ?\ (case i of (x, s) \ f ` (S \ ball x s)))" proof (rule measure_UNION_le [OF \finite K\], clarify) fix x r assume "(x,r) \ K" then have "x \ S" using Csub \K \ C\ by auto show "f ` (S \ ball x r) \ sets lebesgue" by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue) qed also have "\ \ (\(x,s) \ K. (B + e) * ?\ (ball x s))" apply (rule sum_mono) using Csub r \K \ C\ by auto also have "\ = (B + e) * (\(x,s) \ K. ?\ (ball x s))" by (simp add: prod.case_distrib sum_distrib_left) also have "\ = (B + e) * sum ?\ ((\(x, y). ball x y) ` K)" using \B \ 0\ \e > 0\ by (simp add: inj sum.reindex prod.case_distrib) also have "\ = (B + e) * ?\ (\(x,s) \ K. ball x s)" using \B \ 0\ \e > 0\ that by (subst measure_Union') (auto simp: disjnt measure_Union') also have "\ \ (B + e) * ?\ T" using \B \ 0\ \e > 0\ that apply simp apply (rule measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\]) using Csub rT by force+ also have "\ \ (B + e) * (?\ S + d)" using \B \ 0\ \e > 0\ Tless by simp finally show ?thesis . qed have fSUB_mble: "(f ` (S \ ?UB)) \ lmeasurable" unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: fmeasurable_UN_bound [OF \countable C\ _ mle]) have fSUB_meas: "?\ (f ` (S \ ?UB)) \ (B + e) * (?\ S + d)" (is "?MUB") unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: measure_UN_bound [OF \countable C\ _ mle]) have neg: "negligible ((f ` (S \ ?UB) - f ` S) \ (f ` S - f ` (S \ ?UB)))" proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]]) show "f differentiable_on S - (\i\C. ball (fst i) (snd i))" by (meson DiffE differentiable_on_subset subsetI f_diff) qed force show "f ` S \ lmeasurable" by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg]) show ?MD using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp qed show "f ` S \ lmeasurable" using eps_d [of 1] by simp show ?ME proof (rule field_le_epsilon) fix \ :: real assume "0 < \" then show "?\ (f ` S) \ (B + e) * ?\ S + \" using eps_d [of "\ / (B+e)"] \e > 0\ \B \ 0\ by (auto simp: divide_simps mult_ac) qed qed show ?thesis proof (cases "?\ S = 0") case True with eps have "?\ (f ` S) = 0" by (metis mult_zero_right not_le zero_less_measure_iff) then show ?thesis using eps [of 1] by (simp add: True) next case False have "?\ (f ` S) \ B * ?\ S" proof (rule field_le_epsilon) fix e :: real assume "e > 0" then show "?\ (f ` S) \ B * ?\ S + e" using eps [of "e / ?\ S"] False by (auto simp: algebra_simps zero_less_measure_iff) qed with eps [of 1] show ?thesis by auto qed qed then show "f ` S \ lmeasurable" ?M by blast+ qed lemma m_diff_image_weak: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" proof - let ?\ = "measure lebesgue" have aint_S: "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int unfolding absolutely_integrable_on_def by auto define m where "m \ integral S (\x. \det (matrix (f' x))\)" have *: "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" if "e > 0" for e proof - define T where "T \ \n. {x \ S. n * e \ \det (matrix (f' x))\ \ \det (matrix (f' x))\ < (Suc n) * e}" have meas_t: "T n \ lmeasurable" for n proof - have *: "(\x. \det (matrix (f' x))\) \ borel_measurable (lebesgue_on S)" using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def) have [intro]: "x \ sets (lebesgue_on S) \ x \ sets lebesgue" for x using S sets_restrict_space_subset by blast have "{x \ S. real n * e \ \det (matrix (f' x))\} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space) then have 1: "{x \ S. real n * e \ \det (matrix (f' x))\} \ lmeasurable" using S by (simp add: fmeasurableI2) have "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space) then have 2: "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ lmeasurable" using S by (simp add: fmeasurableI2) show ?thesis using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong) qed have aint_T: "\k. (\x. \det (matrix (f' x))\) absolutely_integrable_on T k" using set_integrable_subset [OF aint_S] meas_t T_def by blast have Seq: "S = (\n. T n)" apply (auto simp: T_def) apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI) using that apply auto using of_int_floor_le pos_le_divide_eq apply blast by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt) have meas_ft: "f ` T n \ lmeasurable" for n proof (rule measurable_bounded_differentiable_image) show "T n \ lmeasurable" by (simp add: meas_t) next fix x :: "(real,'n) vec" assume "x \ T n" show "(f has_derivative f' x) (at x within T n)" by (metis (no_types, lifting) \x \ T n\ deriv has_derivative_subset mem_Collect_eq subsetI T_def) show "\det (matrix (f' x))\ \ (Suc n) * e" using \x \ T n\ T_def by auto next show "(\x. \det (matrix (f' x))\) integrable_on T n" using aint_T absolutely_integrable_on_def by blast qed have disT: "disjoint (range T)" unfolding disjoint_def proof clarsimp show "T m \ T n = {}" if "T m \ T n" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) with \e > 0\ show ?case unfolding T_def proof (clarsimp simp add: Collect_conj_eq [symmetric]) fix x assume "e > 0" "m < n" "n * e \ \det (matrix (f' x))\" "\det (matrix (f' x))\ < (1 + real m) * e" then have "n < 1 + real m" - by (metis (no_types, hide_lams) less_le_trans mult.commute not_le real_mult_le_cancel_iff2) + by (metis (no_types, hide_lams) less_le_trans mult.commute not_le mult_le_cancel_iff2) then show "False" using less.hyps by linarith qed qed auto qed have injT: "inj_on T ({n. T n \ {}})" unfolding inj_on_def proof clarsimp show "m = n" if "T m = T n" "T n \ {}" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) have False if "T n \ T m" "x \ T n" for x using \e > 0\ \m < n\ that apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) - by (metis add.commute less_le_trans nat_less_real_le not_le real_mult_le_cancel_iff2) + by (metis add.commute less_le_trans nat_less_real_le not_le mult_le_cancel_iff2) then show ?case using less.prems by blast qed auto qed have sum_eq_Tim: "(\k\n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \ real" and n proof (subst sum.reindex_nontrivial) fix i j assume "i \ {..n}" "j \ {..n}" "i \ j" "T i = T j" with that injT [unfolded inj_on_def] show "f (T i) = 0" by simp metis qed (use atMost_atLeast0 in auto) let ?B = "m + e * ?\ S" have "(\k\n. ?\ (f ` T k)) \ ?B" for n proof - have "(\k\n. ?\ (f ` T k)) \ (\k\n. ((k+1) * e) * ?\(T k))" proof (rule sum_mono [OF measure_bounded_differentiable_image]) show "(f has_derivative f' x) (at x within T k)" if "x \ T k" for k x using that unfolding T_def by (blast intro: deriv has_derivative_subset) show "(\x. \det (matrix (f' x))\) integrable_on T k" for k using absolutely_integrable_on_def aint_T by blast show "\det (matrix (f' x))\ \ real (k + 1) * e" if "x \ T k" for k x using T_def that by auto qed (use meas_t in auto) also have "\ \ (\k\n. (k * e) * ?\(T k)) + (\k\n. e * ?\(T k))" by (simp add: algebra_simps sum.distrib) also have "\ \ ?B" proof (rule add_mono) have "(\k\n. real k * e * ?\ (T k)) = (\k\n. integral (T k) (\x. k * e))" by (simp add: lmeasure_integral [OF meas_t] flip: integral_mult_right integral_mult_left) also have "\ \ (\k\n. integral (T k) (\x. (abs (det (matrix (f' x))))))" proof (rule sum_mono) fix k assume "k \ {..n}" show "integral (T k) (\x. k * e) \ integral (T k) (\x. \det (matrix (f' x))\)" proof (rule integral_le [OF integrable_on_const [OF meas_t]]) show "(\x. \det (matrix (f' x))\) integrable_on T k" using absolutely_integrable_on_def aint_T by blast next fix x assume "x \ T k" show "k * e \ \det (matrix (f' x))\" using \x \ T k\ T_def by blast qed qed also have "\ = sum (\T. integral T (\x. \det (matrix (f' x))\)) (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = integral (\k\n. T k) (\x. \det (matrix (f' x))\)" proof (rule integral_unique [OF has_integral_Union, symmetric]) fix S assume "S \ T ` {..n}" then show "((\x. \det (matrix (f' x))\) has_integral integral S (\x. \det (matrix (f' x))\)) S" using absolutely_integrable_on_def aint_T by blast next show "pairwise (\S S'. negligible (S \ S')) (T ` {..n})" using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible) qed auto also have "\ \ m" unfolding m_def proof (rule integral_subset_le) have "(\x. \det (matrix (f' x))\) absolutely_integrable_on (\k\n. T k)" apply (rule set_integrable_subset [OF aint_S]) apply (intro measurable meas_t fmeasurableD) apply (force simp: Seq) done then show "(\x. \det (matrix (f' x))\) integrable_on (\k\n. T k)" using absolutely_integrable_on_def by blast qed (use Seq int in auto) finally show "(\k\n. real k * e * ?\ (T k)) \ m" . next have "(\k\n. ?\ (T k)) = sum ?\ (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = ?\ (\k\n. T k)" using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric]) also have "\ \ ?\ S" using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable) finally have "(\k\n. ?\ (T k)) \ ?\ S" . then show "(\k\n. e * ?\ (T k)) \ e * ?\ S" by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that) qed finally show "(\k\n. ?\ (f ` T k)) \ ?B" . qed moreover have "measure lebesgue (\k\n. f ` T k) \ (\k\n. ?\ (f ` T k))" for n by (simp add: fmeasurableD meas_ft measure_UNION_le) ultimately have B_ge_m: "?\ (\k\n. (f ` T k)) \ ?B" for n by (meson order_trans) have "(\n. f ` T n) \ lmeasurable" by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m]) moreover have "?\ (\n. f ` T n) \ m + e * ?\ S" by (rule measure_countable_Union_le [OF meas_ft B_ge_m]) ultimately show "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" by (auto simp: Seq image_Union) qed show ?thesis proof show "f ` S \ lmeasurable" using * linordered_field_no_ub by blast let ?x = "m - ?\ (f ` S)" have False if "?\ (f ` S) > integral S (\x. \det (matrix (f' x))\)" proof - have ml: "m < ?\ (f ` S)" using m_def that by blast then have "?\ S \ 0" using "*"(2) bgauge_existence_lemma by fastforce with ml have 0: "0 < - (m - ?\ (f ` S))/2 / ?\ S" using that zero_less_measure_iff by force then show ?thesis using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm) qed then show "?\ (f ` S) \ integral S (\x. \det (matrix (f' x))\)" by fastforce qed qed theorem fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows measurable_differentiable_image: "f ` S \ lmeasurable" and measure_differentiable_image: "measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" (is "?M") proof - let ?I = "\n::nat. cbox (vec (-n)) (vec n) \ S" let ?\ = "measure lebesgue" have "x \ cbox (vec (- real (nat \norm x\))) (vec (real (nat \norm x\)))" for x :: "real^'n::_" apply (auto simp: mem_box_cart) apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans) by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge) then have Seq: "S = (\n. ?I n)" by auto have fIn: "f ` ?I n \ lmeasurable" and mfIn: "?\ (f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" (is ?MN) for n proof - have In: "?I n \ lmeasurable" by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int) moreover have "\x. x \ ?I n \ (f has_derivative f' x) (at x within ?I n)" by (meson Int_iff deriv has_derivative_subset subsetI) moreover have int_In: "(\x. \det (matrix (f' x))\) integrable_on ?I n" proof - have "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int absolutely_integrable_integrable_bound by force then have "(\x. \det (matrix (f' x))\) absolutely_integrable_on ?I n" by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset) then show ?thesis using absolutely_integrable_on_def by blast qed ultimately have "f ` ?I n \ lmeasurable" "?\ (f ` ?I n) \ integral (?I n) (\x. \det (matrix (f' x))\)" using m_diff_image_weak by metis+ moreover have "integral (?I n) (\x. \det (matrix (f' x))\) \ integral S (\x. \det (matrix (f' x))\)" by (simp add: int_In int integral_subset_le) ultimately show "f ` ?I n \ lmeasurable" ?MN by auto qed have "?I k \ ?I n" if "k \ n" for k n by (rule Int_mono) (use that in \auto simp: subset_interval_imp_cart\) then have "(\k\n. f ` ?I k) = f ` ?I n" for n by (fastforce simp add:) with mfIn have "?\ (\k\n. f ` ?I k) \ integral S (\x. \det (matrix (f' x))\)" for n by simp then have "(\n. f ` ?I n) \ lmeasurable" "?\ (\n. f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+ then show "f ` S \ lmeasurable" ?M by (metis Seq image_UN)+ qed lemma borel_measurable_simple_function_limit_increasing: fixes f :: "'a::euclidean_space \ real" shows "(f \ borel_measurable lebesgue \ (\x. 0 \ f x)) \ (\g. (\n x. 0 \ g n x \ g n x \ f x) \ (\n x. g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \ (\x. (\n. g n x) \ f x))" (is "?lhs = ?rhs") proof assume f: ?lhs have leb_f: "{x. a \ f x \ f x < b} \ sets lebesgue" for a b proof - have "{x. a \ f x \ f x < b} = {x. f x < b} - {x. f x < a}" by auto also have "\ \ sets lebesgue" using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto finally show ?thesis . qed have "g n x \ f x" if inc_g: "\n x. 0 \ g n x \ g n x \ g (Suc n) x" and meas_g: "\n. g n \ borel_measurable lebesgue" and fin: "\n. finite(range (g n))" and lim: "\x. (\n. g n x) \ f x" for g n x proof - have "\r>0. \N. \n\N. dist (g n x) (f x) \ r" if "g n x > f x" proof - have g: "g n x \ g (N + n) x" for N by (rule transitive_stepwise_le) (use inc_g in auto) have "\na\N. g n x - f x \ dist (g na x) (f x)" for N apply (rule_tac x="N+n" in exI) using g [of N] by (auto simp: dist_norm) with that show ?thesis using diff_gt_0_iff_gt by blast qed with lim show ?thesis apply (auto simp: lim_sequentially) by (meson less_le_not_le not_le_imp_less) qed moreover let ?\ = "\n k. indicator {y. k/2^n \ f y \ f y < (k+1)/2^n}" let ?g = "\n x. (\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x)" have "\g. (\n x. 0 \ g n x \ g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \(\x. (\n. g n x) \ f x)" proof (intro exI allI conjI) show "0 \ ?g n x" for n x proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg) fix k::real assume "k \ \" and k: "\k\ \ 2 ^ (2*n)" show "0 \ k/2^n * ?\ n k x" using f \k \ \\ apply (auto simp: indicator_def field_split_simps Ints_def) apply (drule spec [where x=x]) using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"] by linarith qed show "?g n x \ ?g (Suc n) x" for n x proof - have "?g n x = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * (indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x + indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x))" by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x)" by (simp add: comm_monoid_add_class.sum.distrib algebra_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \ f y \ f y < (2 * k+1)/2 ^ Suc n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2 * k+1) + 1)/2 ^ Suc n} x)" by (force simp: field_simps indicator_def intro: sum.cong) also have "\ \ (\k | k \ \ \ \k\ \ 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x))" (is "?a + _ \ ?b") proof - have *: "\sum f I \ sum h I; a + sum h I \ b\ \ a + sum f I \ b" for I a b f and h :: "real\real" by linarith let ?h = "\k. (2*k+1)/2 ^ Suc n * (indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2*k+1) + 1)/2 ^ Suc n} x)" show ?thesis proof (rule *) show "(\k | k \ \ \ \k\ \ 2 ^ (2*n). 2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < (2 * k+1 + 1)/2 ^ Suc n} x) \ sum ?h {k \ \. \k\ \ 2 ^ (2*n)}" by (rule sum_mono) (simp add: indicator_def field_split_simps) next have \: "?a = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have \: "sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have 0: "(*) 2 ` {k \ \. P k} \ (\x. 2 * x + 1) ` {k \ \. P k} = {}" for P :: "real \ bool" proof - have "2 * i \ 2 * j + 1" for i j :: int by arith thus ?thesis unfolding Ints_def by auto (use of_int_eq_iff in fastforce) qed have "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" unfolding \ \ using finite_abs_int_segment [of "2 ^ (2*n)"] by (subst sum_Un) (auto simp: 0) also have "\ \ ?b" proof (rule sum_mono2) show "finite {k::real. k \ \ \ \k\ \ 2 ^ (2 * Suc n)}" by (rule finite_abs_int_segment) show "(*) 2 ` {k::real. k \ \ \ \k\ \ 2^(2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2^(2*n)} \ {k \ \. \k\ \ 2 ^ (2 * Suc n)}" apply auto using one_le_power [of "2::real" "2*n"] by linarith have *: "\x \ (S \ T) - U; \x. x \ S \ x \ U; \x. x \ T \ x \ U\ \ P x" for S T U P by blast have "0 \ b" if "b \ \" "f x * (2 * 2^n) < b + 1" for b proof - have "0 \ f x * (2 * 2^n)" by (simp add: f) also have "\ < b+1" by (simp add: that) finally show "0 \ b" using \b \ \\ by (auto simp: elim!: Ints_cases) qed then show "0 \ b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \ f y \ f y < (b + 1)/2 ^ Suc n} x" if "b \ {k \ \. \k\ \ 2 ^ (2 * Suc n)} - ((*) 2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)})" for b using that by (simp add: indicator_def divide_simps) qed finally show "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} \ ?b" . qed qed finally show ?thesis . qed show "?g n \ borel_measurable lebesgue" for n apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum) using leb_f sets_restrict_UNIV by auto show "finite (range (?g n))" for n proof - have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" for x proof (cases "\k. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ f x \ f x < (k+1)/2^n") case True then show ?thesis by (blast intro: indicator_sum_eq) next case False then have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) = 0" by auto then show ?thesis by force qed then have "range (?g n) \ ((\k. (k/2^n)) ` {k. k \ \ \ \k\ \ 2 ^ (2*n)})" by auto moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (intro finite_imageI finite_abs_int_segment) ultimately show ?thesis by (rule finite_subset) qed show "(\n. ?g n x) \ f x" for x proof (clarsimp simp add: lim_sequentially) fix e::real assume "e > 0" obtain N1 where N1: "2 ^ N1 > abs(f x)" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by fastforce have "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) < e" if "N1 + N2 \ n" for n proof - let ?m = "real_of_int \2^n * f x\" have "\?m\ \ 2^n * 2^N1" using N1 apply (simp add: f) by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power) also have "\ \ 2 ^ (2*n)" by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff power_add power_increasing_iff semiring_norm(76)) finally have m_le: "\?m\ \ 2 ^ (2*n)" . have "?m/2^n \ f x" "f x < (?m + 1)/2^n" by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos) then have eq: "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) = dist (?m/2^n) (f x)" by (subst indicator_sum_eq [of ?m]) (auto simp: m_le) have "\2^n\ * \?m/2^n - f x\ = \2^n * (?m/2^n - f x)\" by (simp add: abs_mult) also have "\ < 2 ^ N2 * e" using N2 by (simp add: divide_simps mult.commute) linarith also have "\ \ \2^n\ * e" using that \e > 0\ by auto finally have "dist (?m/2^n) (f x) < e" by (simp add: dist_norm) then show ?thesis using eq by linarith qed then show "\no. \n\no. dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k * ?\ n k x/2^n) (f x) < e" by force qed qed ultimately show ?rhs by metis next assume RHS: ?rhs with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq] show ?lhs by (blast intro: order_trans) qed subsection\Borel measurable Jacobian determinant\ lemma lemma_partial_derivatives0: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim0: "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x. x \ S - {0} \ norm x < e \ k * norm x \ \v \ x\)" shows "f x = 0" proof - interpret linear f by fact have "dim {x. f x = 0} \ DIM('a)" by (rule dim_subset_UNIV) moreover have False if less: "dim {x. f x = 0} < DIM('a)" proof - obtain d where "d \ 0" and d: "\y. f y = 0 \ d \ y = 0" using orthogonal_to_subspace_exists [OF less] orthogonal_def by (metis (mono_tags, lifting) mem_Collect_eq span_base) then obtain k where "k > 0" and k: "\e. e > 0 \ \y. y \ S - {0} \ norm y < e \ k * norm y \ \d \ y\" using lb by blast have "\h. \n. ((h n \ S \ h n \ 0 \ k * norm (h n) \ \d \ h n\) \ norm (h n) < 1 / real (Suc n)) \ norm (h (Suc n)) < norm (h n)" proof (rule dependent_nat_choice) show "\y. (y \ S \ y \ 0 \ k * norm y \ \d \ y\) \ norm y < 1 / real (Suc 0)" by simp (metis DiffE insertCI k not_less not_one_le_zero) qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto) then obtain \ where \: "\n. \ n \ S - {0}" and kd: "\n. k * norm(\ n) \ \d \ \ n\" and norm_lt: "\n. norm(\ n) < 1/(Suc n)" by force let ?\ = "\n. \ n /\<^sub>R norm (\ n)" have com: "\g. (\n. g n \ sphere (0::'a) 1) \ \l \ sphere 0 1. \\::nat\nat. strict_mono \ \ (g \ \) \ l" using compact_sphere compact_def by metis moreover have "\n. ?\ n \ sphere 0 1" using \ by auto ultimately obtain l::'a and \::"nat\nat" where l: "l \ sphere 0 1" and "strict_mono \" and to_l: "(?\ \ \) \ l" by meson moreover have "continuous (at l) (\x. (\d \ x\ - k))" by (intro continuous_intros) ultimately have lim_dl: "((\x. (\d \ x\ - k)) \ (?\ \ \)) \ (\d \ l\ - k)" by (meson continuous_imp_tendsto) have "\\<^sub>F i in sequentially. 0 \ ((\x. \d \ x\ - k) \ ((\n. \ n /\<^sub>R norm (\ n)) \ \)) i" using \ kd by (auto simp: field_split_simps) then have "k \ \d \ l\" using tendsto_lowerbound [OF lim_dl, of 0] by auto moreover have "d \ l = 0" proof (rule d) show "f l = 0" proof (rule LIMSEQ_unique [of "f \ ?\ \ \"]) have "isCont f l" using \linear f\ linear_continuous_at linear_conv_bounded_linear by blast then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ f l" unfolding comp_assoc using to_l continuous_imp_tendsto by blast have "\ \ 0" using norm_lt LIMSEQ_norm_0 by metis with \strict_mono \\ have "(\ \ \) \ 0" by (metis LIMSEQ_subseq_LIMSEQ) with lim0 \ have "((\x. f x /\<^sub>R norm x) \ (\ \ \)) \ 0" by (force simp: tendsto_at_iff_sequentially) then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ 0" by (simp add: o_def scale) qed qed ultimately show False using \k > 0\ by auto qed ultimately have dim: "dim {x. f x = 0} = DIM('a)" by force then show ?thesis using dim_eq_full by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis mem_Collect_eq module_hom_zero span_base span_raw_def) qed lemma lemma_partial_derivatives: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim: "((\x. f (x - a) /\<^sub>R norm (x - a)) \ 0) (at a within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x \ S - {a}. norm(a - x) < e \ k * norm(a - x) \ \v \ (x - a)\)" shows "f x = 0" proof - have "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within (\x. x-a) ` S)" using lim by (simp add: Lim_within dist_norm) then show ?thesis proof (rule lemma_partial_derivatives0 [OF \linear f\]) fix v :: "'a" assume v: "v \ 0" show "\k>0. \e>0. \x. x \ (\x. x - a) ` S - {0} \ norm x < e \ k * norm x \ \v \ x\" using lb [OF v] by (force simp: norm_minus_commute) qed qed proposition borel_measurable_partial_derivatives: fixes f :: "real^'m::{finite,wellorder} \ real^'n" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. (matrix(f' x)$m$n)) \ borel_measurable (lebesgue_on S)" proof - have contf: "continuous_on S f" using continuous_on_eq_continuous_within f has_derivative_continuous by blast have "{x \ S. (matrix (f' x)$m$n) \ b} \ sets lebesgue" for b proof (rule sets_negligible_symdiff) let ?T = "{x \ S. \e>0. \d>0. \A. A$m$n < b \ (\i j. A$i$j \ \) \ (\y \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x))}" let ?U = "S \ (\e \ {e \ \. e > 0}. \A \ {A. A$m$n < b \ (\i j. A$i$j \ \)}. \d \ {d \ \. 0 < d}. S \ (\y \ S. {x \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x)}))" have "?T = ?U" proof (intro set_eqI iffI) fix x assume xT: "x \ ?T" then show "x \ ?U" proof (clarsimp simp add:) fix q :: real assume "q \ \" "q > 0" then obtain d A where "d > 0" and A: "A $ m $ n < b" "\i j. A $ i $ j \ \" "\y. \y\S; norm (y - x) < d\ \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)" using xT by auto then obtain \ where "d > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with A show "\A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\s. s \ \ \ 0 < s \ (\y\S. norm (y - x) < s \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)))" by force qed next fix x assume xU: "x \ ?U" then show "x \ ?T" proof clarsimp fix e :: "real" assume "e > 0" then obtain \ where \: "e > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with xU obtain A r where "x \ S" and Ar: "A $ m $ n < b" "\i j. A $ i $ j \ \" "r \ \" "r > 0" and "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ \ * norm (y - x)" by (auto simp: split: if_split_asm) then have "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)" by (meson \e > \\ less_eq_real_def mult_right_mono norm_ge_zero order_trans) then show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" using \x \ S\ Ar by blast qed qed moreover have "?U \ sets lebesgue" proof - have coQ: "countable {e \ \. 0 < e}" using countable_Collect countable_rat by blast have ne: "{e \ \. (0::real) < e} \ {}" using zero_less_one Rats_1 by blast have coA: "countable {A. A $ m $ n < b \ (\i j. A $ i $ j \ \)}" proof (rule countable_subset) show "countable {A. \i j. A $ i $ j \ \}" using countable_vector [OF countable_vector, of "\i j. \"] by (simp add: countable_rat) qed blast have *: "\U \ {} \ closedin (top_of_set S) (S \ \ U)\ \ closedin (top_of_set S) (S \ \ U)" for U by fastforce have eq: "{x::(real,'m)vec. P x \ (Q x \ R x)} = {x. P x \ \ Q x} \ {x. P x \ R x}" for P Q R by auto have sets: "S \ (\y\S. {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}) \ sets lebesgue" for e A d proof - have clo: "closedin (top_of_set S) {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}" for y proof - have cont1: "continuous_on S (\x. norm (y - x))" and cont2: "continuous_on S (\x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))" by (force intro: contf continuous_intros)+ have clo1: "closedin (top_of_set S) {x \ S. d \ norm (y - x)}" using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def) have clo2: "closedin (top_of_set S) {x \ S. norm (f y - f x - (A *v y - A *v x)) \ e * norm (y - x)}" using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def) show ?thesis by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2) qed show ?thesis by (rule lebesgue_closedin [of S]) (force intro: * S clo)+ qed show ?thesis by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto qed ultimately show "?T \ sets lebesgue" by simp let ?M = "(?T - {x \ S. matrix (f' x) $ m $ n \ b} \ ({x \ S. matrix (f' x) $ m $ n \ b} - ?T))" let ?\ = "\x v. \\>0. \e>0. \y \ S-{x}. norm (x - y) < e \ \v \ (y - x)\ < \ * norm (x - y)" have nN: "negligible {x \ S. \v\0. ?\ x v}" unfolding negligible_eq_zero_density proof clarsimp fix x v and r e :: "real" assume "x \ S" "v \ 0" "r > 0" "e > 0" and Theta [rule_format]: "?\ x v" moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0" by (simp add: \v \ 0\ \e > 0\) ultimately obtain d where "d > 0" and dless: "\y. \y \ S - {x}; norm (x - y) < d\ \ \v \ (y - x)\ < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)" by metis let ?W = "ball x (min d r) \ {y. \v \ (y - x)\ < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}" have "open {x. \v \ (x - a)\ < b}" for a b by (intro open_Collect_less continuous_intros) show "\d>0. d \ r \ (\U. {x' \ S. \v\0. ?\ x' v} \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d))" proof (intro exI conjI) show "0 < min d r" "min d r \ r" using \r > 0\ \d > 0\ by auto show "{x' \ S. \v. v \ 0 \ (\\>0. \e>0. \z\S - {x'}. norm (x' - z) < e \ \v \ (z - x')\ < \ * norm (x' - z))} \ ball x (min d r) \ ?W" proof (clarsimp simp: dist_norm norm_minus_commute) fix y w assume "y \ S" "w \ 0" and less [rule_format]: "\\>0. \e>0. \z\S - {y}. norm (y - z) < e \ \w \ (z - y)\ < \ * norm (y - z)" and d: "norm (y - x) < d" and r: "norm (y - x) < r" show "\v \ (y - x)\ < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" proof (cases "y = x") case True with \r > 0\ \d > 0\ \e > 0\ \v \ 0\ show ?thesis by simp next case False have "\v \ (y - x)\ < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)" apply (rule dless) using False \y \ S\ d by (auto simp: norm_minus_commute) also have "\ \ norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" using d r \e > 0\ by (simp add: field_simps norm_minus_commute mult_left_mono) finally show ?thesis . qed qed show "?W \ lmeasurable" by (simp add: fmeasurable_Int_fmeasurable borel_open) obtain k::'m where True by metis obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis define b where "b \ norm v" have "b > 0" using \v \ 0\ by (auto simp: b_def) obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)" by (metis UNIV_I b_def T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv) let ?v = "\ i. min d r / CARD('m)" let ?v' = "\ i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r" let ?x' = "inv T x" let ?W' = "(ball ?x' (min d r) \ {y. \(y - ?x')$k\ < e * min d r / (2 * CARD('m) ^ CARD('m))})" have abs: "x - e \ y \ y \ x + e \ abs(y - x) \ e" for x y e::real by auto have "?W = T ` ?W'" proof - have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)" by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f) have 2: "{y. \v \ (y - x)\ < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} = T ` {y. \y $ k - ?x' $ k\ < e * min d r / (2 * real CARD('m) ^ CARD('m))}" proof - have *: "\T (b *\<^sub>R axis k 1) \ (y - x)\ = b * \inv T y $ k - ?x' $ k\" for y proof - have "\T (b *\<^sub>R axis k 1) \ (y - x)\ = \(b *\<^sub>R axis k 1) \ inv T (y - x)\" by (metis (no_types, hide_lams) b_def eqb invT orthogonal_transformation_def v) also have "\ = b * \(axis k 1) \ inv T (y - x)\" using \b > 0\ by (simp add: abs_mult) also have "\ = b * \inv T y $ k - ?x' $ k\" using orthogonal_transformation_linear [OF invT] by (simp add: inner_axis' linear_diff) finally show ?thesis by simp qed show ?thesis using v b_def [symmetric] using \b > 0\ by (simp add: * bij_image_Collect_eq [OF \bij T\] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right) qed show ?thesis using \b > 0\ by (simp add: image_Int \inj T\ 1 2 b_def [symmetric]) qed moreover have "?W' \ lmeasurable" by (auto intro: fmeasurable_Int_fmeasurable) ultimately have "measure lebesgue ?W = measure lebesgue ?W'" by (metis measure_orthogonal_image T) also have "\ \ measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))" proof (rule measure_mono_fmeasurable) show "?W' \ cbox (?x' - ?v') (?x' + ?v')" apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff) by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component) qed auto also have "\ \ e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))" proof - have "cbox (?x' - ?v) (?x' + ?v) \ {}" using \r > 0\ \d > 0\ by (auto simp: interval_eq_empty_cart divide_less_0_iff) with \r > 0\ \d > 0\ \e > 0\ show ?thesis apply (simp add: content_cbox_if_cart mem_box_cart) apply (auto simp: prod_nonneg) apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm) done qed also have "\ \ e/2 * measure lebesgue (cball ?x' (min d r))" proof (rule mult_left_mono [OF measure_mono_fmeasurable]) have *: "norm (?x' - y) \ min d r" if y: "\i. \?x' $ i - y $ i\ \ min d r / real CARD('m)" for y proof - have "norm (?x' - y) \ (\i\UNIV. \(?x' - y) $ i\)" by (rule norm_le_l1_cart) also have "\ \ real CARD('m) * (min d r / real CARD('m))" by (rule sum_bounded_above) (use y in auto) finally show ?thesis by simp qed show "cbox (?x' - ?v) (?x' + ?v) \ cball ?x' (min d r)" apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *) by (simp add: abs_diff_le_iff abs_minus_commute) qed (use \e > 0\ in auto) also have "\ < e * content (cball ?x' (min d r))" using \r > 0\ \d > 0\ \e > 0\ by (auto intro: content_cball_pos) also have "\ = e * content (ball x (min d r))" using \r > 0\ \d > 0\ content_ball_conv_unit_ball[of "min d r" "inv T x"] content_ball_conv_unit_ball[of "min d r" x] by (simp add: content_cball_conv_ball) finally show "measure lebesgue ?W < e * content (ball x (min d r))" . qed qed have *: "(\x. (x \ S) \ (x \ T \ x \ U)) \ (T - U) \ (U - T) \ S" for S T U :: "(real,'m) vec set" by blast have MN: "?M \ {x \ S. \v\0. ?\ x v}" proof (rule *) fix x assume x: "x \ {x \ S. \v\0. ?\ x v}" show "(x \ ?T) \ (x \ {x \ S. matrix (f' x) $ m $ n \ b})" proof (cases "x \ S") case True then have x: "\ ?\ x v" if "v \ 0" for v using x that by force show ?thesis proof (rule iffI; clarsimp) assume b: "\e>0. \d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" (is "\e>0. \d>0. \A. ?\ e d A") then have "\k. \d>0. \A. ?\ (1 / Suc k) d A" by (metis (no_types, hide_lams) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff) then obtain \ A where \: "\k. \ k > 0" and Ab: "\k. A k $ m $ n < b" and A: "\k y. \y \ S; norm (y - x) < \ k\ \ norm (f y - f x - A k *v (y - x)) \ 1/(Suc k) * norm (y - x)" by metis have "\i j. \a. (\n. A n $ i $ j) \ a" proof (intro allI) fix i j have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n by (metis cart_eq_inner_axis matrix_vector_mul_component) let ?CA = "{x. Cauchy (\n. (A n) *v x)}" have "subspace ?CA" unfolding subspace_def convergent_eq_Cauchy [symmetric] by (force simp: algebra_simps intro: tendsto_intros) then have CA_eq: "?CA = span ?CA" by (metis span_eq_iff) also have "\ = UNIV" proof - have "dim ?CA \ CARD('m)" using dim_subset_UNIV[of ?CA] by auto moreover have "False" if less: "dim ?CA < CARD('m)" proof - obtain d where "d \ 0" and d: "\y. y \ span ?CA \ orthogonal d y" using less by (force intro: orthogonal_to_subspace_exists [of ?CA]) with x [OF \d \ 0\] obtain \ where "\ > 0" and \: "\e. e > 0 \ \y \ S - {x}. norm (x - y) < e \ \ * norm (x - y) \ \d \ (y - x)\" by (fastforce simp: not_le Bex_def) obtain \ z where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \x: "\ \ x" and z: "(\n. (\ n - x) /\<^sub>R norm (\ n - x)) \ z" proof - have "\\. (\i. (\ i \ S - {x} \ \ * norm(\ i - x) \ \d \ (\ i - x)\ \ norm(\ i - x) < 1/Suc i) \ norm(\(Suc i) - x) < norm(\ i - x))" proof (rule dependent_nat_choice) show "\y. y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1 / Suc 0" using \ [of 1] by (auto simp: dist_norm norm_minus_commute) next fix y i assume "y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1/Suc i" then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0" by auto then obtain y' where "y' \ S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))" "\ * norm (x - y') \ \d \ (y' - x)\" using \ by metis with \ show "\y'. (y' \ S - {x} \ \ * norm (y' - x) \ \d \ (y' - x)\ \ norm (y' - x) < 1/(Suc (Suc i))) \ norm (y' - x) < norm (y - x)" by (auto simp: dist_norm norm_minus_commute) qed then obtain \ where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \conv: "\i. norm(\ i - x) < 1/(Suc i)" by blast let ?f = "\i. (\ i - x) /\<^sub>R norm (\ i - x)" have "?f i \ sphere 0 1" for i using \Sx by auto then obtain l \ where "l \ sphere 0 1" "strict_mono \" and l: "(?f \ \) \ l" using compact_sphere [of "0::(real,'m) vec" 1] unfolding compact_def by meson show thesis proof show "(\ \ \) i \ S - {x}" "\ * norm ((\ \ \) i - x) \ \d \ ((\ \ \) i - x)\" for i using \Sx \le by auto have "\ \ x" proof (clarsimp simp add: LIMSEQ_def dist_norm) fix r :: "real" assume "r > 0" with real_arch_invD obtain no where "no \ 0" "real no > 1/r" by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2) with \conv show "\no. \n\no. norm (\ n - x) < r" by (metis \r > 0\ add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc) qed with \strict_mono \\ show "(\ \ \) \ x" by (metis LIMSEQ_subseq_LIMSEQ) show "(\n. ((\ \ \) n - x) /\<^sub>R norm ((\ \ \) n - x)) \ l" using l by (auto simp: o_def) qed qed have "isCont (\x. (\d \ x\ - \)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\y. \d \ ((\ y - x) /\<^sub>R norm (\ y - x))\ - \) \ \d \ z\ - \" by auto moreover have "\\<^sub>F i in sequentially. 0 \ \d \ ((\ i - x) /\<^sub>R norm (\ i - x))\ - \" proof (rule eventuallyI) fix n show "0 \ \d \ ((\ n - x) /\<^sub>R norm (\ n - x))\ - \" using \le [of n] \Sx by (auto simp: abs_mult divide_simps) qed ultimately have "\ \ \d \ z\" using tendsto_lowerbound [where a=0] by fastforce have "Cauchy (\n. (A n) *v z)" proof (clarsimp simp add: Cauchy_def) fix \ :: "real" assume "0 < \" then obtain N::nat where "N > 0" and N: "\/2 > 1/N" by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse) show "\M. \m\M. \n\M. dist (A m *v z) (A n *v z) < \" proof (intro exI allI impI) fix i j assume ij: "N \ i" "N \ j" let ?V = "\i k. A i *v ((\ k - x) /\<^sub>R norm (\ k - x))" have "\\<^sub>F k in sequentially. dist (\ k) x < min (\ i) (\ j)" using \x [unfolded tendsto_iff] by (meson min_less_iff_conj \) then have even: "\\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \ 0" proof (rule eventually_mono, clarsimp) fix p assume p: "dist (\ p) x < \ i" "dist (\ p) x < \ j" let ?C = "\k. f (\ p) - f x - A k *v (\ p - x)" have "norm ((A i - A j) *v (\ p - x)) = norm (?C j - ?C i)" by (simp add: algebra_simps) also have "\ \ norm (?C j) + norm (?C i)" using norm_triangle_ineq4 by blast also have "\ \ 1/(Suc j) * norm (\ p - x) + 1/(Suc i) * norm (\ p - x)" by (metis A Diff_iff \Sx dist_norm p add_mono) also have "\ \ 1/N * norm (\ p - x) + 1/N * norm (\ p - x)" apply (intro add_mono mult_right_mono) using ij \N > 0\ by (auto simp: field_simps) also have "\ = 2 / N * norm (\ p - x)" by simp finally have no_le: "norm ((A i - A j) *v (\ p - x)) \ 2 / N * norm (\ p - x)" . have "norm (?V i p - ?V j p) = norm ((A i - A j) *v ((\ p - x) /\<^sub>R norm (\ p - x)))" by (simp add: algebra_simps) also have "\ = norm ((A i - A j) *v (\ p - x)) / norm (\ p - x)" by (simp add: divide_inverse matrix_vector_mult_scaleR) also have "\ \ 2 / N" using no_le by (auto simp: field_split_simps) finally show "norm (?V i p - ?V j p) \ 2 / N" . qed have "isCont (\w. (norm(A i *v w - A j *v w) - 2 / N)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\w. norm (A i *v ((\ w - x) /\<^sub>R norm (\ w - x)) - A j *v ((\ w - x) /\<^sub>R norm (\ w - x))) - 2 / N) \ norm (A i *v z - A j *v z) - 2 / N" by auto have "dist (A i *v z) (A j *v z) \ 2 / N" using tendsto_upperbound [OF lim even] by (auto simp: dist_norm) with N show "dist (A i *v z) (A j *v z) < \" by linarith qed qed then have "d \ z = 0" using CA_eq d orthogonal_def by auto then show False using \0 < \\ \\ \ \d \ z\\ by auto qed ultimately show ?thesis using dim_eq_full by fastforce qed finally have "?CA = UNIV" . then have "Cauchy (\n. (A n) *v axis j 1)" by auto then obtain L where "(\n. A n *v axis j 1) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) then have "(\x. (A x *v axis j 1) $ i) \ L $ i" by (rule tendsto_vec_nth) then show "\a. (\n. A n $ i $ j) \ a" by (force simp: vax) qed then obtain B where B: "\i j. (\n. A n $ i $ j) \ B $ i $ j" by (auto simp: lambda_skolem) have lin_df: "linear (f' x)" and lim_df: "((\y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \ 0) (at x within S)" using \x \ S\ assms by (auto simp: has_derivative_within linear_linear) moreover interpret linear "f' x" by fact have "(matrix (f' x) - B) *v w = 0" for w proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"]) show "linear ((*v) (matrix (f' x) - B))" by (rule matrix_vector_mul_linear) have "((\y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" using tendsto_minus [OF lim_df] by (simp add: field_split_simps) then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \ 0) (at x within S)" proof (rule Lim_transform) have "((\y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \ 0) (at x within S)" proof (clarsimp simp add: Lim_within dist_norm) fix e :: "real" assume "e > 0" then obtain q::nat where "q \ 0" and qe2: "1/q < e/2" by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral) let ?g = "\p. sum (\i. sum (\j. abs((A p - B)$i$j)) UNIV) UNIV" have "(\k. onorm (\y. (A k - B) *v y)) \ 0" proof (rule Lim_null_comparison) show "\\<^sub>F k in sequentially. norm (onorm (\y. (A k - B) *v y)) \ ?g k" proof (rule eventually_sequentiallyI) fix k :: "nat" assume "0 \ k" have "0 \ onorm ((*v) (A k - B))" using matrix_vector_mul_bounded_linear by (rule onorm_pos_le) then show "norm (onorm ((*v) (A k - B))) \ (\i\UNIV. \j\UNIV. \(A k - B) $ i $ j\)" by (simp add: onorm_le_matrix_component_sum del: vector_minus_component) qed next show "?g \ 0" using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum) qed with \e > 0\ obtain p where "\n. n \ p \ \onorm ((*v) (A n - B))\ < e/2" unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral) then have pqe2: "\onorm ((*v) (A (p + q) - B))\ < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*) using le_add1 by blast show "\d>0. \y\S. y \ x \ norm (y - x) < d \ inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" proof (intro exI, safe) show "0 < \(p + q)" by (simp add: \) next fix y assume y: "y \ S" "norm (y - x) < \(p + q)" and "y \ x" have *: "\norm(b - c) < e - d; norm(y - x - b) \ d\ \ norm(y - x - c) < e" for b c d e x and y:: "real^'n" using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)" proof (rule *) show "norm (f y - f x - A (p + q) *v (y - x)) \ norm (y - x) / (Suc (p + q))" using A [OF y] by simp have "norm (A (p + q) *v (y - x) - B *v (y - x)) \ onorm(\x. (A(p + q) - B) *v x) * norm(y - x)" by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm) also have "\ < (e/2) * norm (y - x)" using \y \ x\ pqe2 by auto also have "\ \ (e - 1 / (Suc (p + q))) * norm (y - x)" proof (rule mult_right_mono) have "1 / Suc (p + q) \ 1 / q" using \q \ 0\ by (auto simp: field_split_simps) also have "\ < e/2" using qe2 by auto finally show "e / 2 \ e - 1 / real (Suc (p + q))" by linarith qed auto finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))" by (simp add: algebra_simps) qed then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" using \y \ x\ by (simp add: field_split_simps algebra_simps) qed qed then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" by (simp add: algebra_simps diff lin_df scalar_mult_eq_scaleR) qed qed (use x in \simp; auto simp: not_less\) ultimately have "f' x = (*v) B" by (force simp: algebra_simps scalar_mult_eq_scaleR) show "matrix (f' x) $ m $ n \ b" proof (rule tendsto_upperbound [of "\i. (A i $ m $ n)" _ sequentially]) show "(\i. A i $ m $ n) \ matrix (f' x) $ m $ n" by (simp add: B \f' x = (*v) B\) show "\\<^sub>F i in sequentially. A i $ m $ n \ b" by (simp add: Ab less_eq_real_def) qed auto next fix e :: "real" assume "x \ S" and b: "matrix (f' x) $ m $ n \ b" and "e > 0" then obtain d where "d>0" and d: "\y. y\S \ 0 < dist y x \ dist y x < d \ norm (f y - f x - f' x (y - x)) / (norm (y - x)) < e/2" using f [OF \x \ S\] by (simp add: Deriv.has_derivative_at_within Lim_within) (auto simp add: field_simps dest: spec [of _ "e/2"]) let ?A = "matrix(f' x) - (\ i j. if i = m \ j = n then e / 4 else 0)" obtain B where BRats: "\i j. B$i$j \ \" and Bo_e6: "onorm((*v) (?A - B)) < e/6" using matrix_rational_approximation \e > 0\ by (metis zero_less_divide_iff zero_less_numeral) show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" proof (intro exI conjI ballI allI impI) show "d>0" by (rule \d>0\) show "B $ m $ n < b" proof - have "\matrix ((*v) (?A - B)) $ m $ n\ \ onorm ((*v) (?A - B))" using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis then show ?thesis using b Bo_e6 by simp qed show "B $ i $ j \ \" for i j using BRats by auto show "norm (f y - f x - B *v (y - x)) \ e * norm (y - x)" if "y \ S" and y: "norm (y - x) < d" for y proof (cases "y = x") case True then show ?thesis by simp next case False have *: "norm(d' - d) \ e/2 \ norm(y - (x + d')) < e/2 \ norm(y - x - d) \ e" for d d' e and x y::"real^'n" using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp show ?thesis proof (rule *) have split246: "\norm y \ e / 6; norm(x - y) \ e / 4\ \ norm x \ e/2" if "e > 0" for e and x y :: "real^'n" using norm_triangle_le [of y "x-y" "e/2"] \e > 0\ by simp have "linear (f' x)" using True f has_derivative_linear by blast then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))" by (simp add: matrix_vector_mult_diff_rdistrib) also have "\ \ (e * norm (y - x)) / 2" proof (rule split246) have "norm ((?A - B) *v (y - x)) / norm (y - x) \ onorm(\x. (?A - B) *v x)" by (rule le_onorm) auto also have "\ < e/6" by (rule Bo_e6) finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" . then show "norm ((?A - B) *v (y - x)) \ e * norm (y - x) / 6" by (simp add: field_split_simps False) have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x))" by (simp add: algebra_simps) also have "\ = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))" proof - have "(\j\UNIV. (if i = m \ j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i proof (cases "i=m") case True then show ?thesis by (auto simp: if_distrib [of "\z. z * _"] cong: if_cong) next case False then show ?thesis by (simp add: axis_def) qed then have "(\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)" by (auto simp: vec_eq_iff matrix_vector_mult_def) then show ?thesis by metis qed also have "\ \ e * norm (y - x) / 4" using \e > 0\ apply (simp add: norm_mult abs_mult) by (metis component_le_norm_cart vector_minus_component) finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \ e * norm (y - x) / 4" . show "0 < e * norm (y - x)" by (simp add: False \e > 0\) qed finally show "norm (f' x (y - x) - B *v (y - x)) \ (e * norm (y - x)) / 2" . show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2" using False d [OF \y \ S\] y by (simp add: dist_norm field_simps) qed qed qed qed qed auto qed show "negligible ?M" using negligible_subset [OF nN MN] . qed then show ?thesis by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms) qed theorem borel_measurable_det_Jacobian: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. det(matrix(f' x))) \ borel_measurable (lebesgue_on S)" unfolding det_def by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S]) text\The localisation wrt S uses the same argument for many similar results.\ (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) theorem borel_measurable_lebesgue_on_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" proof - have "{x. (if x \ S then f x else 0) \ T} \ sets lebesgue \ {x \ S. f x \ T} \ sets lebesgue" if "T \ sets borel" for T proof (cases "0 \ T") case True then have "{x \ S. f x \ T} = {x. (if x \ S then f x else 0) \ T} \ S" "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T} \ -S" by auto then show ?thesis by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un) next case False then have "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T}" by auto then show ?thesis by auto qed then show ?thesis unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric] by blast qed lemma sets_lebesgue_almost_borel: assumes "S \ sets lebesgue" obtains B N where "B \ sets borel" "negligible N" "B \ N = S" proof - obtain T N N' where "S = T \ N" "N \ N'" "N' \ null_sets lborel" "T \ sets borel" using sets_completionE [OF assms] by auto then show thesis by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) qed lemma double_lebesgue_sets: assumes S: "S \ sets lebesgue" and T: "T \ sets lebesgue" and fim: "f ` S \ T" shows "(\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue) \ f \ borel_measurable (lebesgue_on S) \ (\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue)" (is "?lhs \ _ \ ?rhs") unfolding borel_measurable_lebesgue_on_preimage_borel [OF S] proof (intro iffI allI conjI impI, safe) fix V :: "'b set" assume *: "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "V \ sets borel" then have V: "V \ sets lebesgue" by simp have "{x \ S. f x \ V} = {x \ S. f x \ T \ V}" using fim by blast also have "{x \ S. f x \ T \ V} \ sets lebesgue" using T V * le_inf_iff by blast finally show "{x \ S. f x \ V} \ sets lebesgue" . next fix U :: "'b set" assume "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" "negligible U" "U \ T" then show "{x \ S. f x \ U} \ sets lebesgue" using negligible_imp_sets by blast next fix U :: "'b set" assume 1 [rule_format]: "(\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" and 2 [rule_format]: "\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "U \ sets lebesgue" "U \ T" then obtain C N where C: "C \ sets borel \ negligible N \ C \ N = U" using sets_lebesgue_almost_borel by metis then have "{x \ S. f x \ C} \ sets lebesgue" by (blast intro: 1) moreover have "{x \ S. f x \ N} \ sets lebesgue" using C \U \ T\ by (blast intro: 2) moreover have "{x \ S. f x \ C \ N} = {x \ S. f x \ C} \ {x \ S. f x \ N}" by auto ultimately show "{x \ S. f x \ U} \ sets lebesgue" using C by auto qed subsection\Simplest case of Sard's theorem (we don't need continuity of derivative)\ lemma Sard_lemma00: fixes P :: "'b::euclidean_space set" assumes "a \ 0" and a: "a *\<^sub>R i \ 0" and i: "i \ Basis" and P: "P \ {x. a *\<^sub>R i \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (DIM('b) - 1)" proof - have "a > 0" using assms by simp let ?v = "(\j\Basis. (if j = i then e else m) *\<^sub>R j)" show thesis proof have "- e \ x \ i" "x \ i \ e" if "t \ P" "norm (x - t) \ e" for x t using \a > 0\ that Basis_le_norm [of i "x-t"] P i by (auto simp: inner_commute algebra_simps) moreover have "- m \ x \ j" "x \ j \ m" if "norm x \ m" "t \ P" "norm (x - t) \ e" "j \ Basis" and "j \ i" for x t j using that Basis_le_norm [of j x] by auto ultimately show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ cbox (-?v) ?v" by (auto simp: mem_box) have *: "\k\Basis. - ?v \ k \ ?v \ k" using \0 \ m\ \0 \ e\ by (auto simp: inner_Basis) have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)" by (metis DIM_positive Suc_pred power_Suc) show "measure lebesgue (cbox (-?v) ?v) \ 2 * e * (2 * m) ^ (DIM('b) - 1)" using \i \ Basis\ by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2) qed blast qed text\As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\ lemma Sard_lemma0: fixes P :: "(real^'n::{finite,wellorder}) set" assumes "a \ 0" and P: "P \ {x. a \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis have Tinv [simp]: "T (inv T x) = x" for x by (simp add: T orthogonal_transformation_surj surj_f_inv_f) obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ T-`P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e]) have "norm a *\<^sub>R axis k 1 \ x = 0" if "T x \ P" for x proof - have "a \ T x = 0" using P that by blast then show ?thesis by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def) qed then show "T -` P \ {x. norm a *\<^sub>R axis k 1 \ x = 0}" by auto qed (use assms T in auto) show thesis proof show "T ` S \ lmeasurable" using S measurable_orthogonal_image T by blast have "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" proof clarsimp fix x t assume "norm x \ m" "t \ P" "norm (x - t) \ e" then have "norm (inv T x) \ m" using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm) moreover have "\t\T -` P. norm (inv T x - t) \ e" proof have "T (inv T x - inv T t) = x - t" using T linear_diff orthogonal_transformation_def by (metis (no_types, hide_lams) Tinv) then have "norm (inv T x - inv T t) = norm (x - t)" by (metis T orthogonal_transformation_norm) then show "norm (inv T x - inv T t) \ e" using \norm (x - t) \ e\ by linarith next show "inv T t \ T -` P" using \t \ P\ by force qed ultimately show "x \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" by force qed then show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` S" using image_mono [OF subS] by (rule order_trans) show "measure lebesgue (T ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" using mS T by (simp add: S measure_orthogonal_image) qed qed text\As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\ lemma Sard_lemma1: fixes P :: "(real^'n::{finite,wellorder}) set" assumes P: "dim P < CARD('n)" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm(z - w) \ m \ (\t \ P. norm(z - w - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain a where "a \ 0" "P \ {x. a \ x = 0}" using lowdim_subset_hyperplane [of P] P span_base by auto then obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" by (rule Sard_lemma0 [OF _ _ \0 \ m\ \0 \ e\]) show thesis proof show "(+)w ` S \ lmeasurable" by (metis measurable_translation S) show "{z. norm (z - w) \ m \ (\t\P. norm (z - w - t) \ e)} \ (+)w ` S" using subS by force show "measure lebesgue ((+)w ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" by (metis measure_translation mS) qed qed lemma Sard_lemma2: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" (is "?m \ ?n") and "B > 0" "bounded S" and derS: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" and B: "\x. x \ S \ onorm(f' x) \ B" shows "negligible(f ` S)" proof - have lin_f': "\x. x \ S \ linear(f' x)" using derS has_derivative_linear by blast show ?thesis proof (clarsimp simp add: negligible_outer_le) fix e :: "real" assume "e > 0" obtain c where csub: "S \ cbox (- (vec c)) (vec c)" and "c > 0" proof - obtain b where b: "\x. x \ S \ norm x \ b" using \bounded S\ by (auto simp: bounded_iff) show thesis proof have "- \b\ - 1 \ x $ i \ x $ i \ \b\ + 1" if "x \ S" for x i using component_le_norm_cart [of x i] b [OF that] by auto then show "S \ cbox (- vec (\b\ + 1)) (vec (\b\ + 1))" by (auto simp: mem_box_cart) qed auto qed then have box_cc: "box (- (vec c)) (vec c) \ {}" and cbox_cc: "cbox (- (vec c)) (vec c) \ {}" by (auto simp: interval_eq_empty_cart) obtain d where "d > 0" "d \ B" and d: "(d * 2) * (4 * B) ^ (?n - 1) \ e / (2*c) ^ ?m / ?m ^ ?m" apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"]) using \B > 0\ \c > 0\ \e > 0\ by (simp_all add: divide_simps min_mult_distrib_right) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. y \ S \ norm(y - x) < r \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)))" for x proof (cases "x \ S") case True then obtain r where "r > 0" and "\y. \y \ S; norm (y - x) < r\ \ norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using derS \d > 0\ by (force simp: has_derivative_within_alt) then show ?thesis by (rule_tac x="min r (1/2)" in exI) simp next case False then show ?thesis by (rule_tac x="1/2" in exI) simp qed then obtain r where r12: "\x. 0 < r x \ r x \ 1/2" and r: "\x y. \x \ S; y \ S; norm(y - x) < r x\ \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)" by metis then have ga: "gauge (\x. ball x (r x))" by (auto simp: gauge_def) obtain \ where \: "countable \" and sub_cc: "\\ \ cbox (- vec c) (vec c)" and cbox: "\K. K \ \ \ interior K \ {} \ (\u v. K = cbox u v)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ ball x (r x)" and close: "\u v. cbox u v \ \ \ \n. \i::'m. v $ i - u $ i = 2*c / 2^n" and covers: "S \ \\" apply (rule covering_lemma [OF csub box_cc ga]) apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric]) done let ?\ = "measure lebesgue" have "\T. T \ lmeasurable \ f ` (K \ S) \ T \ ?\ T \ e / (2*c) ^ ?m * ?\ K" if "K \ \" for K proof - obtain u v where uv: "K = cbox u v" using cbox \K \ \\ by blast then have uv_ne: "cbox u v \ {}" using cbox that by fastforce obtain x where x: "x \ S \ cbox u v" "cbox u v \ ball x (r x)" using \K \ \\ covered uv by blast then have "dim (range (f' x)) < ?n" using rank_dim_range [of "matrix (f' x)"] x rank[of x] by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f') then obtain T where T: "T \ lmeasurable" and subT: "{z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))} \ T" and measT: "?\ T \ (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" (is "_ \ ?DVU") apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"]) using \B > 0\ \d > 0\ by simp_all show ?thesis proof (intro exI conjI) have "f ` (K \ S) \ {z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))}" unfolding uv proof (clarsimp simp: mult.assoc, intro conjI) fix y assume y: "y \ cbox u v" and "y \ S" then have "norm (y - x) < r x" by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2)) then have le_dyx: "norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using r [of x y] x \y \ S\ by blast have yx_le: "norm (y - x) \ norm (v - u)" proof (rule norm_le_componentwise_cart) show "norm ((y - x) $ i) \ norm ((v - u) $ i)" for i using x y by (force simp: mem_box_cart dest!: spec [where x=i]) qed have *: "\norm(y - x - z) \ d; norm z \ B; d \ B\ \ norm(y - x) \ 2 * B" for x y z :: "real^'n::_" and d B using norm_triangle_ineq2 [of "y - x" z] by auto show "norm (f y - f x) \ 2 * (B * norm (v - u))" proof (rule * [OF le_dyx]) have "norm (f' x (y - x)) \ onorm (f' x) * norm (y - x)" using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1)) also have "\ \ B * norm (v - u)" proof (rule mult_mono) show "onorm (f' x) \ B" using B x by blast qed (use \B > 0\ yx_le in auto) finally show "norm (f' x (y - x)) \ B * norm (v - u)" . show "d * norm (y - x) \ B * norm (v - u)" using \B > 0\ by (auto intro: mult_mono [OF \d \ B\ yx_le]) qed show "\t. norm (f y - f x - f' x t) \ d * norm (v - u)" apply (rule_tac x="y-x" in exI) using \d > 0\ yx_le le_dyx mult_left_mono [where c=d] - by (meson order_trans real_mult_le_cancel_iff2) + by (meson order_trans mult_le_cancel_iff2) qed with subT show "f ` (K \ S) \ T" by blast show "?\ T \ e / (2*c) ^ ?m * ?\ K" proof (rule order_trans [OF measT]) have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n" using \c > 0\ apply (simp add: algebra_simps) by (metis Suc_pred power_Suc zero_less_card_finite) also have "\ \ (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n" by (rule mult_right_mono [OF d]) auto also have "\ \ e / (2*c) ^ ?m * ?\ K" proof - have "u \ ball (x) (r x)" "v \ ball x (r x)" using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+ moreover have "r x \ 1/2" using r12 by auto ultimately have "norm (v - u) \ 1" using norm_triangle_half_r [of x u 1 v] by (metis (no_types, hide_lams) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball) then have "norm (v - u) ^ ?n \ norm (v - u) ^ ?m" by (simp add: power_decreasing [OF mlen]) also have "\ \ ?\ K * real (?m ^ ?m)" proof - obtain n where n: "\i. v$i - u$i = 2 * c / 2^n" using close [of u v] \K \ \\ uv by blast have "norm (v - u) ^ ?m \ (\i\UNIV. \(v - u) $ i\) ^ ?m" by (intro norm_le_l1_cart power_mono) auto also have "\ \ (\i\UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)" by (simp add: n field_simps \c > 0\ less_eq_real_def) also have "\ = ?\ K * real (?m ^ ?m)" by (simp add: uv uv_ne content_cbox_cart) finally show ?thesis . qed finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \ ?\ K" by (simp add: field_split_simps) show ?thesis using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \c > 0\ \e > 0\ by auto qed finally show "?DVU \ e / (2*c) ^ ?m * ?\ K" . qed qed (use T in auto) qed then obtain g where meas_g: "\K. K \ \ \ g K \ lmeasurable" and sub_g: "\K. K \ \ \ f ` (K \ S) \ g K" and le_g: "\K. K \ \ \ ?\ (g K) \ e / (2*c)^?m * ?\ K" by metis have le_e: "?\ (\i\\. g i) \ e" if "\ \ \" "finite \" for \ proof - have "?\ (\i\\. g i) \ (\i\\. ?\ (g i))" using meas_g \\ \ \\ by (auto intro: measure_UNION_le [OF \finite \\]) also have "\ \ (\K\\. e / (2*c) ^ ?m * ?\ K)" using \\ \ \\ sum_mono [OF le_g] by (meson le_g subsetCE sum_mono) also have "\ = e / (2*c) ^ ?m * (\K\\. ?\ K)" by (simp add: sum_distrib_left) also have "\ \ e" proof - have "\ division_of \\" proof (rule division_ofI) show "K \ \\" "K \ {}" "\a b. K = cbox a b" if "K \ \" for K using \K \ \\ covered cbox \\ \ \\ by (auto simp: Union_upper) show "interior K \ interior L = {}" if "K \ \" and "L \ \" and "K \ L" for K L by (metis (mono_tags, lifting) \\ \ \\ pairwiseD djointish pairwise_subset that) qed (use that in auto) then have "sum ?\ \ \ ?\ (\\)" by (simp add: content_division) also have "\ \ ?\ (cbox (- vec c) (vec c) :: (real, 'm) vec set)" proof (rule measure_mono_fmeasurable) show "\\ \ cbox (- vec c) (vec c)" by (meson Sup_subset_mono sub_cc order_trans \\ \ \\) qed (use \\ division_of \\\ lmeasurable_division in auto) also have "\ = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)" by simp also have "\ \ (2 ^ ?m * c ^ ?m)" using \c > 0\ by (simp add: content_cbox_if_cart) finally have "sum ?\ \ \ (2 ^ ?m * c ^ ?m)" . then show ?thesis using \e > 0\ \c > 0\ by (auto simp: field_split_simps) qed finally show ?thesis . qed show "\T. f ` S \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "f ` S \ \ (g ` \)" using covers sub_g by force show "\ (g ` \) \ lmeasurable" by (rule fmeasurable_UN_bound [OF \countable \\ meas_g le_e]) show "?\ (\ (g ` \)) \ e" by (rule measure_UN_bound [OF \countable \\ meas_g le_e]) qed qed qed theorem baby_Sard: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" and der: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" shows "negligible(f ` S)" proof - let ?U = "\n. {x \ S. norm(x) \ n \ onorm(f' x) \ real n}" have "\x. x \ S \ \n. norm x \ real n \ onorm (f' x) \ real n" by (meson linear order_trans real_arch_simple) then have eq: "S = (\n. ?U n)" by auto have "negligible (f ` ?U n)" for n proof (rule Sard_lemma2 [OF mlen]) show "0 < real n + 1" by auto show "bounded (?U n)" using bounded_iff by blast show "(f has_derivative f' x) (at x within ?U n)" if "x \ ?U n" for x using der that by (force intro: has_derivative_subset) qed (use rank in auto) then show ?thesis by (subst eq) (simp add: image_Union negligible_Union_nat) qed subsection\A one-way version of change-of-variables not assuming injectivity. \ lemma integral_on_image_ubound_weak: fixes f :: "real^'n::{finite,wellorder} \ real" assumes S: "S \ sets lebesgue" and f: "f \ borel_measurable (lebesgue_on (g ` S))" and nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and det_int_fg: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" and meas_gim: "\T. \T \ g ` S; T \ sets lebesgue\ \ {x \ S. g x \ T} \ sets lebesgue" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. \det (matrix (g' x))\" have cont_g: "continuous_on S g" using der_g has_derivative_continuous_on by blast have [simp]: "space (lebesgue_on S) = S" by (simp add: S) have gS_in_sets_leb: "g ` S \ sets lebesgue" apply (rule differentiable_image_in_sets_lebesgue) using der_g by (auto simp: S differentiable_def differentiable_on_def) obtain h where nonneg_h: "\n x. 0 \ h n x" and h_le_f: "\n x. x \ S \ h n (g x) \ f (g x)" and h_inc: "\n x. h n x \ h (Suc n) x" and h_meas: "\n. h n \ borel_measurable lebesgue" and fin_R: "\n. finite(range (h n))" and lim: "\x. x \ g ` S \ (\n. h n x) \ f x" proof - let ?f = "\x. if x \ g ` S then f x else 0" have "?f \ borel_measurable lebesgue \ (\x. 0 \ ?f x)" by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric]) then show ?thesis apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing) apply (rename_tac h) by (rule_tac h=h in that) (auto split: if_split_asm) qed have h_lmeas: "{t. h n (g t) = y} \ S \ sets lebesgue" for y n proof - have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV" by simp then have "((h n) -`{y} \ g ` S) \ sets (lebesgue_on (g ` S))" by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel) then have "({u. h n u = y} \ g ` S) \ sets lebesgue" using gS_in_sets_leb by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def) then have "{x \ S. g x \ ({u. h n u = y} \ g ` S)} \ sets lebesgue" using meas_gim[of "({u. h n u = y} \ g ` S)"] by force moreover have "{t. h n (g t) = y} \ S = {x \ S. g x \ ({u. h n u = y} \ g ` S)}" by blast ultimately show ?thesis by auto qed have hint: "h n integrable_on g ` S \ integral (g ` S) (h n) \ integral S (\x. ?D x * h n (g x))" (is "?INT \ ?lhs \ ?rhs") for n proof - let ?R = "range (h n)" have hn_eq: "h n = (\x. \y\?R. y * indicat_real {x. h n x = y} x)" by (simp add: indicator_def if_distrib fin_R cong: if_cong) have yind: "(\t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \ (integral (g ` S) (\t. y * indicator {x. h n x = y} t)) \ integral S (\t. \det (matrix (g' t))\ * y * indicator {x. h n x = y} (g t))" if y: "y \ ?R" for y::real proof (cases "y=0") case True then show ?thesis using gS_in_sets_leb integrable_0 by force next case False with that have "y > 0" using less_eq_real_def nonneg_h by fastforce have "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. ?D x) \ borel_measurable (lebesgue_on ({t. h n (g t) = y} \ S))" apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g]) by (meson der_g IntD2 has_derivative_subset inf_le2) then have "(\x. if x \ {t. h n (g t) = y} \ S then ?D x else 0) \ borel_measurable lebesgue" by (rule borel_measurable_if_I [OF _ h_lmeas]) then show "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) \ borel_measurable (lebesgue_on S)" by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric]) show "(\x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S" by (rule integrable_cmul) (use det_int_fg in auto) show "norm (if x \ {t. h n (g t) = y} then ?D x else 0) \ ?D x *\<^sub>R f (g x) /\<^sub>R y" if "x \ S" for x using nonneg_h [of n x] \y > 0\ nonneg_fg [of x] h_le_f [of x n] that by (auto simp: divide_simps mult_left_mono) qed (use S in auto) then have int_det: "(\t. \det (matrix (g' t))\) integrable_on ({t. h n (g t) = y} \ S)" using integrable_restrict_Int by force have "(g ` ({t. h n (g t) = y} \ S)) \ lmeasurable" apply (rule measurable_differentiable_image [OF h_lmeas]) apply (blast intro: has_derivative_subset [OF der_g]) apply (rule int_det) done moreover have "g ` ({t. h n (g t) = y} \ S) = {x. h n x = y} \ g ` S" by blast moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \ S)) \ integral ({t. h n (g t) = y} \ S) (\t. \det (matrix (g' t))\)" apply (rule measure_differentiable_image [OF h_lmeas _ int_det]) apply (blast intro: has_derivative_subset [OF der_g]) done ultimately show ?thesis using \y > 0\ integral_restrict_Int [of S "{t. h n (g t) = y}" "\t. \det (matrix (g' t))\ * y"] apply (simp add: integrable_on_indicator integral_indicator) apply (simp add: indicator_def if_distrib cong: if_cong) done qed have hn_int: "h n integrable_on g ` S" apply (subst hn_eq) using yind by (force intro: integrable_sum [OF fin_R]) then show ?thesis proof have "?lhs = integral (g ` S) (\x. \y\range (h n). y * indicat_real {x. h n x = y} x)" by (metis hn_eq) also have "\ = (\y\range (h n). integral (g ` S) (\x. y * indicat_real {x. h n x = y} x))" by (rule integral_sum [OF fin_R]) (use yind in blast) also have "\ \ (\y\range (h n). integral S (\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)))" using yind by (force intro: sum_mono) also have "\ = integral S (\u. \y\range (h n). \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u))" proof (rule integral_sum [OF fin_R, symmetric]) fix y assume y: "y \ ?R" with nonneg_h have "y \ 0" by auto show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. indicat_real {x. h n x = y} (g x)) \ borel_measurable (lebesgue_on S)" using h_lmeas S by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff) then show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) \ borel_measurable (lebesgue_on S)" by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g]) next fix x assume "x \ S" have "y * indicat_real {x. h n x = y} (g x) \ f (g x)" by (metis (full_types) \x \ S\ h_le_f indicator_def mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg) with \y \ 0\ show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \ ?D x * f(g x)" by (simp add: abs_mult mult.assoc mult_left_mono) qed (use S det_int_fg in auto) qed also have "\ = integral S (\T. \det (matrix (g' T))\ * (\y\range (h n). y * indicat_real {x. h n x = y} (g T)))" by (simp add: sum_distrib_left mult.assoc) also have "\ = ?rhs" by (metis hn_eq) finally show "integral (g ` S) (h n) \ ?rhs" . qed qed have le: "integral S (\T. \det (matrix (g' T))\ * h n (g T)) \ ?b" for n proof (rule integral_le) show "(\T. \det (matrix (g' T))\ * h n (g T)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\T. \det (matrix (g' T))\ *\<^sub>R h n (g T)) \ borel_measurable (lebesgue_on S)" proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \S \ sets lebesgue\) have eq: "{x \ S. f x \ a} = (\b \ (f ` S) \ atMost a. {x. f x = b} \ S)" for f and a::real by auto have "finite ((\x. h n (g x)) ` S \ {..a})" for a by (force intro: finite_subset [OF _ fin_R]) with h_lmeas [of n] show "(\x. h n (g x)) \ borel_measurable (lebesgue_on S)" apply (simp add: borel_measurable_vimage_halfspace_component_le \S \ sets lebesgue\ sets_restrict_space_iff eq) by (metis (mono_tags) SUP_inf sets.finite_UN) qed (use der_g in blast) then show "(\T. \det (matrix (g' T))\ * h n (g T)) \ borel_measurable (lebesgue_on S)" by simp show "norm (?D x * h n (g x)) \ ?D x *\<^sub>R f (g x)" if "x \ S" for x by (simp add: h_le_f mult_left_mono nonneg_h that) qed (use S det_int_fg in auto) show "?D x * h n (g x) \ ?D x * f (g x)" if "x \ S" for x by (simp add: \x \ S\ h_le_f mult_left_mono) show "(\x. ?D x * f (g x)) integrable_on S" using det_int_fg by blast qed have "f integrable_on g ` S \ (\k. integral (g ` S) (h k)) \ integral (g ` S) f" proof (rule monotone_convergence_increasing) have "\integral (g ` S) (h n)\ \ integral S (\x. ?D x * f (g x))" for n proof - have "\integral (g ` S) (h n)\ = integral (g ` S) (h n)" using hint by (simp add: integral_nonneg nonneg_h) also have "\ \ integral S (\x. ?D x * f (g x))" using hint le by (meson order_trans) finally show ?thesis . qed then show "bounded (range (\k. integral (g ` S) (h k)))" by (force simp: bounded_iff) qed (use h_inc lim hint in auto) moreover have "integral (g ` S) (h n) \ integral S (\x. ?D x * f (g x))" for n using hint by (blast intro: le order_trans) ultimately show ?thesis by (auto intro: Lim_bounded) qed lemma integral_on_image_ubound_nonneg: fixes f :: "real^'n::{finite,wellorder} \ real" assumes nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. det (matrix (g' x))" define S' where "S' \ {x \ S. ?D x * f(g x) \ 0}" then have der_gS': "\x. x \ S' \ (g has_derivative g' x) (at x within S')" by (metis (mono_tags, lifting) der_g has_derivative_subset mem_Collect_eq subset_iff) have "(\x. if x \ S then \?D x\ * f (g x) else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV intS) then have Df_borel: "(\x. if x \ S then \?D x\ * f (g x) else 0) \ borel_measurable lebesgue" using integrable_imp_measurable lebesgue_on_UNIV_eq by force have S': "S' \ sets lebesgue" proof - from Df_borel borel_measurable_vimage_open [of _ UNIV] have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ T} \ sets lebesgue" if "open T" for T using that unfolding lebesgue_on_UNIV_eq by (fastforce simp add: dest!: spec) then have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ -{0}} \ sets lebesgue" using open_Compl by blast then show ?thesis by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong) qed then have gS': "g ` S' \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "g differentiable_on S'" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_subset) qed auto have f: "f \ borel_measurable (lebesgue_on (g ` S'))" proof (clarsimp simp add: borel_measurable_vimage_open) fix T :: "real set" assume "open T" have "{x \ g ` S'. f x \ T} = g ` {x \ S'. f(g x) \ T}" by blast moreover have "g ` {x \ S'. f(g x) \ T} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) let ?h = "\x. \?D x\ * f (g x) /\<^sub>R \?D x\" have "(\x. if x \ S' then \?D x\ * f (g x) else 0) = (\x. if x \ S then \?D x\ * f (g x) else 0)" by (auto simp: S'_def) also have "\ \ borel_measurable lebesgue" by (rule Df_borel) finally have *: "(\x. \?D x\ * f (g x)) \ borel_measurable (lebesgue_on S')" by (simp add: borel_measurable_if_D) have "?h \ borel_measurable (lebesgue_on S')" by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS') moreover have "?h x = f(g x)" if "x \ S'" for x using that by (auto simp: S'_def) ultimately have "(\x. f(g x)) \ borel_measurable (lebesgue_on S')" by (metis (no_types, lifting) measurable_lebesgue_cong) then show "{x \ S'. f (g x) \ T} \ sets lebesgue" by (simp add: \S' \ sets lebesgue\ \open T\ borel_measurable_vimage_open sets_restrict_space_iff) show "g differentiable_on {x \ S'. f (g x) \ T}" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_subset) qed auto ultimately have "{x \ g ` S'. f x \ T} \ sets lebesgue" by metis then show "{x \ g ` S'. f x \ T} \ sets (lebesgue_on (g ` S'))" by (simp add: \g ` S' \ sets lebesgue\ sets_restrict_space_iff) qed have intS': "(\x. \?D x\ * f (g x)) integrable_on S'" using intS by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible) have lebS': "{x \ S'. g x \ T} \ sets lebesgue" if "T \ g ` S'" "T \ sets lebesgue" for T proof - have "g \ borel_measurable (lebesgue_on S')" using der_gS' has_derivative_continuous_on S' by (blast intro: continuous_imp_measurable_on_sets_lebesgue) moreover have "{x \ S'. g x \ U} \ sets lebesgue" if "negligible U" "U \ g ` S'" for U proof (intro negligible_imp_sets negligible_differentiable_vimage that) fix x assume x: "x \ S'" then have "linear (g' x)" using der_gS' has_derivative_linear by blast with x show "inj (g' x)" by (auto simp: S'_def det_nz_iff_inj) qed (use der_gS' in auto) ultimately show ?thesis using double_lebesgue_sets [OF S' gS' order_refl] that by blast qed have int_gS': "f integrable_on g ` S' \ integral (g ` S') f \ integral S' (\x. \?D x\ * f(g x))" using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast have "negligible (g ` {x \ S. det(matrix(g' x)) = 0})" proof (rule baby_Sard, simp_all) fix x assume x: "x \ S \ det (matrix (g' x)) = 0" then show "(g has_derivative g' x) (at x within {x \ S. det (matrix (g' x)) = 0})" by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI) then show "rank (matrix (g' x)) < CARD('n)" using det_nz_iff_inj matrix_vector_mul_linear x by (fastforce simp add: less_rank_noninjective) qed then have negg: "negligible (g ` S - g ` {x \ S. ?D x \ 0})" by (rule negligible_subset) (auto simp: S'_def) have null: "g ` {x \ S. ?D x \ 0} - g ` S = {}" by (auto simp: S'_def) let ?F = "{x \ S. f (g x) \ 0}" have eq: "g ` S' = g ` ?F \ g ` {x \ S. ?D x \ 0}" by (auto simp: S'_def image_iff) show ?thesis proof have "((\x. if x \ g ` ?F then f x else 0) integrable_on g ` {x \ S. ?D x \ 0})" using int_gS' eq integrable_restrict_Int [where f=f] by simp then have "f integrable_on g ` {x \ S. ?D x \ 0}" by (auto simp: image_iff elim!: integrable_eq) then show "f integrable_on g ` S" apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) using negg null by auto have "integral (g ` S) f = integral (g ` {x \ S. ?D x \ 0}) f" using negg by (auto intro: negligible_subset integral_spike_set) also have "\ = integral (g ` {x \ S. ?D x \ 0}) (\x. if x \ g ` ?F then f x else 0)" by (auto simp: image_iff intro!: integral_cong) also have "\ = integral (g ` S') f" using eq integral_restrict_Int by simp also have "\ \ integral S' (\x. \?D x\ * f(g x))" by (metis int_gS') also have "\ \ ?b" by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto) finally show "integral (g ` S) f \ ?b" . qed qed lemma absolutely_integrable_on_image_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" proof - let ?D = "\x. \det (matrix (g' x))\ * f (g x)" let ?N = "{x \ S. f (g x) < 0}" and ?P = "{x \ S. f (g x) > 0}" have eq: "{x. (if x \ S then ?D x else 0) > 0} = {x \ S. ?D x > 0}" "{x. (if x \ S then ?D x else 0) < 0} = {x \ S. ?D x < 0}" by auto have "?D integrable_on S" using intS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?D x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?D x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have Dlt: "{x \ S. ?D x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have Dgt: "{x \ S. ?D x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have dfgbm: "?D \ borel_measurable (lebesgue_on S)" using intS absolutely_integrable_on_def integrable_imp_measurable by blast have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \ ?N" for x using der_g has_derivative_subset that by force have "(\x. - f x) integrable_on g ` ?N \ integral (g ` ?N) (\x. - f x) \ integral ?N (\x. \det (matrix (g' x))\ * - f (g x))" proof (rule integral_on_image_ubound_nonneg [OF _ der_gN]) have 1: "?D integrable_on {x \ S. ?D x < 0}" using Dlt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) have "uminus \ (\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1]) then show "(\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: integrable_neg_iff o_def) qed auto then have "f integrable_on g ` ?N" by (simp add: integrable_neg_iff) moreover have "g ` ?N = {y \ g ` S. f y < 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y < 0}" by simp then have N: "f absolutely_integrable_on {y \ g ` S. f y < 0}" by (rule absolutely_integrable_absolutely_integrable_ubound) auto have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \ ?P" for x using der_g has_derivative_subset that by force have "f integrable_on g ` ?P \ integral (g ` ?P) f \ integral ?P ?D" proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) have "?D integrable_on {x \ S. 0 < ?D x}" using Dgt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) then show "?D integrable_on ?P" apply (rule integrable_spike_set) by (auto simp: zero_less_mult_iff empty_imp_negligible) qed auto then have "f integrable_on g ` ?P" by metis moreover have "g ` ?P = {y \ g ` S. f y > 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y > 0}" by simp then have P: "f absolutely_integrable_on {y \ g ` S. f y > 0}" by (rule absolutely_integrable_absolutely_integrable_lbound) auto have "(\x. if x \ g ` S \ f x < 0 \ x \ g ` S \ 0 < f x then f x else 0) = (\x. if x \ g ` S then f x else 0)" by auto then show ?thesis using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f] by simp qed proposition absolutely_integrable_on_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]]) using absolutely_integrable_component [OF intS] by auto proposition integral_on_image_ubound: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes"\x. x \ S \ 0 \ f(g x)" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" using integral_on_image_ubound_nonneg [OF assms] by simp subsection\Change-of-variables theorem\ text\The classic change-of-variables theorem. We have two versions with quite general hypotheses, the first that the transforming function has a continuous inverse, the second that the base set is Lebesgue measurable.\ lemma cov_invertible_nonneg_le: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and f0: "\y. y \ T \ 0 \ f y" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "f integrable_on T \ (integral T f) \ b \ (\x. \det (matrix (g' x))\ * f(g x)) integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) \ b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ have gS: "g differentiable_on S" by (meson der_g differentiable_def differentiable_on_def) let ?D = "\x. \det (matrix (g' x))\ * f (g x)" show ?thesis proof assume ?lhs then have fT: "f integrable_on T" and intf: "integral T f \ b" by blast+ show ?rhs proof let ?fgh = "\x. \det (matrix (h' x))\ * (\det (matrix (g' (h x)))\ * f (g (h x)))" have ddf: "?fgh x = f x" if "x \ T" for x proof - have "matrix (h' x) ** matrix (g' (h x)) = mat 1" using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ then have "\det (matrix (h' x))\ * \det (matrix (g' (h x)))\ = 1" by (metis abs_1 abs_mult det_I det_mul) then show ?thesis by (simp add: gh that) qed have "?D integrable_on (h ` T)" proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real) show "(\x. ?fgh x) absolutely_integrable_on T" proof (subst absolutely_integrable_on_iff_nonneg) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (simp add: zero_le_mult_iff f0 gh) qed (use der_h in auto) with Seq show "(\x. ?D x) integrable_on S" by simp have "integral S (\x. ?D x) \ integral T (\x. ?fgh x)" unfolding Seq proof (rule integral_on_image_ubound) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (use f0 gh der_h in auto) also have "\ = integral T f" by (force simp: ddf intro: integral_cong) also have "\ \ b" by (rule intf) finally show "integral S (\x. ?D x) \ b" . qed next assume R: ?rhs then have "f integrable_on g ` S" using der_g f0 hg integral_on_image_ubound_nonneg by blast moreover have "integral (g ` S) f \ integral S (\x. ?D x)" by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto) ultimately show ?lhs using R by (simp add: Teq) qed qed lemma cov_invertible_nonneg_eq: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "\y. y \ T \ (h has_derivative h' y) (at y within T)" and "\y. y \ T \ 0 \ f y" and "\x. x \ S \ g x \ T \ h(g x) = x" and "\y. y \ T \ h y \ S \ g(h y) = y" and "\y. y \ T \ h' y \ g'(h y) = id" shows "((\x. \det (matrix (g' x))\ * f(g x)) has_integral b) S \ (f has_integral b) T" using cov_invertible_nonneg_le [OF assms] by (simp add: has_integral_iff) (meson eq_iff) lemma cov_invertible_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ let ?DP = "\x. \det (matrix (g' x))\ * f(g x)" and ?DN = "\x. \det (matrix (g' x))\ * -f(g x)" have "+": "(?DP has_integral b) {x \ S. f (g x) > 0} \ (f has_integral b) {y \ T. f y > 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. f (g x)) -` Y \ {x \ S. f (g x) > 0} = ((\x. f (g x)) -` Y \ S) \ {x \ S. f (g x) > 0}" for Y by auto show "(g has_derivative g' x) (at x within {x \ S. f (g x) > 0})" if "x \ {x \ S. f (g x) > 0}" for x using that der_g has_derivative_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y > 0})" if "y \ {y \ T. f y > 0}" for y using that der_h has_derivative_subset by fastforce qed (use gh hg id in auto) have "-": "(?DN has_integral b) {x \ S. f (g x) < 0} \ ((\x. - f x) has_integral b) {y \ T. f y < 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. - f (g x)) -` y \ {x \ S. f (g x) < 0} = ((\x. f (g x)) -` uminus ` y \ S) \ {x \ S. f (g x) < 0}" for y using image_iff by fastforce show "(g has_derivative g' x) (at x within {x \ S. f (g x) < 0})" if "x \ {x \ S. f (g x) < 0}" for x using that der_g has_derivative_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y < 0})" if "y \ {y \ T. f y < 0}" for y using that der_h has_derivative_subset by fastforce qed (use gh hg id in auto) show ?thesis proof assume LHS: ?lhs have eq: "{x. (if x \ S then ?DP x else 0) > 0} = {x \ S. ?DP x > 0}" "{x. (if x \ S then ?DP x else 0) < 0} = {x \ S. ?DP x < 0}" by auto have "?DP integrable_on S" using LHS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?DP x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?DP x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have SN: "{x \ S. ?DP x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have SP: "{x \ S. ?DP x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have "?DP absolutely_integrable_on {x \ S. ?DP x > 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP) then have aP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible) have "?DP absolutely_integrable_on {x \ S. ?DP x < 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN) then have aN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible) have fN: "f integrable_on {y \ T. f y < 0}" "integral {y \ T. f y < 0} f = integral {x \ S. f (g x) < 0} ?DP" using "-" [of "integral {x \ S. f(g x) < 0} ?DN"] aN by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faN: "f absolutely_integrable_on {y \ T. f y < 0}" apply (rule absolutely_integrable_integrable_bound [where g = "\x. - f x"]) using fN by (auto simp: integrable_neg_iff) have fP: "f integrable_on {y \ T. f y > 0}" "integral {y \ T. f y > 0} f = integral {x \ S. f (g x) > 0} ?DP" using "+" [of "integral {x \ S. f(g x) > 0} ?DP"] aP by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faP: "f absolutely_integrable_on {y \ T. f y > 0}" apply (rule absolutely_integrable_integrable_bound [where g = f]) using fP by auto have fa: "f absolutely_integrable_on ({y \ T. f y < 0} \ {y \ T. f y > 0})" by (rule absolutely_integrable_Un [OF faN faP]) show ?rhs proof have eq: "((if x \ T \ f x < 0 \ x \ T \ 0 < f x then 1 else 0) * f x) = (if x \ T then 1 else 0) * f x" for x by auto show "f absolutely_integrable_on T" using fa by (simp add: indicator_def set_integrable_def eq) have [simp]: "{y \ T. f y < 0} \ {y \ T. 0 < f y} = {}" for T and f :: "(real^'n::_) \ real" by auto have "integral T f = integral ({y \ T. f y < 0} \ {y \ T. f y > 0}) f" by (intro empty_imp_negligible integral_spike_set) (auto simp: eq) also have "\ = integral {y \ T. f y < 0} f + integral {y \ T. f y > 0} f" using fN fP by simp also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" by (simp add: fN fP) also have "\ = integral ({x \ S. f (g x) < 0} \ {x \ S. 0 < f (g x)}) ?DP" using aP aN by (simp add: set_lebesgue_integral_eq_integral) also have "\ = integral S ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using LHS by simp finally show "integral T f = b" . qed next assume RHS: ?rhs have eq: "{x. (if x \ T then f x else 0) > 0} = {x \ T. f x > 0}" "{x. (if x \ T then f x else 0) < 0} = {x \ T. f x < 0}" by auto have "f integrable_on T" using RHS absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have TN: "{x \ T. f x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have TP: "{x \ T. f x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have aint: "f absolutely_integrable_on {y. y \ T \ 0 < (f y)}" "f absolutely_integrable_on {y. y \ T \ (f y) < 0}" and intT: "integral T f = b" using set_integrable_subset [of _ T] TP TN RHS by blast+ show ?lhs proof have fN: "f integrable_on {v \ T. f v < 0}" using absolutely_integrable_on_def aint by blast then have DN: "(?DN has_integral integral {y \ T. f y < 0} (\x. - f x)) {x \ S. f (g x) < 0}" using "-" [of "integral {y \ T. f y < 0} (\x. - f x)"] by (simp add: has_integral_neg_iff integrable_integral) have aDN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DN]) using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+ have fP: "f integrable_on {v \ T. f v > 0}" using absolutely_integrable_on_def aint by blast then have DP: "(?DP has_integral integral {y \ T. f y > 0} f) {x \ S. f (g x) > 0}" using "+" [of "integral {y \ T. f y > 0} f"] by (simp add: has_integral_neg_iff integrable_integral) have aDP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DP]) using DP hg by (fastforce simp: integrable_neg_iff)+ have eq: "(if x \ S then 1 else 0) * ?DP x = (if x \ S \ f (g x) < 0 \ x \ S \ f (g x) > 0 then 1 else 0) * ?DP x" for x by force have "?DP absolutely_integrable_on ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0})" by (rule absolutely_integrable_Un [OF aDN aDP]) then show I: "?DP absolutely_integrable_on S" by (simp add: indicator_def eq set_integrable_def) have [simp]: "{y \ S. f y < 0} \ {y \ S. 0 < f y} = {}" for S and f :: "(real^'n::_) \ real" by auto have "integral S ?DP = integral ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0}) ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" using aDN aDP by (simp add: set_lebesgue_integral_eq_integral) also have "\ = - integral {y \ T. f y < 0} (\x. - f x) + integral {y \ T. f y > 0} f" using DN DP by (auto simp: has_integral_iff) also have "\ = integral ({x \ T. f x < 0} \ {x \ T. 0 < f x}) f" by (simp add: fN fP) also have "\ = integral T f" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using intT by simp finally show "integral S ?DP = b" . qed qed qed lemma cv_inv_version3: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have "((\x. \det (matrix (g' x))\ * f(g x) $ i) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * (f(g x) $ i)) = b $ i) \ ((\x. f x $ i) absolutely_integrable_on T \ integral T (\x. f x $ i) = b $ i)" for i by (rule cov_invertible_real [OF der_g der_h hg gh id]) then have "?D absolutely_integrable_on S \ (?D has_integral b) S \ f absolutely_integrable_on T \ (f has_integral b) T" unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f] absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D] by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric] has_integral_iff set_lebesgue_integral_eq_integral) then show ?thesis using absolutely_integrable_on_def by blast qed lemma cv_inv_version4: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S) \ invertible(matrix(g' x))" and hg: "\x. x \ S \ continuous_on (g ` S) h \ h(g x) = x" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "\x. \h'. x \ S \ (g has_derivative g' x) (at x within S) \ linear h' \ g' x \ h' = id \ h' \ g' x = id" using der_g matrix_invertible has_derivative_linear by blast then obtain h' where h': "\x. x \ S \ (g has_derivative g' x) (at x within S) \ linear (h' x) \ g' x \ (h' x) = id \ (h' x) \ g' x = id" by metis show ?thesis proof (rule cv_inv_version3) show "\y. y \ g ` S \ (h has_derivative h' (h y)) (at y within g ` S)" using h' hg by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within) qed (use h' hg in auto) qed theorem has_absolute_integral_change_of_variables_invertible: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and hg: "\x. x \ S \ h(g x) = x" and conth: "continuous_on (g ` S) h" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" (is "?lhs = ?rhs") proof - let ?S = "{x \ S. invertible (matrix (g' x))}" and ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have *: "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b" proof (rule cv_inv_version4) show "(g has_derivative g' x) (at x within ?S) \ invertible (matrix (g' x))" if "x \ ?S" for x using der_g that has_derivative_subset that by fastforce show "continuous_on (g ` ?S) h \ h (g x) = x" if "x \ ?S" for x using that continuous_on_subset [OF conth] by (simp add: hg image_mono) qed have "(g has_derivative g' x) (at x within {x \ S. rank (matrix (g' x)) < CARD('m)})" if "x \ S" for x by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI that) then have "negligible (g ` {x \ S. \ invertible (matrix (g' x))})" by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard) then have neg: "negligible {x \ g ` S. x \ g ` ?S \ f x \ 0}" by (auto intro: negligible_subset) have [simp]: "{x \ g ` ?S. x \ g ` S \ f x \ 0} = {}" by auto have "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" apply (intro conj_cong absolutely_integrable_spike_set_eq) apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg) done moreover have "f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg) ultimately show ?thesis using "*" by blast qed theorem has_absolute_integral_change_of_variables_compact: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "compact S" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b)" proof - obtain h where hg: "\x. x \ S \ h(g x) = x" using inj by (metis the_inv_into_f_f) have conth: "continuous_on (g ` S) h" by (metis \compact S\ continuous_on_inv der_g has_derivative_continuous_on hg) show ?thesis by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth]) qed lemma has_absolute_integral_change_of_variables_compact_family: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes compact: "\n::nat. compact (F n)" and der_g: "\x. x \ (\n. F n) \ (g has_derivative g' x) (at x within (\n. F n))" and inj: "inj_on g (\n. F n)" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on (\n. F n) \ integral (\n. F n) (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` (\n. F n)) \ integral (g ` (\n. F n)) f = b)" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" let ?U = "\n. \m\n. F m" let ?lift = "vec::real\real^1" have F_leb: "F m \ sets lebesgue" for m by (simp add: compact borel_compact) have iff: "(\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \ integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) = b \ f absolutely_integrable_on (g ` (?U n)) \ integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \ real^'k" proof (rule has_absolute_integral_change_of_variables_compact) show "compact (?U n)" by (simp add: compact compact_UN) show "(g has_derivative g' x) (at x within (?U n))" if "x \ ?U n" for x using that by (blast intro!: has_derivative_subset [OF der_g]) show "inj_on g (?U n)" using inj by (auto simp: inj_on_def) qed show ?thesis unfolding image_UN proof safe assume DS: "?D absolutely_integrable_on (\n. F n)" and b: "b = integral (\n. F n) ?D" have DU: "\n. ?D absolutely_integrable_on (?U n)" "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" using integral_countable_UN [OF DS F_leb] by auto with iff have fag: "f absolutely_integrable_on g ` (?U n)" and fg_int: "integral (\m\n. g ` F m) f = integral (?U n) ?D" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\m. g ` F m) then norm(f x) else 0" have "(\x. if x \ (\m. g ` F m) then f x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ (\m\k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using fag by (simp add: image_UN) let ?nf = "\n x. if x \ (\m\n. g ` F m) then norm(f x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nf k integrable_on UNIV" for k using fag unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n have "(norm \ ?D) absolutely_integrable_on ?U n" by (intro absolutely_integrable_norm DU) then have "integral (g ` ?U n) (norm \ f) = integral (?U n) (norm \ ?D)" using iff [of n "vec \ norm \ f" "integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R (?lift \ norm \ f) (g x))"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def) } moreover have "bounded (range (\k. integral (?U k) (norm \ ?D)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (?U k) (norm \ ?D)) \ integral (\n. F n) (norm \ ?D)" unfolding integral_restrict_UNIV [of _ "norm \ ?D", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ \ (F ` {..k}) then (norm \ ?D) x else 0) integrable_on UNIV" "(\x. if x \ (\n. F n) then (norm \ ?D) x else 0) integrable_on UNIV" using DU(1) DS unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nf k)))" by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def) next show "(\k. if x \ (\m\k. g ` F m) then norm (f x) else 0) \ (if x \ (\m. g ` F m) then norm (f x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ (\m\k. g ` F m) then f x else 0) \ (if x \ (\m. g ` F m) then f x else 0)" for x proof clarsimp fix m y assume "y \ F m" show "(\k. if \x\{..k}. g y \ g ` F x then f (g y) else 0) \ f (g y)" using \y \ F m\ by (force intro: tendsto_eventually eventually_sequentiallyI [of m]) qed qed auto then show fai: "f absolutely_integrable_on (\m. g ` F m)" using absolutely_integrable_restrict_UNIV by blast show "integral ((\x. g ` F x)) f = integral (\n. F n) ?D" proof (rule LIMSEQ_unique) show "(\n. integral (?U n) ?D) \ integral (\x. g ` F x) f" unfolding fg_int [symmetric] proof (rule integral_countable_UN [OF fai]) show "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI) qed auto qed next show "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" by (rule DU) qed next assume fs: "f absolutely_integrable_on (\x. g ` F x)" and b: "b = integral ((\x. g ` F x)) f" have gF_leb: "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" using der_g unfolding differentiable_def differentiable_on_def by (meson Sup_upper UNIV_I UnionI has_derivative_subset image_eqI) qed auto have fgU: "\n. f absolutely_integrable_on (\m\n. g ` F m)" "(\n. integral (\m\n. g ` F m) f) \ integral (\m. g ` F m) f" using integral_countable_UN [OF fs gF_leb] by auto with iff have DUn: "?D absolutely_integrable_on ?U n" and D_int: "integral (?U n) ?D = integral (\m\n. g ` F m) f" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\n. F n) then norm(?D x) else 0" have "(\x. if x \ (\n. F n) then ?D x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using DUn by simp let ?nD = "\n x. if x \ ?U n then norm(?D x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nD k integrable_on UNIV" for k using DUn unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n::nat have "(norm \ f) absolutely_integrable_on (\m\n. g ` F m)" apply (rule absolutely_integrable_norm) using fgU by blast then have "integral (?U n) (norm \ ?D) = integral (g ` ?U n) (norm \ f)" using iff [of n "?lift \ norm \ f" "integral (g ` ?U n) (?lift \ norm \ f)"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def) } moreover have "bounded (range (\k. integral (g ` ?U k) (norm \ f)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (g ` ?U k) (norm \ f)) \ integral (g ` (\n. F n)) (norm \ f)" unfolding integral_restrict_UNIV [of _ "norm \ f", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ g ` ?U k then (norm \ f) x else 0) integrable_on UNIV" "(\x. if x \ g ` (\n. F n) then (norm \ f) x else 0) integrable_on UNIV" using fgU fs unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by (auto simp: image_UN) qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nD k)))" unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp next show "(\k. if x \ ?U k then norm (?D x) else 0) \ (if x \ (\n. F n) then norm (?D x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ ?U k then ?D x else 0) \ (if x \ (\n. F n) then ?D x else 0)" for x proof clarsimp fix n assume "x \ F n" show "(\m. if \j\{..m}. x \ F j then ?D x else 0) \ ?D x" using \x \ F n\ by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n]) qed qed auto then show Dai: "?D absolutely_integrable_on (\n. F n)" unfolding absolutely_integrable_restrict_UNIV by simp show "integral (\n. F n) ?D = integral ((\x. g ` F x)) f" proof (rule LIMSEQ_unique) show "(\n. integral (\m\n. g ` F m) f) \ integral (\x. g ` F x) f" by (rule fgU) show "(\n. integral (\m\n. g ` F m) f) \ integral (\n. F n) ?D" unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb]) qed qed qed theorem has_absolute_integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - obtain C N where "fsigma C" and N: "N \ null_sets lebesgue" and CNS: "C \ N = S" and "disjnt C N" using lebesgue_set_almost_fsigma [OF S] . then obtain F :: "nat \ (real^'m::_) set" where F: "range F \ Collect compact" and Ceq: "C = Union(range F)" using fsigma_Union_compact by metis have "negligible N" using N by (simp add: negligible_iff_null_sets) let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" have "?D absolutely_integrable_on C \ integral C ?D = b \ f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b" unfolding Ceq proof (rule has_absolute_integral_change_of_variables_compact_family) fix n x assume "x \ \(F ` UNIV)" then show "(g has_derivative g' x) (at x within \(F ` UNIV))" using Ceq \C \ N = S\ der_g has_derivative_subset by blast next have "\(F ` UNIV) \ S" using Ceq \C \ N = S\ by blast then show "inj_on g (\(F ` UNIV))" using inj by (meson inj_on_subset) qed (use F in auto) moreover have "?D absolutely_integrable_on C \ integral C ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" proof (rule conj_cong) have neg: "negligible {x \ C - S. ?D x \ 0}" "negligible {x \ S - C. ?D x \ 0}" using CNS by (blast intro: negligible_subset [OF \negligible N\])+ then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)" by (rule absolutely_integrable_spike_set_eq) show "(integral C ?D = b) \ (integral S ?D = b)" using integral_spike_set [OF neg] by simp qed moreover have "f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (rule conj_cong) have "g differentiable_on N" by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2) with \negligible N\ have neg_gN: "negligible (g ` N)" by (blast intro: negligible_differentiable_image_negligible) have neg: "negligible {x \ g ` C - g ` S. f x \ 0}" "negligible {x \ g ` S - g ` C. f x \ 0}" using CNS by (blast intro: negligible_subset [OF neg_gN])+ then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)" by (rule absolutely_integrable_spike_set_eq) show "(integral (g ` C) f = b) \ (integral (g ` S) f = b)" using integral_spike_set [OF neg] by simp qed ultimately show ?thesis by simp qed corollary absolutely_integrable_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "inj_on g S" shows "f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" using assms has_absolute_integral_change_of_variables by blast corollary integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" and disj: "(f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S)" shows "integral (g ` S) f = integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x))" using has_absolute_integral_change_of_variables [OF S der_g inj] disj by blast lemma has_absolute_integral_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - let ?lift = "vec :: real \ real^1" let ?drop = "(\x::real^1. x $ 1)" have S': "?lift ` S \ sets lebesgue" by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec) have "((\x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)" if "z \ S" for z using der_g [OF that] by (simp add: has_vector_derivative_def has_derivative_vector_1) then have der': "\x. x \ ?lift ` S \ (?lift \ g \ ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)" by (auto simp: o_def) have inj': "inj_on (vec \ g \ ?drop) (vec ` S)" using inj by (simp add: inj_on_def) let ?fg = "\x. \g' x\ *\<^sub>R f(g x)" have "((\x. ?fg x $ i) absolutely_integrable_on S \ ((\x. ?fg x $ i) has_integral b $ i) S \ (\x. f x $ i) absolutely_integrable_on g ` S \ ((\x. f x $ i) has_integral b $ i) (g ` S))" for i using has_absolute_integral_change_of_variables [OF S' der' inj', of "\x. ?lift(f (?drop x) $ i)" "?lift (b$i)"] unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff) then have "?fg absolutely_integrable_on S \ (?fg has_integral b) S \ f absolutely_integrable_on (g ` S) \ (f has_integral b) (g ` S)" unfolding has_integral_componentwise_iff [where y=b] absolutely_integrable_componentwise_iff [where f=f] absolutely_integrable_componentwise_iff [where f = ?fg] by (force simp: Basis_vec_def cart_eq_inner_axis) then show ?thesis using absolutely_integrable_on_def by blast qed corollary absolutely_integrable_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(f absolutely_integrable_on g ` S \ (\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S)" using has_absolute_integral_change_of_variables_1 [OF assms] by auto subsection\Change of variables for integrals: special case of linear function\ lemma has_absolute_integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix g)\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (cases "det(matrix g) = 0") case True then have "negligible(g ` S)" using assms det_nz_iff_inj negligible_linear_singular_image by blast with True show ?thesis by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible) next case False then obtain h where h: "\x. x \ S \ h (g x) = x" "linear h" using assms det_nz_iff_inj linear_injective_isomorphism by metis show ?thesis proof (rule has_absolute_integral_change_of_variables_invertible) show "(g has_derivative g) (at x within S)" for x by (simp add: assms linear_imp_has_derivative) show "continuous_on (g ` S) h" using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast qed (use h in auto) qed lemma absolutely_integrable_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ f absolutely_integrable_on (g ` S)" using assms has_absolute_integral_change_of_variables_linear by blast lemma absolutely_integrable_on_linear_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S \ det(matrix g) = 0" unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff by (auto simp: set_integrable_def) lemma integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" and "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S" shows "integral (g ` S) f = \det (matrix g)\ *\<^sub>R integral S (f \ g)" proof - have "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S) \ (f absolutely_integrable_on g ` S)" using absolutely_integrable_on_linear_image assms by blast moreover have ?thesis if "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)" using has_absolute_integral_change_of_variables_linear [OF \linear g\] that by (auto simp: o_def) ultimately show ?thesis using absolutely_integrable_change_of_variables_linear [OF \linear g\] by blast qed subsection\Change of variable for measure\ lemma has_measure_differentiable_image: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) = m \ ((\x. \det (matrix (f' x))\) has_integral m) S" using has_absolute_integral_change_of_variables [OF assms, of "\x. (1::real^1)" "vec m"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) lemma measurable_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) integrable_on S" using has_measure_differentiable_image [OF assms] by blast lemma measurable_differentiable_image_alt: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) absolutely_integrable_on S" using measurable_differentiable_image_eq [OF assms] by (simp only: absolutely_integrable_on_iff_nonneg) lemma measure_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and der_f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and inj: "inj_on f S" and intS: "(\x. \det (matrix (f' x))\) integrable_on S" shows "measure lebesgue (f ` S) = integral S (\x. \det (matrix (f' x))\)" using measurable_differentiable_image_eq [OF S der_f inj] assms has_measure_differentiable_image by blast end diff --git a/src/HOL/Analysis/Derivative.thy b/src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy +++ b/src/HOL/Analysis/Derivative.thy @@ -1,3475 +1,3475 @@ (* Title: HOL/Analysis/Derivative.thy Author: John Harrison Author: Robert Himmelmann, TU Muenchen (translation from HOL Light); tidied by LCP *) section \Derivative\ theory Derivative imports Bounded_Linear_Function Line_Segment Convex_Euclidean_Space begin declare bounded_linear_inner_left [intro] declare has_derivative_bounded_linear[dest] subsection \Derivatives\ lemma has_derivative_add_const: "(f has_derivative f') net \ ((\x. f x + c) has_derivative f') net" by (intro derivative_eq_intros) auto subsection\<^marker>\tag unimportant\ \Derivative with composed bilinear function\ text \More explicit epsilon-delta forms.\ proposition has_derivative_within': "(f has_derivative f')(at x within s) \ bounded_linear f' \ (\e>0. \d>0. \x'\s. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" unfolding has_derivative_within Lim_within dist_norm by (simp add: diff_diff_eq) lemma has_derivative_at': "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \x'. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" using has_derivative_within' [of f f' x UNIV] by simp lemma has_derivative_componentwise_within: "(f has_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_derivative (\x. f' x \ i)) (at a within S))" apply (simp add: has_derivative_within) apply (subst tendsto_componentwise_iff) apply (simp add: bounded_linear_componentwise_iff [symmetric] ball_conj_distrib) apply (simp add: algebra_simps) done lemma has_derivative_at_withinI: "(f has_derivative f') (at x) \ (f has_derivative f') (at x within s)" unfolding has_derivative_within' has_derivative_at' by blast lemma has_derivative_right: fixes f :: "real \ real" and y :: "real" shows "(f has_derivative ((*) y)) (at x within ({x <..} \ I)) \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x <..} \ I))" proof - have "((\t. (f t - (f x + y * (t - x))) / \t - x\) \ 0) (at x within ({x<..} \ I)) \ ((\t. (f t - f x) / (t - x) - y) \ 0) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (auto simp add: diff_divide_distrib add_divide_distrib) also have "\ \ ((\t. (f t - f x) / (t - x)) \ y) (at x within ({x<..} \ I))" by (simp add: Lim_null[symmetric]) also have "\ \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (simp_all add: field_simps) finally show ?thesis by (simp add: bounded_linear_mult_right has_derivative_within) qed subsubsection \Caratheodory characterization\ lemma DERIV_caratheodory_within: "(f has_field_derivative l) (at x within S) \ (\g. (\z. f z - f x = g z * (z - x)) \ continuous (at x within S) g \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs proof (intro exI conjI) let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z-x)" by simp show "continuous (at x within S) ?g" using \?lhs\ by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) show "?g x = l" by simp qed next assume ?rhs then obtain g where "(\z. f z - f x = g z * (z-x))" and "continuous (at x within S) g" and "g x = l" by blast thus ?lhs by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) qed subsection \Differentiability\ definition\<^marker>\tag important\ differentiable_on :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a set \ bool" (infix "differentiable'_on" 50) where "f differentiable_on s \ (\x\s. f differentiable (at x within s))" lemma differentiableI: "(f has_derivative f') net \ f differentiable net" unfolding differentiable_def by auto lemma differentiable_onD: "\f differentiable_on S; x \ S\ \ f differentiable (at x within S)" using differentiable_on_def by blast lemma differentiable_at_withinI: "f differentiable (at x) \ f differentiable (at x within s)" unfolding differentiable_def using has_derivative_at_withinI by blast lemma differentiable_at_imp_differentiable_on: "(\x. x \ s \ f differentiable at x) \ f differentiable_on s" by (metis differentiable_at_withinI differentiable_on_def) corollary\<^marker>\tag unimportant\ differentiable_iff_scaleR: fixes f :: "real \ 'a::real_normed_vector" shows "f differentiable F \ (\d. (f has_derivative (\x. x *\<^sub>R d)) F)" by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR) lemma differentiable_on_eq_differentiable_at: "open s \ f differentiable_on s \ (\x\s. f differentiable at x)" unfolding differentiable_on_def by (metis at_within_interior interior_open) lemma differentiable_transform_within: assumes "f differentiable (at x within s)" and "0 < d" and "x \ s" and "\x'. \x'\s; dist x' x < d\ \ f x' = g x'" shows "g differentiable (at x within s)" using assms has_derivative_transform_within unfolding differentiable_def by blast lemma differentiable_on_ident [simp, derivative_intros]: "(\x. x) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_on_id [simp, derivative_intros]: "id differentiable_on S" by (simp add: id_def) lemma differentiable_on_const [simp, derivative_intros]: "(\z. c) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_mult [simp, derivative_intros]: fixes f :: "'M::real_normed_vector \ 'a::real_normed_algebra" shows "\f differentiable_on S; g differentiable_on S\ \ (\z. f z * g z) differentiable_on S" unfolding differentiable_on_def differentiable_def using differentiable_def differentiable_mult by blast lemma differentiable_on_compose: "\g differentiable_on S; f differentiable_on (g ` S)\ \ (\x. f (g x)) differentiable_on S" by (simp add: differentiable_in_compose differentiable_on_def) lemma bounded_linear_imp_differentiable_on: "bounded_linear f \ f differentiable_on S" by (simp add: differentiable_on_def bounded_linear_imp_differentiable) lemma linear_imp_differentiable_on: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable_on S" by (simp add: differentiable_on_def linear_imp_differentiable) lemma differentiable_on_minus [simp, derivative_intros]: "f differentiable_on S \ (\z. -(f z)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_add [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z + g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_diff [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z - g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_inverse [simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" shows "f differentiable_on S \ (\x. x \ S \ f x \ 0) \ (\x. inverse (f x)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_scaleR [derivative_intros, simp]: "\f differentiable_on S; g differentiable_on S\ \ (\x. f x *\<^sub>R g x) differentiable_on S" unfolding differentiable_on_def by (blast intro: differentiable_scaleR) lemma has_derivative_sqnorm_at [derivative_intros, simp]: "((\x. (norm x)\<^sup>2) has_derivative (\x. 2 *\<^sub>R (a \ x))) (at a)" using bounded_bilinear.FDERIV [of "(\)" id id a _ id id] by (auto simp: inner_commute dot_square_norm bounded_bilinear_inner) lemma differentiable_sqnorm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "(\x. (norm x)\<^sup>2) differentiable (at a)" by (force simp add: differentiable_def intro: has_derivative_sqnorm_at) lemma differentiable_on_sqnorm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "(\x. (norm x)\<^sup>2) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_norm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "a \ 0 \ norm differentiable (at a)" using differentiableI has_derivative_norm by blast lemma differentiable_on_norm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "0 \ S \ norm differentiable_on S" by (metis differentiable_at_imp_differentiable_on differentiable_norm_at) subsection \Frechet derivative and Jacobian matrix\ definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)" proposition frechet_derivative_works: "f differentiable net \ (f has_derivative (frechet_derivative f net)) net" unfolding frechet_derivative_def differentiable_def unfolding some_eq_ex[of "\ f' . (f has_derivative f') net"] .. lemma linear_frechet_derivative: "f differentiable net \ linear (frechet_derivative f net)" unfolding frechet_derivative_works has_derivative_def by (auto intro: bounded_linear.linear) lemma frechet_derivative_const [simp]: "frechet_derivative (\x. c) (at a) = (\x. 0)" using differentiable_const frechet_derivative_works has_derivative_const has_derivative_unique by blast lemma frechet_derivative_id [simp]: "frechet_derivative id (at a) = id" using differentiable_def frechet_derivative_works has_derivative_id has_derivative_unique by blast lemma frechet_derivative_ident [simp]: "frechet_derivative (\x. x) (at a) = (\x. x)" by (metis eq_id_iff frechet_derivative_id) subsection \Differentiability implies continuity\ proposition differentiable_imp_continuous_within: "f differentiable (at x within s) \ continuous (at x within s) f" by (auto simp: differentiable_def intro: has_derivative_continuous) lemma differentiable_imp_continuous_on: "f differentiable_on s \ continuous_on s f" unfolding differentiable_on_def continuous_on_eq_continuous_within using differentiable_imp_continuous_within by blast lemma differentiable_on_subset: "f differentiable_on t \ s \ t \ f differentiable_on s" unfolding differentiable_on_def using differentiable_within_subset by blast lemma differentiable_on_empty: "f differentiable_on {}" unfolding differentiable_on_def by auto lemma has_derivative_continuous_on: "(\x. x \ s \ (f has_derivative f' x) (at x within s)) \ continuous_on s f" by (auto intro!: differentiable_imp_continuous_on differentiableI simp: differentiable_on_def) text \Results about neighborhoods filter.\ lemma eventually_nhds_metric_le: "eventually P (nhds a) = (\d>0. \x. dist x a \ d \ P x)" unfolding eventually_nhds_metric by (safe, rule_tac x="d / 2" in exI, auto) lemma le_nhds: "F \ nhds a \ (\S. open S \ a \ S \ eventually (\x. x \ S) F)" unfolding le_filter_def eventually_nhds by (fast elim: eventually_mono) lemma le_nhds_metric: "F \ nhds a \ (\e>0. eventually (\x. dist x a < e) F)" unfolding le_filter_def eventually_nhds_metric by (fast elim: eventually_mono) lemma le_nhds_metric_le: "F \ nhds a \ (\e>0. eventually (\x. dist x a \ e) F)" unfolding le_filter_def eventually_nhds_metric_le by (fast elim: eventually_mono) text \Several results are easier using a "multiplied-out" variant. (I got this idea from Dieudonne's proof of the chain rule).\ lemma has_derivative_within_alt: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. \d>0. \y\s. norm(y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_within_alt2: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. eventually (\y. norm (f y - f x - f' (y - x)) \ e * norm (y - x)) (at x within s))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_at_alt: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \y. norm(y - x) < d \ norm (f y - f x - f'(y - x)) \ e * norm (y - x))" using has_derivative_within_alt[where s=UNIV] by simp subsection \The chain rule\ proposition diff_chain_within[derivative_intros]: assumes "(f has_derivative f') (at x within s)" and "(g has_derivative g') (at (f x) within (f ` s))" shows "((g \ f) has_derivative (g' \ f'))(at x within s)" using has_derivative_in_compose[OF assms] by (simp add: comp_def) lemma diff_chain_at[derivative_intros]: "(f has_derivative f') (at x) \ (g has_derivative g') (at (f x)) \ ((g \ f) has_derivative (g' \ f')) (at x)" using has_derivative_compose[of f f' x UNIV g g'] by (simp add: comp_def) lemma has_vector_derivative_within_open: "a \ S \ open S \ (f has_vector_derivative f') (at a within S) \ (f has_vector_derivative f') (at a)" by (simp only: at_within_interior interior_open) lemma field_vector_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_field_derivative g') (at (f x) within f ` S)" shows "((g \ f) has_vector_derivative (f' * g')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_derivative g') (at (f x) within f`S)" shows "((g \ f) has_vector_derivative (g' f')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg] linear.scaleR[OF has_derivative_linear[OF Dg]] unfolding has_vector_derivative_def o_def by (auto simp: o_def mult.commute has_vector_derivative_def) subsection\<^marker>\tag unimportant\ \Composition rules stated just for differentiability\ lemma differentiable_chain_at: "f differentiable (at x) \ g differentiable (at (f x)) \ (g \ f) differentiable (at x)" unfolding differentiable_def by (meson diff_chain_at) lemma differentiable_chain_within: "f differentiable (at x within S) \ g differentiable (at(f x) within (f ` S)) \ (g \ f) differentiable (at x within S)" unfolding differentiable_def by (meson diff_chain_within) subsection \Uniqueness of derivative\ text\<^marker>\tag important\ \ The general result is a bit messy because we need approachability of the limit point from any direction. But OK for nontrivial intervals etc. \ proposition frechet_derivative_unique_within: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes 1: "(f has_derivative f') (at x within S)" and 2: "(f has_derivative f'') (at x within S)" and S: "\i e. \i\Basis; e>0\ \ \d. 0 < \d\ \ \d\ < e \ (x + d *\<^sub>R i) \ S" shows "f' = f''" proof - note as = assms(1,2)[unfolded has_derivative_def] then interpret f': bounded_linear f' by auto from as interpret f'': bounded_linear f'' by auto have "x islimpt S" unfolding islimpt_approachable proof (intro allI impI) fix e :: real assume "e > 0" obtain d where "0 < \d\" and "\d\ < e" and "x + d *\<^sub>R (SOME i. i \ Basis) \ S" using assms(3) SOME_Basis \e>0\ by blast then show "\x'\S. x' \ x \ dist x' x < e" by (rule_tac x="x + d *\<^sub>R (SOME i. i \ Basis)" in bexI) (auto simp: dist_norm SOME_Basis nonzero_Basis) qed then have *: "netlimit (at x within S) = x" by (simp add: Lim_ident_at trivial_limit_within) show ?thesis proof (rule linear_eq_stdbasis) show "linear f'" "linear f''" unfolding linear_conv_bounded_linear using as by auto next fix i :: 'a assume i: "i \ Basis" define e where "e = norm (f' i - f'' i)" show "f' i = f'' i" proof (rule ccontr) assume "f' i \ f'' i" then have "e > 0" unfolding e_def by auto obtain d where d: "0 < d" "(\y. y\S \ 0 < dist y x \ dist y x < d \ dist ((f y - f x - f' (y - x)) /\<^sub>R norm (y - x) - (f y - f x - f'' (y - x)) /\<^sub>R norm (y - x)) (0 - 0) < e)" using tendsto_diff [OF as(1,2)[THEN conjunct2]] unfolding * Lim_within using \e>0\ by blast obtain c where c: "0 < \c\" "\c\ < d \ x + c *\<^sub>R i \ S" using assms(3) i d(1) by blast have *: "norm (- ((1 / \c\) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \c\) *\<^sub>R f'' (c *\<^sub>R i)) = norm ((1 / \c\) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))" unfolding scaleR_right_distrib by auto also have "\ = norm ((1 / \c\) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))" unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto also have "\ = e" unfolding e_def using c(1) using norm_minus_cancel[of "f' i - f'' i"] by auto finally show False using c using d(2)[of "x + c *\<^sub>R i"] unfolding dist_norm unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib using i by (auto simp: inverse_eq_divide) qed qed qed proposition frechet_derivative_unique_within_closed_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes ab: "\i. i\Basis \ a\i < b\i" and x: "x \ cbox a b" and "(f has_derivative f' ) (at x within cbox a b)" and "(f has_derivative f'') (at x within cbox a b)" shows "f' = f''" proof (rule frechet_derivative_unique_within) fix e :: real fix i :: 'a assume "e > 0" and i: "i \ Basis" then show "\d. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R i \ cbox a b" proof (cases "x\i = a\i") case True with ab[of i] \e>0\ x i show ?thesis by (rule_tac x="(min (b\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) next case False moreover have "a \ i < x \ i" using False i mem_box(2) x by force moreover { have "a \ i * 2 + min (x \ i - a \ i) e \ a\i *2 + x\i - a\i" by auto also have "\ = a\i + x\i" by auto also have "\ \ 2 * (x\i)" using \a \ i < x \ i\ by auto finally have "a \ i * 2 + min (x \ i - a \ i) e \ x \ i * 2" by auto } moreover have "min (x \ i - a \ i) e \ 0" by (simp add: \0 < e\ \a \ i < x \ i\ less_eq_real_def) then have "x \ i * 2 \ b \ i * 2 + min (x \ i - a \ i) e" using i mem_box(2) x by force ultimately show ?thesis using ab[of i] \e>0\ x i by (rule_tac x="- (min (x\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) qed qed (use assms in auto) lemma frechet_derivative_unique_within_open_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes x: "x \ box a b" and f: "(f has_derivative f' ) (at x within box a b)" "(f has_derivative f'') (at x within box a b)" shows "f' = f''" proof - have "at x within box a b = at x" by (metis x at_within_interior interior_open open_box) with f show "f' = f''" by (simp add: has_derivative_unique) qed lemma frechet_derivative_at: "(f has_derivative f') (at x) \ f' = frechet_derivative f (at x)" using differentiable_def frechet_derivative_works has_derivative_unique by blast lemma frechet_derivative_compose: "frechet_derivative (f o g) (at x) = frechet_derivative (f) (at (g x)) o frechet_derivative g (at x)" if "g differentiable at x" "f differentiable at (g x)" by (metis diff_chain_at frechet_derivative_at frechet_derivative_works that) lemma frechet_derivative_within_cbox: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "\i. i\Basis \ a\i < b\i" and "x \ cbox a b" and "(f has_derivative f') (at x within cbox a b)" shows "frechet_derivative f (at x within cbox a b) = f'" using assms by (metis Derivative.differentiableI frechet_derivative_unique_within_closed_interval frechet_derivative_works) lemma frechet_derivative_transform_within_open: "frechet_derivative f (at x) = frechet_derivative g (at x)" if "f differentiable at x" "open X" "x \ X" "\x. x \ X \ f x = g x" by (meson frechet_derivative_at frechet_derivative_works has_derivative_transform_within_open that) subsection \Derivatives of local minima and maxima are zero\ lemma has_derivative_local_min: fixes f :: "'a::real_normed_vector \ real" assumes deriv: "(f has_derivative f') (at x)" assumes min: "eventually (\y. f x \ f y) (at x)" shows "f' = (\h. 0)" proof fix h :: 'a interpret f': bounded_linear f' using deriv by (rule has_derivative_bounded_linear) show "f' h = 0" proof (cases "h = 0") case False from min obtain d where d1: "0 < d" and d2: "\y\ball x d. f x \ f y" unfolding eventually_at by (force simp: dist_commute) have "FDERIV (\r. x + r *\<^sub>R h) 0 :> (\r. r *\<^sub>R h)" by (intro derivative_eq_intros) auto then have "FDERIV (\r. f (x + r *\<^sub>R h)) 0 :> (\k. f' (k *\<^sub>R h))" by (rule has_derivative_compose, simp add: deriv) then have "DERIV (\r. f (x + r *\<^sub>R h)) 0 :> f' h" unfolding has_field_derivative_def by (simp add: f'.scaleR mult_commute_abs) moreover have "0 < d / norm h" using d1 and \h \ 0\ by simp moreover have "\y. \0 - y\ < d / norm h \ f (x + 0 *\<^sub>R h) \ f (x + y *\<^sub>R h)" using \h \ 0\ by (auto simp add: d2 dist_norm pos_less_divide_eq) ultimately show "f' h = 0" by (rule DERIV_local_min) qed simp qed lemma has_derivative_local_max: fixes f :: "'a::real_normed_vector \ real" assumes "(f has_derivative f') (at x)" assumes "eventually (\y. f y \ f x) (at x)" shows "f' = (\h. 0)" using has_derivative_local_min [of "\x. - f x" "\h. - f' h" "x"] using assms unfolding fun_eq_iff by simp lemma differential_zero_maxmin: fixes f::"'a::real_normed_vector \ real" assumes "x \ S" and "open S" and deriv: "(f has_derivative f') (at x)" and mono: "(\y\S. f y \ f x) \ (\y\S. f x \ f y)" shows "f' = (\v. 0)" using mono proof assume "\y\S. f y \ f x" with \x \ S\ and \open S\ have "eventually (\y. f y \ f x) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_max) next assume "\y\S. f x \ f y" with \x \ S\ and \open S\ have "eventually (\y. f x \ f y) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_min) qed lemma differential_zero_maxmin_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" and ball: "0 < e" "(\y \ ball x e. (f y)\k \ (f x)\k) \ (\y\ball x e. (f x)\k \ (f y)\k)" and diff: "f differentiable (at x)" shows "(\j\Basis. (frechet_derivative f (at x) j \ k) *\<^sub>R j) = (0::'a)" (is "?D k = 0") proof - let ?f' = "frechet_derivative f (at x)" have "x \ ball x e" using \0 < e\ by simp moreover have "open (ball x e)" by simp moreover have "((\x. f x \ k) has_derivative (\h. ?f' h \ k)) (at x)" using bounded_linear_inner_left diff[unfolded frechet_derivative_works] by (rule bounded_linear.has_derivative) ultimately have "(\h. frechet_derivative f (at x) h \ k) = (\v. 0)" using ball(2) by (rule differential_zero_maxmin) then show ?thesis unfolding fun_eq_iff by simp qed subsection \One-dimensional mean value theorem\ lemma mvt_simple: fixes f :: "real \ real" assumes "a < b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a<.. real" assumes "a \ b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a..b}. f b - f a = f' x (b - a)" proof (cases "a = b") interpret bounded_linear "f' b" using assms(2) assms(1) by auto case True then show ?thesis by force next case False then show ?thesis using mvt_simple[OF _ derf] by (metis \a \ b\ atLeastAtMost_iff dual_order.order_iff_strict greaterThanLessThan_iff) qed text \A nice generalization (see Havin's proof of 5.19 from Rudin's book).\ lemma mvt_general: fixes f :: "real \ 'a::real_inner" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\x\{a<.. norm (f' x (b - a))" proof - have "\x\{a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" apply (rule mvt [OF \a < b\, where f = "\x. (f b - f a) \ f x"]) apply (intro continuous_intros contf) using derf apply (auto intro: has_derivative_inner_right) done then obtain x where x: "x \ {a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" .. show ?thesis proof (cases "f a = f b") case False have "norm (f b - f a) * norm (f b - f a) = (norm (f b - f a))\<^sup>2" by (simp add: power2_eq_square) also have "\ = (f b - f a) \ (f b - f a)" unfolding power2_norm_eq_inner .. also have "\ = (f b - f a) \ f' x (b - a)" using x(2) by (simp only: inner_diff_right) also have "\ \ norm (f b - f a) * norm (f' x (b - a))" by (rule norm_cauchy_schwarz) finally show ?thesis using False x(1) by (auto simp add: mult_left_cancel) next case True then show ?thesis using \a < b\ by (rule_tac x="(a + b) /2" in bexI) auto qed qed subsection \More general bound theorems\ proposition differentiable_bound_general: fixes f :: "real \ 'a::real_normed_vector" assumes "a < b" and f_cont: "continuous_on {a..b} f" and phi_cont: "continuous_on {a..b} \" and f': "\x. a < x \ x < b \ (f has_vector_derivative f' x) (at x)" and phi': "\x. a < x \ x < b \ (\ has_vector_derivative \' x) (at x)" and bnd: "\x. a < x \ x < b \ norm (f' x) \ \' x" shows "norm (f b - f a) \ \ b - \ a" proof - { fix x assume x: "a < x" "x < b" have "0 \ norm (f' x)" by simp also have "\ \ \' x" using x by (auto intro!: bnd) finally have "0 \ \' x" . } note phi'_nonneg = this note f_tendsto = assms(2)[simplified continuous_on_def, rule_format] note phi_tendsto = assms(3)[simplified continuous_on_def, rule_format] { fix e::real assume "e > 0" define e2 where "e2 = e / 2" with \e > 0\ have "e2 > 0" by simp let ?le = "\x1. norm (f x1 - f a) \ \ x1 - \ a + e * (x1 - a) + e" define A where "A = {x2. a \ x2 \ x2 \ b \ (\x1\{a ..< x2}. ?le x1)}" have A_subset: "A \ {a..b}" by (auto simp: A_def) { fix x2 assume a: "a \ x2" "x2 \ b" and le: "\x1\{a..e > 0\ proof cases assume "x2 \ a" with a have "a < x2" by simp have "at x2 within {a <.. bot" using \a < x2\ by (auto simp: trivial_limit_within islimpt_in_closure) moreover have "((\x1. (\ x1 - \ a) + e * (x1 - a) + e) \ (\ x2 - \ a) + e * (x2 - a) + e) (at x2 within {a <..x1. norm (f x1 - f a)) \ norm (f x2 - f a)) (at x2 within {a <..x. x > a) (at x2 within {a <.. A" using assms by (auto simp: A_def) hence [simp]: "A \ {}" by auto have A_ivl: "\x1 x2. x2 \ A \ x1 \ {a ..x2} \ x1 \ A" by (simp add: A_def) have [simp]: "bdd_above A" by (auto simp: A_def) define y where "y = Sup A" have "y \ b" unfolding y_def by (simp add: cSup_le_iff) (simp add: A_def) have leI: "\x x1. a \ x1 \ x \ A \ x1 < x \ ?le x1" by (auto simp: A_def intro!: le_cont) have y_all_le: "\x1\{a.. y" by (metis \a \ A\ \bdd_above A\ cSup_upper y_def) have "y \ A" using y_all_le \a \ y\ \y \ b\ by (auto simp: A_def) hence "A = {a .. y}" using A_subset by (auto simp: subset_iff y_def cSup_upper intro: A_ivl) from le_cont[OF \a \ y\ \y \ b\ y_all_le] have le_y: "?le y" . have "y = b" proof (cases "a = y") case True with \a < b\ have "y < b" by simp with \a = y\ f_cont phi_cont \e2 > 0\ have 1: "\\<^sub>F x in at y within {y..b}. dist (f x) (f y) < e2" and 2: "\\<^sub>F x in at y within {y..b}. dist (\ x) (\ y) < e2" by (auto simp: continuous_on_def tendsto_iff) have 3: "eventually (\x. y < x) (at y within {y..b})" by (auto simp: eventually_at_filter) have 4: "eventually (\x::real. x < b) (at y within {y..b})" using _ \y < b\ by (rule order_tendstoD) (auto intro!: tendsto_eq_intros) from 1 2 3 4 have eventually_le: "eventually (\x. ?le x) (at y within {y .. b})" proof eventually_elim case (elim x1) have "norm (f x1 - f a) = norm (f x1 - f y)" by (simp add: \a = y\) also have "norm (f x1 - f y) \ e2" using elim \a = y\ by (auto simp : dist_norm intro!: less_imp_le) also have "\ \ e2 + (\ x1 - \ a + e2 + e * (x1 - a))" using \0 < e\ elim by (intro add_increasing2[OF add_nonneg_nonneg order.refl]) (auto simp: \a = y\ dist_norm intro!: mult_nonneg_nonneg) also have "\ = \ x1 - \ a + e * (x1 - a) + e" by (simp add: e2_def) finally show "?le x1" . qed from this[unfolded eventually_at_topological] \?le y\ obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y..b} \ ?le x" by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min b (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a = y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" by (simp add: d'_def) fix x1 assume "x1 \ {a.. S" "x1 \ {y..b}" by (auto simp: \a = y\ d'_def dist_real_def intro!: d ) thus "?le x1" by (rule S) qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp then show "y = b" using \d > 0\ \y < b\ by (simp add: d'_def) next case False with \a \ y\ have "a < y" by simp show "y = b" proof (rule ccontr) assume "y \ b" hence "y < b" using \y \ b\ by simp let ?F = "at y within {y.. has_vector_derivative \' y) ?F" using \a < y\ \y < b\ by (auto simp add: at_within_open[of _ "{a<..\<^sub>F x1 in ?F. norm (f x1 - f y - (x1 - y) *\<^sub>R f' y) \ e2 * \x1 - y\" "\\<^sub>F x1 in ?F. norm (\ x1 - \ y - (x1 - y) *\<^sub>R \' y) \ e2 * \x1 - y\" using \e2 > 0\ by (auto simp: has_derivative_within_alt2 has_vector_derivative_def) moreover have "\\<^sub>F x1 in ?F. y \ x1" "\\<^sub>F x1 in ?F. x1 < b" by (auto simp: eventually_at_filter) ultimately have "\\<^sub>F x1 in ?F. norm (f x1 - f y) \ (\ x1 - \ y) + e * \x1 - y\" (is "\\<^sub>F x1 in ?F. ?le' x1") proof eventually_elim case (elim x1) from norm_triangle_ineq2[THEN order_trans, OF elim(1)] have "norm (f x1 - f y) \ norm (f' y) * \x1 - y\ + e2 * \x1 - y\" by (simp add: ac_simps) also have "norm (f' y) \ \' y" using bnd \a < y\ \y < b\ by simp also have "\' y * \x1 - y\ \ \ x1 - \ y + e2 * \x1 - y\" using elim by (simp add: ac_simps) finally have "norm (f x1 - f y) \ \ x1 - \ y + e2 * \x1 - y\ + e2 * \x1 - y\" by (auto simp: mult_right_mono) thus ?case by (simp add: e2_def) qed moreover have "?le' y" by simp ultimately obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y.. ?le' x" unfolding eventually_at_topological by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min ((y + b)/2) (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a < y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" using \y < b\ by (simp add: d'_def min_def) fix x1 assume x1: "x1 \ {a..y \ A\ local.leI x1 by auto next case False hence x1': "x1 \ S" "x1 \ {y.. norm (f x1 - f y) + norm (f y - f a)" by (rule order_trans[OF _ norm_triangle_ineq]) simp also note S(3)[OF x1'] also note le_y finally show "?le x1" using False by (auto simp: algebra_simps) qed qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp thus False using \d > 0\ \y < b\ by (simp add: d'_def min_def split: if_split_asm) qed qed with le_y have "norm (f b - f a) \ \ b - \ a + e * (b - a + 1)" by (simp add: algebra_simps) } note * = this show ?thesis proof (rule field_le_epsilon) fix e::real assume "e > 0" then show "norm (f b - f a) \ \ b - \ a + e" using *[of "e / (b - a + 1)"] \a < b\ by simp qed qed lemma differentiable_bound: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\x. x\S \ (f has_derivative f' x) (at x within S)" and B: "\x. x \ S \ onorm (f' x) \ B" and x: "x \ S" and y: "y \ S" shows "norm (f x - f y) \ B * norm (x - y)" proof - let ?p = "\u. x + u *\<^sub>R (y - x)" let ?\ = "\h. h * B * norm (x - y)" have *: "x + u *\<^sub>R (y - x) \ S" if "u \ {0..1}" for u proof - have "u *\<^sub>R y = u *\<^sub>R (y - x) + u *\<^sub>R x" by (simp add: scale_right_diff_distrib) then show "x + u *\<^sub>R (y - x) \ S" using that \convex S\ x y by (simp add: convex_alt) (metis pth_b(2) pth_c(1) scaleR_collapse) qed have "\z. z \ (\u. x + u *\<^sub>R (y - x)) ` {0..1} \ (f has_derivative f' z) (at z within (\u. x + u *\<^sub>R (y - x)) ` {0..1})" by (auto intro: * has_derivative_subset [OF derf]) then have "continuous_on (?p ` {0..1}) f" unfolding continuous_on_eq_continuous_within by (meson has_derivative_continuous) with * have 1: "continuous_on {0 .. 1} (f \ ?p)" by (intro continuous_intros)+ { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret linear "(f' ?u)" using u by (auto intro!: has_derivative_linear derf *) have "(f \ ?p has_derivative (f' ?u) \ (\u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)" by (intro derivative_intros has_derivative_subset [OF derf]) (use u * in auto) hence "((f \ ?p) has_vector_derivative f' ?u (y - x)) (at u)" by (simp add: at_within_open[OF u open_greaterThanLessThan] scaleR has_vector_derivative_def o_def) } note 2 = this have 3: "continuous_on {0..1} ?\" by (rule continuous_intros)+ have 4: "(?\ has_vector_derivative B * norm (x - y)) (at u)" for u by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros) { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret bounded_linear "(f' ?u)" using u by (auto intro!: has_derivative_bounded_linear derf *) have "norm (f' ?u (y - x)) \ onorm (f' ?u) * norm (y - x)" by (rule onorm) (rule bounded_linear) also have "onorm (f' ?u) \ B" using u by (auto intro!: assms(3)[rule_format] *) finally have "norm ((f' ?u) (y - x)) \ B * norm (x - y)" by (simp add: mult_right_mono norm_minus_commute) } note 5 = this have "norm (f x - f y) = norm ((f \ (\u. x + u *\<^sub>R (y - x))) 1 - (f \ (\u. x + u *\<^sub>R (y - x))) 0)" by (auto simp add: norm_minus_commute) also from differentiable_bound_general[OF zero_less_one 1, OF 3 2 4 5] have "norm ((f \ ?p) 1 - (f \ ?p) 0) \ B * norm (x - y)" by simp finally show ?thesis . qed lemma field_differentiable_bound: fixes S :: "'a::real_normed_field set" assumes cvs: "convex S" and df: "\z. z \ S \ (f has_field_derivative f' z) (at z within S)" and dn: "\z. z \ S \ norm (f' z) \ B" and "x \ S" "y \ S" shows "norm(f x - f y) \ B * norm(x - y)" apply (rule differentiable_bound [OF cvs]) apply (erule df [unfolded has_field_derivative_def]) apply (rule onorm_le, simp_all add: norm_mult mult_right_mono assms) done lemma differentiable_bound_segment: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes "\t. t \ {0..1} \ x0 + t *\<^sub>R a \ G" assumes f': "\x. x \ G \ (f has_derivative f' x) (at x within G)" assumes B: "\x. x \ {0..1} \ onorm (f' (x0 + x *\<^sub>R a)) \ B" shows "norm (f (x0 + a) - f x0) \ norm a * B" proof - let ?G = "(\x. x0 + x *\<^sub>R a) ` {0..1}" have "?G = (+) x0 ` (\x. x *\<^sub>R a) ` {0..1}" by auto also have "convex \" by (intro convex_translation convex_scaled convex_real_interval) finally have "convex ?G" . moreover have "?G \ G" "x0 \ ?G" "x0 + a \ ?G" using assms by (auto intro: image_eqI[where x=1]) ultimately show ?thesis using has_derivative_subset[OF f' \?G \ G\] B differentiable_bound[of "(\x. x0 + x *\<^sub>R a) ` {0..1}" f f' B "x0 + a" x0] by (force simp: ac_simps) qed lemma differentiable_bound_linearization: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes S: "\t. t \ {0..1} \ a + t *\<^sub>R (b - a) \ S" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes B: "\x. x \ S \ onorm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - f' x0 (b - a)) \ norm (b - a) * B" proof - define g where [abs_def]: "g x = f x - f' x0 x" for x have g: "\x. x \ S \ (g has_derivative (\i. f' x i - f' x0 i)) (at x within S)" unfolding g_def using assms by (auto intro!: derivative_eq_intros bounded_linear.has_derivative[OF has_derivative_bounded_linear, OF f']) from B have "\x\{0..1}. onorm (\i. f' (a + x *\<^sub>R (b - a)) i - f' x0 i) \ B" using assms by (auto simp: fun_diff_def) with differentiable_bound_segment[OF S g] \x0 \ S\ show ?thesis by (simp add: g_def field_simps linear_diff[OF has_derivative_linear[OF f']]) qed lemma vector_differentiable_bound_linearization: fixes f::"real \ 'b::real_normed_vector" assumes f': "\x. x \ S \ (f has_vector_derivative f' x) (at x within S)" assumes "closed_segment a b \ S" assumes B: "\x. x \ S \ norm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - (b - a) *\<^sub>R f' x0) \ norm (b - a) * B" using assms by (intro differentiable_bound_linearization[of a b S f "\x h. h *\<^sub>R f' x" x0 B]) (force simp: closed_segment_real_eq has_vector_derivative_def scaleR_diff_right[symmetric] mult.commute[of B] intro!: onorm_le mult_left_mono)+ text \In particular.\ lemma has_derivative_zero_constant: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" shows "\c. \x\s. f x = c" proof - { fix x y assume "x \ s" "y \ s" then have "norm (f x - f y) \ 0 * norm (x - y)" using assms by (intro differentiable_bound[of s]) (auto simp: onorm_zero) then have "f x = f y" by simp } then show ?thesis by metis qed lemma has_field_derivative_zero_constant: assumes "convex s" "\x. x \ s \ (f has_field_derivative 0) (at x within s)" shows "\c. \x\s. f (x) = (c :: 'a :: real_normed_field)" proof (rule has_derivative_zero_constant) have A: "(*) 0 = (\_. 0 :: 'a)" by (intro ext) simp fix x assume "x \ s" thus "(f has_derivative (\h. 0)) (at x within s)" using assms(2)[of x] by (simp add: has_field_derivative_def A) qed fact lemma has_vector_derivative_zero_constant: assumes "convex s" assumes "\x. x \ s \ (f has_vector_derivative 0) (at x within s)" obtains c where "\x. x \ s \ f x = c" using has_derivative_zero_constant[of s f] assms by (auto simp: has_vector_derivative_def) lemma has_derivative_zero_unique: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" and "x \ s" "y \ s" shows "f x = f y" using has_derivative_zero_constant[OF assms(1,2)] assms(3-) by force lemma has_derivative_zero_unique_connected: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open s" "connected s" assumes f: "\x. x \ s \ (f has_derivative (\x. 0)) (at x)" assumes "x \ s" "y \ s" shows "f x = f y" proof (rule connected_local_const[where f=f, OF \connected s\ \x\s\ \y\s\]) show "\a\s. eventually (\b. f a = f b) (at a within s)" proof fix a assume "a \ s" with \open s\ obtain e where "0 < e" "ball a e \ s" by (rule openE) then have "\c. \x\ball a e. f x = c" by (intro has_derivative_zero_constant) (auto simp: at_within_open[OF _ open_ball] f) with \0 have "\x\ball a e. f a = f x" by auto then show "eventually (\b. f a = f b) (at a within s)" using \0 unfolding eventually_at_topological by (intro exI[of _ "ball a e"]) auto qed qed subsection \Differentiability of inverse function (most basic form)\ lemma has_derivative_inverse_basic: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes derf: "(f has_derivative f') (at (g y))" and ling': "bounded_linear g'" and "g' \ f' = id" and contg: "continuous (at y) g" and "open T" and "y \ T" and fg: "\z. z \ T \ f (g z) = z" shows "(g has_derivative g') (at y)" proof - interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto interpret g': bounded_linear g' using assms by auto obtain C where C: "0 < C" "\x. norm (g' x) \ norm x * C" using bounded_linear.pos_bounded[OF assms(2)] by blast have lem1: "\e>0. \d>0. \z. norm (z - y) < d \ norm (g z - g y - g'(z - y)) \ e * norm (g z - g y)" proof (intro allI impI) fix e :: real assume "e > 0" with C(1) have *: "e / C > 0" by auto obtain d0 where "0 < d0" and d0: "\u. norm (u - g y) < d0 \ norm (f u - f (g y) - f' (u - g y)) \ e / C * norm (u - g y)" using derf * unfolding has_derivative_at_alt by blast obtain d1 where "0 < d1" and d1: "\x. \0 < dist x y; dist x y < d1\ \ dist (g x) (g y) < d0" using contg \0 < d0\ unfolding continuous_at Lim_at by blast obtain d2 where "0 < d2" and d2: "\u. dist u y < d2 \ u \ T" using \open T\ \y \ T\ unfolding open_dist by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF \0 < d1\ \0 < d2\] by blast show "\d>0. \z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < d" then have "z \ T" using d2 d unfolding dist_norm by auto have "norm (g z - g y - g' (z - y)) \ norm (g' (f (g z) - y - f' (g z - g y)))" unfolding g'.diff f'.diff unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] fg[OF \z\T\] by (simp add: norm_minus_commute) also have "\ \ norm (f (g z) - y - f' (g z - g y)) * C" by (rule C(2)) also have "\ \ (e / C) * norm (g z - g y) * C" proof - have "norm (g z - g y) < d0" by (metis as cancel_comm_monoid_add_class.diff_cancel d(2) \0 < d0\ d1 diff_gt_0_iff_gt diff_strict_mono dist_norm dist_self zero_less_dist_iff) then show ?thesis - by (metis C(1) \y \ T\ d0 fg real_mult_le_cancel_iff1) + by (metis C(1) \y \ T\ d0 fg mult_le_cancel_iff1) qed also have "\ \ e * norm (g z - g y)" using C by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" by simp qed (use d in auto) qed have *: "(0::real) < 1 / 2" by auto obtain d where "0 < d" and d: "\z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ 1/2 * norm (g z - g y)" using lem1 * by blast define B where "B = C * 2" have "B > 0" unfolding B_def using C by auto have lem2: "norm (g z - g y) \ B * norm (z - y)" if z: "norm(z - y) < d" for z proof - have "norm (g z - g y) \ norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" by (rule norm_triangle_sub) also have "\ \ norm (g' (z - y)) + 1 / 2 * norm (g z - g y)" by (rule add_left_mono) (use d z in auto) also have "\ \ norm (z - y) * C + 1 / 2 * norm (g z - g y)" by (rule add_right_mono) (use C in auto) finally show "norm (g z - g y) \ B * norm (z - y)" unfolding B_def by (auto simp add: field_simps) qed show ?thesis unfolding has_derivative_at_alt proof (intro conjI assms allI impI) fix e :: real assume "e > 0" then have *: "e / B > 0" by (metis \B > 0\ divide_pos_pos) obtain d' where "0 < d'" and d': "\z. norm (z - y) < d' \ norm (g z - g y - g' (z - y)) \ e / B * norm (g z - g y)" using lem1 * by blast obtain k where k: "0 < k" "k < d" "k < d'" using field_lbound_gt_zero[OF \0 < d\ \0 < d'\] by blast show "\d>0. \ya. norm (ya - y) < d \ norm (g ya - g y - g' (ya - y)) \ e * norm (ya - y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < k" then have "norm (g z - g y - g' (z - y)) \ e / B * norm(g z - g y)" using d' k by auto also have "\ \ e * norm (z - y)" unfolding times_divide_eq_left pos_divide_le_eq[OF \B>0\] using lem2[of z] k as \e > 0\ by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (z - y)" by simp qed (use k in auto) qed qed text\<^marker>\tag unimportant\\Inverse function theorem for complex derivatives\ lemma has_field_derivative_inverse_basic: shows "DERIV f (g y) :> f' \ f' \ 0 \ continuous (at y) g \ open t \ y \ t \ (\z. z \ t \ f (g z) = z) \ DERIV g y :> inverse (f')" unfolding has_field_derivative_def apply (rule has_derivative_inverse_basic) apply (auto simp: bounded_linear_mult_right) done text \Simply rewrite that based on the domain point x.\ lemma has_derivative_inverse_basic_x: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" and "continuous (at (f x)) g" and "g (f x) = x" and "open T" and "f x \ T" and "\y. y \ T \ f (g y) = y" shows "(g has_derivative g') (at (f x))" by (rule has_derivative_inverse_basic) (use assms in auto) text \This is the version in Dieudonne', assuming continuity of f and g.\ lemma has_derivative_inverse_dieudonne: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open S" and "open (f ` S)" and "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g (f x) = x" and "x \ S" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" apply (rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)]) using assms(3-6) unfolding continuous_on_eq_continuous_at[OF assms(1)] continuous_on_eq_continuous_at[OF assms(2)] apply auto done text \Here's the simplest way of not assuming much about g.\ proposition has_derivative_inverse: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "compact S" and "x \ S" and fx: "f x \ interior (f ` S)" and "continuous_on S f" and gf: "\y. y \ S \ g (f y) = y" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" proof - have *: "\y. y \ interior (f ` S) \ f (g y) = y" by (metis gf image_iff interior_subset subsetCE) show ?thesis apply (rule has_derivative_inverse_basic_x[OF assms(6-8), where T = "interior (f ` S)"]) apply (rule continuous_on_interior[OF _ fx]) apply (rule continuous_on_inv) apply (simp_all add: assms *) done qed text \Invertible derivative continuous at a point implies local injectivity. It's only for this we need continuity of the derivative, except of course if we want the fact that the inverse derivative is also continuous. So if we know for some other reason that the inverse function exists, it's OK.\ proposition has_derivative_locally_injective: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "a \ S" and "open S" and bling: "bounded_linear g'" and "g' \ f' a = id" and derf: "\x. x \ S \ (f has_derivative f' x) (at x)" and "\e. e > 0 \ \d>0. \x. dist a x < d \ onorm (\v. f' x v - f' a v) < e" obtains r where "r > 0" "ball a r \ S" "inj_on f (ball a r)" proof - interpret bounded_linear g' using assms by auto note f'g' = assms(4)[unfolded id_def o_def,THEN cong] have "g' (f' a (\Basis)) = (\Basis)" "(\Basis) \ (0::'n)" using f'g' by auto then have *: "0 < onorm g'" unfolding onorm_pos_lt[OF assms(3)] by fastforce define k where "k = 1 / onorm g' / 2" have *: "k > 0" unfolding k_def using * by auto obtain d1 where d1: "0 < d1" "\x. dist a x < d1 \ onorm (\v. f' x v - f' a v) < k" using assms(6) * by blast from \open S\ obtain d2 where "d2 > 0" "ball a d2 \ S" using \a\S\ .. obtain d2 where d2: "0 < d2" "ball a d2 \ S" using \0 < d2\ \ball a d2 \ S\ by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF d1(1) d2(1)] by blast show ?thesis proof show "0 < d" by (fact d) show "ball a d \ S" using \d < d2\ \ball a d2 \ S\ by auto show "inj_on f (ball a d)" unfolding inj_on_def proof (intro strip) fix x y assume as: "x \ ball a d" "y \ ball a d" "f x = f y" define ph where [abs_def]: "ph w = w - g' (f w - f x)" for w have ph':"ph = g' \ (\w. f' a w - (f w - f x))" unfolding ph_def o_def by (simp add: diff f'g') have "norm (ph x - ph y) \ (1 / 2) * norm (x - y)" proof (rule differentiable_bound[OF convex_ball _ _ as(1-2)]) fix u assume u: "u \ ball a d" then have "u \ S" using d d2 by auto have *: "(\v. v - g' (f' u v)) = g' \ (\w. f' a w - f' u w)" unfolding o_def and diff using f'g' by auto have blin: "bounded_linear (f' a)" using \a \ S\ derf by blast show "(ph has_derivative (\v. v - g' (f' u v))) (at u within ball a d)" unfolding ph' * comp_def by (rule \u \ S\ derivative_eq_intros has_derivative_at_withinI [OF derf] bounded_linear.has_derivative [OF blin] bounded_linear.has_derivative [OF bling] |simp)+ have **: "bounded_linear (\x. f' u x - f' a x)" "bounded_linear (\x. f' a x - f' u x)" using \u \ S\ blin bounded_linear_sub derf by auto then have "onorm (\v. v - g' (f' u v)) \ onorm g' * onorm (\w. f' a w - f' u w)" by (simp add: "*" bounded_linear_axioms onorm_compose) also have "\ \ onorm g' * k" apply (rule mult_left_mono) using d1(2)[of u] using onorm_neg[where f="\x. f' u x - f' a x"] d u onorm_pos_le[OF bling] apply (auto simp: algebra_simps) done also have "\ \ 1 / 2" unfolding k_def by auto finally show "onorm (\v. v - g' (f' u v)) \ 1 / 2" . qed moreover have "norm (ph y - ph x) = norm (y - x)" by (simp add: as(3) ph_def) ultimately show "x = y" unfolding norm_minus_commute by auto qed qed qed subsection \Uniformly convergent sequence of derivatives\ lemma has_derivative_sequence_lipschitz_lemma: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\n x h. \n\N; x \ S\ \ norm (f' n x h - g' x h) \ e * norm h" and "0 \ e" shows "\m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof clarify fix m n x y assume as: "N \ m" "N \ n" "x \ S" "y \ S" show "norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof (rule differentiable_bound[where f'="\x h. f' m x h - f' n x h", OF \convex S\ _ _ as(3-4)]) fix x assume "x \ S" show "((\a. f m a - f n a) has_derivative (\h. f' m x h - f' n x h)) (at x within S)" by (rule derivative_intros derf \x\S\)+ show "onorm (\h. f' m x h - f' n x h) \ 2 * e" proof (rule onorm_bound) fix h have "norm (f' m x h - f' n x h) \ norm (f' m x h - g' x h) + norm (f' n x h - g' x h)" using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] by (auto simp add: algebra_simps norm_minus_commute) also have "\ \ e * norm h + e * norm h" using nle[OF \N \ m\ \x \ S\, of h] nle[OF \N \ n\ \x \ S\, of h] by (auto simp add: field_simps) finally show "norm (f' m x h - f' n x h) \ 2 * e * norm h" by auto qed (simp add: \0 \ e\) qed qed lemma has_derivative_sequence_Lipschitz: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "e > 0" shows "\N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" proof - have *: "2 * (e/2) = e" using \e > 0\ by auto obtain N where "\n\N. \x\S. \h. norm (f' n x h - g' x h) \ (e/2) * norm h" using nle \e > 0\ unfolding eventually_sequentially by (metis less_divide_eq_numeral1(1) mult_zero_left) then show "\N. \m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" apply (rule_tac x=N in exI) apply (rule has_derivative_sequence_lipschitz_lemma[where e="e/2", unfolded *]) using assms \e > 0\ apply auto done qed proposition has_derivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "x0 \ S" and lim: "((\n. f n x0) \ l) sequentially" shows "\g. \x\S. (\n. f n x) \ g x \ (g has_derivative g'(x)) (at x within S)" proof - have lem1: "\e. e > 0 \ \N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" using assms(1,2,3) by (rule has_derivative_sequence_Lipschitz) have "\g. \x\S. ((\n. f n x) \ g x) sequentially" proof (intro ballI bchoice) fix x assume "x \ S" show "\y. (\n. f n x) \ y" unfolding convergent_eq_Cauchy proof (cases "x = x0") case True then show "Cauchy (\n. f n x)" using LIMSEQ_imp_Cauchy[OF lim] by auto next case False show "Cauchy (\n. f n x)" unfolding Cauchy_def proof (intro allI impI) fix e :: real assume "e > 0" hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto obtain M where M: "\m\M. \n\M. dist (f m x0) (f n x0) < e / 2" using LIMSEQ_imp_Cauchy[OF lim] * unfolding Cauchy_def by blast obtain N where N: "\m\N. \n\N. \u\S. \y\S. norm (f m u - f n u - (f m y - f n y)) \ e / 2 / norm (x - x0) * norm (u - y)" using lem1 *(2) by blast show "\M. \m\M. \n\M. dist (f m x) (f n x) < e" proof (intro exI allI impI) fix m n assume as: "max M N \m" "max M N\n" have "dist (f m x) (f n x) \ norm (f m x0 - f n x0) + norm (f m x - f n x - (f m x0 - f n x0))" unfolding dist_norm by (rule norm_triangle_sub) also have "\ \ norm (f m x0 - f n x0) + e / 2" using N \x\S\ \x0\S\ as False by fastforce also have "\ < e / 2 + e / 2" by (rule add_strict_right_mono) (use as M in \auto simp: dist_norm\) finally show "dist (f m x) (f n x) < e" by auto qed qed qed qed then obtain g where g: "\x\S. (\n. f n x) \ g x" .. have lem2: "\N. \n\N. \x\S. \y\S. norm ((f n x - f n y) - (g x - g y)) \ e * norm (x - y)" if "e > 0" for e proof - obtain N where N: "\m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" using lem1 \e > 0\ by blast show "\N. \n\N. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" proof (intro exI ballI allI impI) fix n x y assume as: "N \ n" "x \ S" "y \ S" have "((\m. norm (f n x - f n y - (f m x - f m y))) \ norm (f n x - f n y - (g x - g y))) sequentially" by (intro tendsto_intros g[rule_format] as) moreover have "eventually (\m. norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)) sequentially" unfolding eventually_sequentially proof (intro exI allI impI) fix m assume "N \ m" then show "norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)" using N as by (auto simp add: algebra_simps) qed ultimately show "norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" by (simp add: tendsto_upperbound) qed qed have "\x\S. ((\n. f n x) \ g x) sequentially \ (g has_derivative g' x) (at x within S)" unfolding has_derivative_within_alt2 proof (intro ballI conjI allI impI) fix x assume "x \ S" then show "(\n. f n x) \ g x" by (simp add: g) have tog': "(\n. f' n x u) \ g' x u" for u unfolding filterlim_def le_nhds_metric_le eventually_filtermap dist_norm proof (intro allI impI) fix e :: real assume "e > 0" show "eventually (\n. norm (f' n x u - g' x u) \ e) sequentially" proof (cases "u = 0") case True have "eventually (\n. norm (f' n x u - g' x u) \ e * norm u) sequentially" using nle \0 < e\ \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u = 0\ \0 < e\ by (auto elim: eventually_mono) next case False with \0 < e\ have "0 < e / norm u" by simp then have "eventually (\n. norm (f' n x u - g' x u) \ e / norm u * norm u) sequentially" using nle \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u \ 0\ by simp qed qed show "bounded_linear (g' x)" proof fix x' y z :: 'a fix c :: real note lin = assms(2)[rule_format,OF \x\S\,THEN has_derivative_bounded_linear] show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_cmul] apply (intro tendsto_intros tog') done show "g' x (y + z) = g' x y + g' x z" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_add] apply (rule tendsto_add) apply (rule tog')+ done obtain N where N: "\h. norm (f' N x h - g' x h) \ 1 * norm h" using nle \x \ S\ unfolding eventually_sequentially by (fast intro: zero_less_one) have "bounded_linear (f' N x)" using derf \x \ S\ by fast from bounded_linear.bounded [OF this] obtain K where K: "\h. norm (f' N x h) \ norm h * K" .. { fix h have "norm (g' x h) = norm (f' N x h - (f' N x h - g' x h))" by simp also have "\ \ norm (f' N x h) + norm (f' N x h - g' x h)" by (rule norm_triangle_ineq4) also have "\ \ norm h * K + 1 * norm h" using N K by (fast intro: add_mono) finally have "norm (g' x h) \ norm h * (K + 1)" by (simp add: ring_distribs) } then show "\K. \h. norm (g' x h) \ norm h * K" by fast qed show "eventually (\y. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)) (at x within S)" if "e > 0" for e proof - have *: "e / 3 > 0" using that by auto obtain N1 where N1: "\n\N1. \x\S. \h. norm (f' n x h - g' x h) \ e / 3 * norm h" using nle * unfolding eventually_sequentially by blast obtain N2 where N2[rule_format]: "\n\N2. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e / 3 * norm (x - y)" using lem2 * by blast let ?N = "max N1 N2" have "eventually (\y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)) (at x within S)" using derf[unfolded has_derivative_within_alt2] and \x \ S\ and * by fast moreover have "eventually (\y. y \ S) (at x within S)" unfolding eventually_at by (fast intro: zero_less_one) ultimately show "\\<^sub>F y in at x within S. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" proof (rule eventually_elim2) fix y assume "y \ S" assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)" moreover have "norm (g y - g x - (f ?N y - f ?N x)) \ e / 3 * norm (y - x)" using N2[OF _ \y \ S\ \x \ S\] by (simp add: norm_minus_commute) ultimately have "norm (g y - g x - f' ?N x (y - x)) \ 2 * e / 3 * norm (y - x)" using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] by (auto simp add: algebra_simps) moreover have " norm (f' ?N x (y - x) - g' x (y - x)) \ e / 3 * norm (y - x)" using N1 \x \ S\ by auto ultimately show "norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by (auto simp add: algebra_simps) qed qed qed then show ?thesis by fast qed text \Can choose to line up antiderivatives if we want.\ lemma has_antiderivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and der: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and no: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof (cases "S = {}") case False then obtain a where "a \ S" by auto have *: "\P Q. \g. \x\S. P g x \ Q g x \ \g. \x\S. Q g x" by auto show ?thesis apply (rule *) apply (rule has_derivative_sequence [OF \convex S\ _ no, of "\n x. f n x + (f 0 a - f n a)"]) apply (metis assms(2) has_derivative_add_const) using \a \ S\ apply auto done qed auto lemma has_antiderivative_limit: fixes g' :: "'a::real_normed_vector \ 'a \ 'b::banach" assumes "convex S" and "\e. e>0 \ \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm (f' x h - g' x h) \ e * norm h)" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof - have *: "\n. \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm(f' x h - g' x h) \ inverse (real (Suc n)) * norm h)" by (simp add: assms(2)) obtain f where *: "\x. \f'. \xa\S. (f x has_derivative f' xa) (at xa within S) \ (\h. norm (f' xa h - g' xa h) \ inverse (real (Suc x)) * norm h)" using * by metis obtain f' where f': "\x. \z\S. (f x has_derivative f' x z) (at z within S) \ (\h. norm (f' x z h - g' z h) \ inverse (real (Suc x)) * norm h)" using * by metis show ?thesis proof (rule has_antiderivative_sequence[OF \convex S\, of f f']) fix e :: real assume "e > 0" obtain N where N: "inverse (real (Suc N)) < e" using reals_Archimedean[OF \e>0\] .. show "\\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" unfolding eventually_sequentially proof (intro exI allI ballI impI) fix n x h assume n: "N \ n" and x: "x \ S" have *: "inverse (real (Suc n)) \ e" apply (rule order_trans[OF _ N[THEN less_imp_le]]) using n apply (auto simp add: field_simps) done show "norm (f' n x h - g' x h) \ e * norm h" by (meson "*" mult_right_mono norm_ge_zero order.trans x f') qed qed (use f' in auto) qed subsection \Differentiation of a series\ proposition has_derivative_series: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and "\e. e>0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (sum (\i. f' i x h) {.. e * norm h" and "x \ S" and "(\n. f n x) sums l" shows "\g. \x\S. (\n. f n x) sums (g x) \ (g has_derivative g' x) (at x within S)" unfolding sums_def apply (rule has_derivative_sequence[OF assms(1) _ assms(3)]) apply (metis assms(2) has_derivative_sum) using assms(4-5) unfolding sums_def apply auto done lemma has_field_derivative_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniform_limit S (\n x. \i S" "summable (\n. f n x0)" shows "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" unfolding has_field_derivative_def proof (rule has_derivative_series) show "\\<^sub>F n in sequentially. \x\S. \h. norm ((\i e * norm h" if "e > 0" for e unfolding eventually_sequentially proof - from that assms(3) obtain N where N: "\n x. n \ N \ x \ S \ norm ((\i N" "x \ S" have "norm ((\iii e" by simp hence "norm ((\i e * norm h" by (intro mult_right_mono) simp_all finally have "norm ((\i e * norm h" . } thus "\N. \n\N. \x\S. \h. norm ((\i e * norm h" by blast qed qed (use assms in \auto simp: has_field_derivative_def\) lemma has_field_derivative_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" "x \ interior S" shows "summable (\n. f n x)" "((\x. \n. f n x) has_field_derivative (\n. f' n x)) (at x)" proof - from \x \ interior S\ have "x \ S" using interior_subset by blast define g' where [abs_def]: "g' x = (\i. f' i x)" for x from assms(3) have "uniform_limit S (\n x. \ix. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g(1)[OF \x \ S\] show "summable (\n. f n x)" by (simp add: sums_iff) from g(2)[OF \x \ S\] \x \ interior S\ have "(g has_field_derivative g' x) (at x)" by (simp add: at_within_interior[of x S]) also have "(g has_field_derivative g' x) (at x) \ ((\x. \n. f n x) has_field_derivative g' x) (at x)" using eventually_nhds_in_nhd[OF \x \ interior S\] interior_subset[of S] g(1) by (intro DERIV_cong_ev) (auto elim!: eventually_mono simp: sums_iff) finally show "((\x. \n. f n x) has_field_derivative g' x) (at x)" . qed lemma differentiable_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" and x: "x \ S" shows "summable (\n. f n x)" and "(\x. \n. f n x) differentiable (at x)" proof - from assms(4) obtain g' where A: "uniform_limit S (\n x. \iopen S\ have S: "at x within S = at x" by (rule at_within_open) have "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g[OF x] show "summable (\n. f n x)" by (auto simp: summable_def) from g(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)" by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative (*) (g' x)) (at x)" by (rule has_derivative_transform_within_open[OF g' \open S\ x]) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def differentiable_def has_field_derivative_def) qed lemma differentiable_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" shows "(\x. \n. f n x) differentiable (at x0)" using differentiable_series[OF assms, of x0] \x0 \ S\ by blast+ subsection \Derivative as a vector\ text \Considering derivative \<^typ>\real \ 'b::real_normed_vector\ as a vector.\ definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)" lemma vector_derivative_unique_within: assumes not_bot: "at x within S \ bot" and f': "(f has_vector_derivative f') (at x within S)" and f'': "(f has_vector_derivative f'') (at x within S)" shows "f' = f''" proof - have "(\x. x *\<^sub>R f') = (\x. x *\<^sub>R f'')" proof (rule frechet_derivative_unique_within, simp_all) show "\d. d \ 0 \ \d\ < e \ x + d \ S" if "0 < e" for e proof - from that obtain x' where "x' \ S" "x' \ x" "\x' - x\ < e" using islimpt_approachable_real[of x S] not_bot by (auto simp add: trivial_limit_within) then show ?thesis using eq_iff_diff_eq_0 by fastforce qed qed (use f' f'' in \auto simp: has_vector_derivative_def\) then show ?thesis unfolding fun_eq_iff by (metis scaleR_one) qed lemma vector_derivative_unique_at: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f'') (at x) \ f' = f''" by (rule vector_derivative_unique_within) auto lemma differentiableI_vector: "(f has_vector_derivative y) F \ f differentiable F" by (auto simp: differentiable_def has_vector_derivative_def) proposition vector_derivative_works: "f differentiable net \ (f has_vector_derivative (vector_derivative f net)) net" (is "?l = ?r") proof assume ?l obtain f' where f': "(f has_derivative f') net" using \?l\ unfolding differentiable_def .. then interpret bounded_linear f' by auto show ?r unfolding vector_derivative_def has_vector_derivative_def by (rule someI[of _ "f' 1"]) (simp add: scaleR[symmetric] f') qed (auto simp: vector_derivative_def has_vector_derivative_def differentiable_def) lemma vector_derivative_within: assumes not_bot: "at x within S \ bot" and y: "(f has_vector_derivative y) (at x within S)" shows "vector_derivative f (at x within S) = y" using y by (intro vector_derivative_unique_within[OF not_bot vector_derivative_works[THEN iffD1] y]) (auto simp: differentiable_def has_vector_derivative_def) lemma frechet_derivative_eq_vector_derivative: assumes "f differentiable (at x)" shows "(frechet_derivative f (at x)) = (\r. r *\<^sub>R vector_derivative f (at x))" using assms by (auto simp: differentiable_iff_scaleR vector_derivative_def has_vector_derivative_def intro: someI frechet_derivative_at [symmetric]) lemma has_real_derivative: fixes f :: "real \ real" assumes "(f has_derivative f') F" obtains c where "(f has_real_derivative c) F" proof - obtain c where "f' = (\x. x * c)" by (metis assms has_derivative_bounded_linear real_bounded_linear) then show ?thesis by (metis assms that has_field_derivative_def mult_commute_abs) qed lemma has_real_derivative_iff: fixes f :: "real \ real" shows "(\c. (f has_real_derivative c) F) = (\D. (f has_derivative D) F)" by (metis has_field_derivative_def has_real_derivative) lemma has_vector_derivative_cong_ev: assumes *: "eventually (\x. x \ S \ f x = g x) (nhds x)" "f x = g x" shows "(f has_vector_derivative f') (at x within S) = (g has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def has_derivative_def using * apply (cases "at x within S \ bot") apply (intro refl conj_cong filterlim_cong) apply (auto simp: Lim_ident_at eventually_at_filter elim: eventually_mono) done lemma islimpt_closure_open: fixes s :: "'a::perfect_space set" assumes "open s" and t: "t = closure s" "x \ t" shows "x islimpt t" proof cases assume "x \ s" { fix T assume "x \ T" "open T" then have "open (s \ T)" using \open s\ by auto then have "s \ T \ {x}" using not_open_singleton[of x] by auto with \x \ T\ \x \ s\ have "\y\t. y \ T \ y \ x" using closure_subset[of s] by (auto simp: t) } then show ?thesis by (auto intro!: islimptI) next assume "x \ s" with t show ?thesis unfolding t closure_def by (auto intro: islimpt_subset) qed lemma vector_derivative_unique_within_closed_interval: assumes ab: "a < b" "x \ cbox a b" assumes D: "(f has_vector_derivative f') (at x within cbox a b)" "(f has_vector_derivative f'') (at x within cbox a b)" shows "f' = f''" using ab by (intro vector_derivative_unique_within[OF _ D]) (auto simp: trivial_limit_within intro!: islimpt_closure_open[where s="{a <..< b}"]) lemma vector_derivative_at: "(f has_vector_derivative f') (at x) \ vector_derivative f (at x) = f'" by (intro vector_derivative_within at_neq_bot) lemma has_vector_derivative_id_at [simp]: "vector_derivative (\x. x) (at a) = 1" by (simp add: vector_derivative_at) lemma vector_derivative_minus_at [simp]: "f differentiable at a \ vector_derivative (\x. - f x) (at a) = - vector_derivative f (at a)" by (simp add: vector_derivative_at has_vector_derivative_minus vector_derivative_works [symmetric]) lemma vector_derivative_add_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x + g x) (at a) = vector_derivative f (at a) + vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_add vector_derivative_works [symmetric]) lemma vector_derivative_diff_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x - g x) (at a) = vector_derivative f (at a) - vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_diff vector_derivative_works [symmetric]) lemma vector_derivative_mult_at [simp]: fixes f g :: "real \ 'a :: real_normed_algebra" shows "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x * g x) (at a) = f a * vector_derivative g (at a) + vector_derivative f (at a) * g a" by (simp add: vector_derivative_at has_vector_derivative_mult vector_derivative_works [symmetric]) lemma vector_derivative_scaleR_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x *\<^sub>R g x) (at a) = f a *\<^sub>R vector_derivative g (at a) + vector_derivative f (at a) *\<^sub>R g a" apply (rule vector_derivative_at) apply (rule has_vector_derivative_scaleR) apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs) done lemma vector_derivative_within_cbox: assumes ab: "a < b" "x \ cbox a b" assumes f: "(f has_vector_derivative f') (at x within cbox a b)" shows "vector_derivative f (at x within cbox a b) = f'" by (intro vector_derivative_unique_within_closed_interval[OF ab _ f] vector_derivative_works[THEN iffD1] differentiableI_vector) fact lemma vector_derivative_within_closed_interval: fixes f::"real \ 'a::euclidean_space" assumes "a < b" and "x \ {a..b}" assumes "(f has_vector_derivative f') (at x within {a..b})" shows "vector_derivative f (at x within {a..b}) = f'" using assms vector_derivative_within_cbox by fastforce lemma has_vector_derivative_within_subset: "(f has_vector_derivative f') (at x within S) \ T \ S \ (f has_vector_derivative f') (at x within T)" by (auto simp: has_vector_derivative_def intro: has_derivative_subset) lemma has_vector_derivative_at_within: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def by (rule has_derivative_at_withinI) lemma has_vector_derivative_weaken: fixes x D and f g S T assumes f: "(f has_vector_derivative D) (at x within T)" and "x \ S" "S \ T" and "\x. x \ S \ f x = g x" shows "(g has_vector_derivative D) (at x within S)" proof - have "(f has_vector_derivative D) (at x within S) \ (g has_vector_derivative D) (at x within S)" unfolding has_vector_derivative_def has_derivative_iff_norm using assms by (intro conj_cong Lim_cong_within refl) auto then show ?thesis using has_vector_derivative_within_subset[OF f \S \ T\] by simp qed lemma has_vector_derivative_transform_within: assumes "(f has_vector_derivative f') (at x within S)" and "0 < d" and "x \ S" and "\x'. \x'\S; dist x' x < d\ \ f x' = g x'" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within) lemma has_vector_derivative_transform_within_open: assumes "(f has_vector_derivative f') (at x)" and "open S" and "x \ S" and "\y. y\S \ f y = g y" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within_open) lemma has_vector_derivative_transform: assumes "x \ S" "\x. x \ S \ g x = f x" assumes f': "(f has_vector_derivative f') (at x within S)" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform) lemma vector_diff_chain_at: assumes "(f has_vector_derivative f') (at x)" and "(g has_vector_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x)" using assms has_vector_derivative_at_within has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_diff_chain_within: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at (f x) within f ` s)" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)" using assms has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_derivative_const_at [simp]: "vector_derivative (\x. c) (at a) = 0" by (simp add: vector_derivative_at) lemma vector_derivative_at_within_ivl: "(f has_vector_derivative f') (at x) \ a \ x \ x \ b \ a vector_derivative f (at x within {a..b}) = f'" using has_vector_derivative_at_within vector_derivative_within_cbox by fastforce lemma vector_derivative_chain_at: assumes "f differentiable at x" "(g differentiable at (f x))" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" by (metis vector_diff_chain_at vector_derivative_at vector_derivative_works assms) lemma field_vector_diff_chain_at: (*thanks to Wenda Li*) assumes Df: "(f has_vector_derivative f') (at x)" and Dg: "(g has_field_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' * g')) (at x)" using diff_chain_at[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_chain_within: assumes "at x within S \ bot" "f differentiable (at x within S)" "(g has_derivative g') (at (f x) within f ` S)" shows "vector_derivative (g \ f) (at x within S) = g' (vector_derivative f (at x within S)) " apply (rule vector_derivative_within [OF \at x within S \ bot\]) apply (rule vector_derivative_diff_chain_within) using assms(2-3) vector_derivative_works by auto subsection \Field differentiability\ definition\<^marker>\tag important\ field_differentiable :: "['a \ 'a::real_normed_field, 'a filter] \ bool" (infixr "(field'_differentiable)" 50) where "f field_differentiable F \ \f'. (f has_field_derivative f') F" lemma field_differentiable_imp_differentiable: "f field_differentiable F \ f differentiable F" unfolding field_differentiable_def differentiable_def using has_field_derivative_imp_has_derivative by auto lemma field_differentiable_imp_continuous_at: "f field_differentiable (at x within S) \ continuous (at x within S) f" by (metis DERIV_continuous field_differentiable_def) lemma field_differentiable_within_subset: "\f field_differentiable (at x within S); T \ S\ \ f field_differentiable (at x within T)" by (metis DERIV_subset field_differentiable_def) lemma field_differentiable_at_within: "\f field_differentiable (at x)\ \ f field_differentiable (at x within S)" unfolding field_differentiable_def by (metis DERIV_subset top_greatest) lemma field_differentiable_linear [simp,derivative_intros]: "((*) c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def mult_commute_abs by (force intro: has_derivative_mult_right) lemma field_differentiable_const [simp,derivative_intros]: "(\z. c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_const has_field_derivative_imp_has_derivative by blast lemma field_differentiable_ident [simp,derivative_intros]: "(\z. z) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_ident has_field_derivative_def by blast lemma field_differentiable_id [simp,derivative_intros]: "id field_differentiable F" unfolding id_def by (rule field_differentiable_ident) lemma field_differentiable_minus [derivative_intros]: "f field_differentiable F \ (\z. - (f z)) field_differentiable F" unfolding field_differentiable_def by (metis field_differentiable_minus) lemma field_differentiable_add [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z + g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_add) lemma field_differentiable_add_const [simp,derivative_intros]: "(+) c field_differentiable F" by (simp add: field_differentiable_add) lemma field_differentiable_sum [derivative_intros]: "(\i. i \ I \ (f i) field_differentiable F) \ (\z. \i\I. f i z) field_differentiable F" by (induct I rule: infinite_finite_induct) (auto intro: field_differentiable_add field_differentiable_const) lemma field_differentiable_diff [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z - g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_diff) lemma field_differentiable_inverse [derivative_intros]: assumes "f field_differentiable (at a within S)" "f a \ 0" shows "(\z. inverse (f z)) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_inverse_fun) lemma field_differentiable_mult [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" shows "(\z. f z * g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_mult [of f _ a S g]) lemma field_differentiable_divide [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" "g a \ 0" shows "(\z. f z / g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_divide [of f _ a S g]) lemma field_differentiable_power [derivative_intros]: assumes "f field_differentiable (at a within S)" shows "(\z. f z ^ n) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_power) lemma field_differentiable_transform_within: "0 < d \ x \ S \ (\x'. x' \ S \ dist x' x < d \ f x' = g x') \ f field_differentiable (at x within S) \ g field_differentiable (at x within S)" unfolding field_differentiable_def has_field_derivative_def by (blast intro: has_derivative_transform_within) lemma field_differentiable_compose_within: assumes "f field_differentiable (at a within S)" "g field_differentiable (at (f a) within f`S)" shows "(g o f) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_image_chain) lemma field_differentiable_compose: "f field_differentiable at z \ g field_differentiable at (f z) \ (g o f) field_differentiable at z" by (metis field_differentiable_at_within field_differentiable_compose_within) lemma field_differentiable_within_open: "\a \ S; open S\ \ f field_differentiable at a within S \ f field_differentiable at a" unfolding field_differentiable_def by (metis at_within_open) lemma exp_scaleR_has_vector_derivative_right: "((\t. exp (t *\<^sub>R A)) has_vector_derivative exp (t *\<^sub>R A) * A) (at t within T)" unfolding has_vector_derivative_def proof (rule has_derivativeI) let ?F = "at t within (T \ {t - 1 <..< t + 1})" have *: "at t within T = ?F" by (rule at_within_nhd[where S="{t - 1 <..< t + 1}"]) auto let ?e = "\i x. (inverse (1 + real i) * inverse (fact i) * (x - t) ^ i) *\<^sub>R (A * A ^ i)" have "\\<^sub>F n in sequentially. \x\T \ {t - 1<.. norm (A ^ (n + 1) /\<^sub>R fact (n + 1))" apply (auto simp: algebra_split_simps intro!: eventuallyI) apply (rule mult_left_mono) apply (auto simp add: field_simps power_abs intro!: divide_right_mono power_le_one) done then have "uniform_limit (T \ {t - 1<..n x. \ix. \i. ?e i x) sequentially" by (rule Weierstrass_m_test_ev) (intro summable_ignore_initial_segment summable_norm_exp) moreover have "\\<^sub>F x in sequentially. x > 0" by (metis eventually_gt_at_top) then have "\\<^sub>F n in sequentially. ((\x. \i A) ?F" by eventually_elim (auto intro!: tendsto_eq_intros simp: power_0_left if_distrib if_distribR cong: if_cong) ultimately have [tendsto_intros]: "((\x. \i. ?e i x) \ A) ?F" by (auto intro!: swap_uniform_limit[where f="\n x. \i < n. ?e i x" and F = sequentially]) have [tendsto_intros]: "((\x. if x = t then 0 else 1) \ 1) ?F" by (rule tendsto_eventually) (simp add: eventually_at_filter) have "((\y. ((y - t) / abs (y - t)) *\<^sub>R ((\n. ?e n y) - A)) \ 0) (at t within T)" unfolding * by (rule tendsto_norm_zero_cancel) (auto intro!: tendsto_eq_intros) moreover have "\\<^sub>F x in at t within T. x \ t" by (simp add: eventually_at_filter) then have "\\<^sub>F x in at t within T. ((x - t) / \x - t\) *\<^sub>R ((\n. ?e n x) - A) = (exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" proof eventually_elim case (elim x) have "(exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = ((\n. (x - t) *\<^sub>R ?e n x) - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" unfolding exp_first_term by (simp add: ac_simps) also have "summable (\n. ?e n x)" proof - from elim have "?e n x = (((x - t) *\<^sub>R A) ^ (n + 1)) /\<^sub>R fact (n + 1) /\<^sub>R (x - t)" for n by simp then show ?thesis by (auto simp only: intro!: summable_scaleR_right summable_ignore_initial_segment summable_exp_generic) qed then have "(\n. (x - t) *\<^sub>R ?e n x) = (x - t) *\<^sub>R (\n. ?e n x)" by (rule suminf_scaleR_right[symmetric]) also have "(\ - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = (x - t) *\<^sub>R ((\n. ?e n x) - A) /\<^sub>R norm (x - t)" by (simp add: algebra_simps) finally show ?case by simp (simp add: field_simps) qed ultimately have "((\y. (exp ((y - t) *\<^sub>R A) - 1 - (y - t) *\<^sub>R A) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) from tendsto_mult_right_zero[OF this, where c="exp (t *\<^sub>R A)"] show "((\y. (exp (y *\<^sub>R A) - exp (t *\<^sub>R A) - (y - t) *\<^sub>R (exp (t *\<^sub>R A) * A)) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) (auto simp: field_split_simps exp_add_commuting[symmetric]) qed (rule bounded_linear_scaleR_left) lemma exp_times_scaleR_commute: "exp (t *\<^sub>R A) * A = A * exp (t *\<^sub>R A)" using exp_times_arg_commute[symmetric, of "t *\<^sub>R A"] by (auto simp: algebra_simps) lemma exp_scaleR_has_vector_derivative_left: "((\t. exp (t *\<^sub>R A)) has_vector_derivative A * exp (t *\<^sub>R A)) (at t)" using exp_scaleR_has_vector_derivative_right[of A t] by (simp add: exp_times_scaleR_commute) lemma field_differentiable_series: fixes f :: "nat \ 'a::{real_normed_field,banach} \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" and x: "x \ S" shows "(\x. \n. f n x) field_differentiable (at x)" proof - from assms(4) obtain g' where A: "uniform_limit S (\n x. \iopen S\ have S: "at x within S = at x" by (rule at_within_open) have "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)" by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative (*) (g' x)) (at x)" by (rule has_derivative_transform_within_open[OF g' \open S\ x]) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) field_differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def field_differentiable_def has_field_derivative_def) qed subsubsection\<^marker>\tag unimportant\\Caratheodory characterization\ lemma field_differentiable_caratheodory_at: "f field_differentiable (at z) \ (\g. (\w. f(w) - f(z) = g(w) * (w - z)) \ continuous (at z) g)" using CARAT_DERIV [of f] by (simp add: field_differentiable_def has_field_derivative_def) lemma field_differentiable_caratheodory_within: "f field_differentiable (at z within s) \ (\g. (\w. f(w) - f(z) = g(w) * (w - z)) \ continuous (at z within s) g)" using DERIV_caratheodory_within [of f] by (simp add: field_differentiable_def has_field_derivative_def) subsection \Field derivative\ definition\<^marker>\tag important\ deriv :: "('a \ 'a::real_normed_field) \ 'a \ 'a" where "deriv f x \ SOME D. DERIV f x :> D" lemma DERIV_imp_deriv: "DERIV f x :> f' \ deriv f x = f'" unfolding deriv_def by (metis some_equality DERIV_unique) lemma DERIV_deriv_iff_has_field_derivative: "DERIV f x :> deriv f x \ (\f'. (f has_field_derivative f') (at x))" by (auto simp: has_field_derivative_def DERIV_imp_deriv) lemma DERIV_deriv_iff_real_differentiable: fixes x :: real shows "DERIV f x :> deriv f x \ f differentiable at x" unfolding differentiable_def by (metis DERIV_imp_deriv has_real_derivative_iff) lemma deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "deriv f x = deriv g y" proof - have "(\D. (f has_field_derivative D) (at x)) = (\D. (g has_field_derivative D) (at y))" by (intro ext DERIV_cong_ev refl assms) thus ?thesis by (simp add: deriv_def assms) qed lemma higher_deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "(deriv ^^ n) f x = (deriv ^^ n) g y" proof - from assms(1) have "eventually (\x. (deriv ^^ n) f x = (deriv ^^ n) g x) (nhds x)" proof (induction n arbitrary: f g) case (Suc n) from Suc.prems have "eventually (\y. eventually (\z. f z = g z) (nhds y)) (nhds x)" by (simp add: eventually_eventually) hence "eventually (\x. deriv f x = deriv g x) (nhds x)" by eventually_elim (rule deriv_cong_ev, simp_all) thus ?case by (auto intro!: deriv_cong_ev Suc simp: funpow_Suc_right simp del: funpow.simps) qed auto from eventually_nhds_x_imp_x[OF this] assms(2) show ?thesis by simp qed lemma real_derivative_chain: fixes x :: real shows "f differentiable at x \ g differentiable at (f x) \ deriv (g o f) x = deriv g (f x) * deriv f x" by (metis DERIV_deriv_iff_real_differentiable DERIV_chain DERIV_imp_deriv) lemma field_derivative_eq_vector_derivative: "(deriv f x) = vector_derivative f (at x)" by (simp add: mult.commute deriv_def vector_derivative_def has_vector_derivative_def has_field_derivative_def) proposition field_differentiable_derivI: "f field_differentiable (at x) \ (f has_field_derivative deriv f x) (at x)" by (simp add: field_differentiable_def DERIV_deriv_iff_has_field_derivative) lemma vector_derivative_chain_at_general: assumes "f differentiable at x" "g field_differentiable at (f x)" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) * deriv g (f x)" apply (rule vector_derivative_at [OF field_vector_diff_chain_at]) using assms vector_derivative_works by (auto simp: field_differentiable_derivI) lemma DERIV_deriv_iff_field_differentiable: "DERIV f x :> deriv f x \ f field_differentiable at x" unfolding field_differentiable_def by (metis DERIV_imp_deriv) lemma deriv_chain: "f field_differentiable at x \ g field_differentiable at (f x) \ deriv (g o f) x = deriv g (f x) * deriv f x" by (metis DERIV_deriv_iff_field_differentiable DERIV_chain DERIV_imp_deriv) lemma deriv_linear [simp]: "deriv (\w. c * w) = (\z. c)" by (metis DERIV_imp_deriv DERIV_cmult_Id) lemma deriv_uminus [simp]: "deriv (\w. -w) = (\z. -1)" using deriv_linear[of "-1"] by (simp del: deriv_linear) lemma deriv_ident [simp]: "deriv (\w. w) = (\z. 1)" by (metis DERIV_imp_deriv DERIV_ident) lemma deriv_id [simp]: "deriv id = (\z. 1)" by (simp add: id_def) lemma deriv_const [simp]: "deriv (\w. c) = (\z. 0)" by (metis DERIV_imp_deriv DERIV_const) lemma deriv_add [simp]: "\f field_differentiable at z; g field_differentiable at z\ \ deriv (\w. f w + g w) z = deriv f z + deriv g z" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_intros) lemma deriv_diff [simp]: "\f field_differentiable at z; g field_differentiable at z\ \ deriv (\w. f w - g w) z = deriv f z - deriv g z" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_intros) lemma deriv_mult [simp]: "\f field_differentiable at z; g field_differentiable at z\ \ deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_eq_intros) lemma deriv_cmult: "f field_differentiable at z \ deriv (\w. c * f w) z = c * deriv f z" by simp lemma deriv_cmult_right: "f field_differentiable at z \ deriv (\w. f w * c) z = deriv f z * c" by simp lemma deriv_inverse [simp]: "\f field_differentiable at z; f z \ 0\ \ deriv (\w. inverse (f w)) z = - deriv f z / f z ^ 2" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (safe intro!: DERIV_imp_deriv derivative_eq_intros) (auto simp: field_split_simps power2_eq_square) lemma deriv_divide [simp]: "\f field_differentiable at z; g field_differentiable at z; g z \ 0\ \ deriv (\w. f w / g w) z = (deriv f z * g z - f z * deriv g z) / g z ^ 2" by (simp add: field_class.field_divide_inverse field_differentiable_inverse) (simp add: field_split_simps power2_eq_square) lemma deriv_cdivide_right: "f field_differentiable at z \ deriv (\w. f w / c) z = deriv f z / c" by (simp add: field_class.field_divide_inverse) lemma deriv_compose_linear: "f field_differentiable at (c * z) \ deriv (\w. f (c * w)) z = c * deriv f (c * z)" apply (rule DERIV_imp_deriv) unfolding DERIV_deriv_iff_field_differentiable [symmetric] by (metis (full_types) DERIV_chain2 DERIV_cmult_Id mult.commute) lemma nonzero_deriv_nonconstant: assumes df: "DERIV f \ :> df" and S: "open S" "\ \ S" and "df \ 0" shows "\ f constant_on S" unfolding constant_on_def by (metis \df \ 0\ has_field_derivative_transform_within_open [OF df S] DERIV_const DERIV_unique) subsection \Relation between convexity and derivative\ (* TODO: Generalise to real vector spaces? *) proposition convex_on_imp_above_tangent: assumes convex: "convex_on A f" and connected: "connected A" assumes c: "c \ interior A" and x : "x \ A" assumes deriv: "(f has_field_derivative f') (at c within A)" shows "f x - f c \ f' * (x - c)" proof (cases x c rule: linorder_cases) assume xc: "x > c" let ?A' = "interior A \ {c<..}" from c have "c \ interior A \ closure {c<..}" by auto also have "\ \ closure (interior A \ {c<..})" by (intro open_Int_closure_subset) auto finally have "at c within ?A' \ bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_right_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_right c)" proof eventually_elim fix y assume y: "y \ {c<.. (f x - f c) / (x - c) * (y - c) + f c" using interior_subset[of A] by (intro convex_onD_Icc' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) / (x - c) * (y - c)" by simp thus "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_upperbound) thus ?thesis using xc by (simp add: field_simps) next assume xc: "x < c" let ?A' = "interior A \ {.. interior A \ closure {.. \ closure (interior A \ {.. bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_left_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_left c)" proof eventually_elim fix y assume y: "y \ {x<.. (f x - f c) / (c - x) * (c - y) + f c" using interior_subset[of A] by (intro convex_onD_Icc'' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) * ((c - y) / (c - x))" by simp also have "(c - y) / (c - x) = (y - c) / (x - c)" using y xc by (simp add: field_simps) finally show "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_lowerbound) thus ?thesis using xc by (simp add: field_simps) qed simp_all subsection \Partial derivatives\ lemma eventually_at_Pair_within_TimesI1: fixes x::"'a::metric_space" assumes "\\<^sub>F x' in at x within X. P x'" assumes "P x" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "x \ S" "\x'. x' \ X \ x' \ S \ P x'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" unfolding eventually_at_topological by (auto intro!: exI[where x="S \ UNIV"] S open_Times) qed lemma eventually_at_Pair_within_TimesI2: fixes x::"'a::metric_space" assumes "\\<^sub>F y' in at y within Y. P y'" "P y" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "y \ S" "\y'. y' \ Y \ y' \ S \ P y'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" unfolding eventually_at_topological by (auto intro!: exI[where x="UNIV \ S"] S open_Times) qed proposition has_derivative_partialsI: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector \ 'c::real_normed_vector" assumes fx: "((\x. f x y) has_derivative fx) (at x within X)" assumes fy: "\x y. x \ X \ y \ Y \ ((\y. f x y) has_derivative blinfun_apply (fy x y)) (at y within Y)" assumes fy_cont[unfolded continuous_within]: "continuous (at (x, y) within X \ Y) (\(x, y). fy x y)" assumes "y \ Y" "convex Y" shows "((\(x, y). f x y) has_derivative (\(tx, ty). fx tx + fy x y ty)) (at (x, y) within X \ Y)" proof (safe intro!: has_derivativeI tendstoI, goal_cases) case (2 e') interpret fx: bounded_linear "fx" using fx by (rule has_derivative_bounded_linear) define e where "e = e' / 9" have "e > 0" using \e' > 0\ by (simp add: e_def) from fy_cont[THEN tendstoD, OF \e > 0\] have "\\<^sub>F (x', y') in at (x, y) within X \ Y. dist (fy x' y') (fy x y) < e" by (auto simp: split_beta') from this[unfolded eventually_at] obtain d' where "d' > 0" "\x' y'. x' \ X \ y' \ Y \ (x', y') \ (x, y) \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" by auto then have d': "x' \ X \ y' \ Y \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" for x' y' using \0 < e\ by (cases "(x', y') = (x, y)") auto define d where "d = d' / sqrt 2" have "d > 0" using \0 < d'\ by (simp add: d_def) have d: "x' \ X \ y' \ Y \ dist x' x < d \ dist y' y < d \ dist (fy x' y') (fy x y) < e" for x' y' by (auto simp: dist_prod_def d_def intro!: d' real_sqrt_sum_squares_less) let ?S = "ball y d \ Y" have "convex ?S" by (auto intro!: convex_Int \convex Y\) { fix x'::'a and y'::'b assume x': "x' \ X" and y': "y' \ Y" assume dx': "dist x' x < d" and dy': "dist y' y < d" have "norm (fy x' y' - fy x' y) \ dist (fy x' y') (fy x y) + dist (fy x' y) (fy x y)" by norm also have "dist (fy x' y') (fy x y) < e" by (rule d; fact) also have "dist (fy x' y) (fy x y) < e" by (auto intro!: d simp: dist_prod_def x' \d > 0\ \y \ Y\ dx') finally have "norm (fy x' y' - fy x' y) < e + e" by arith then have "onorm (blinfun_apply (fy x' y') - blinfun_apply (fy x' y)) < e + e" by (auto simp: norm_blinfun.rep_eq blinfun.diff_left[abs_def] fun_diff_def) } note onorm = this have ev_mem: "\\<^sub>F (x', y') in at (x, y) within X \ Y. (x', y') \ X \ Y" using \y \ Y\ by (auto simp: eventually_at intro!: zero_less_one) moreover have ev_dist: "\\<^sub>F xy in at (x, y) within X \ Y. dist xy (x, y) < d" if "d > 0" for d using eventually_at_ball[OF that] by (rule eventually_elim2) (auto simp: dist_commute intro!: eventually_True) note ev_dist[OF \0 < d\] ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" proof (eventually_elim, safe) fix x' y' assume "x' \ X" and y': "y' \ Y" assume dist: "dist (x', y') (x, y) < d" then have dx: "dist x' x < d" and dy: "dist y' y < d" unfolding dist_prod_def fst_conv snd_conv atomize_conj by (metis le_less_trans real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) { fix t::real assume "t \ {0 .. 1}" then have "y + t *\<^sub>R (y' - y) \ closed_segment y y'" by (auto simp: closed_segment_def algebra_simps intro!: exI[where x=t]) also have "\ \ ball y d \ Y" using \y \ Y\ \0 < d\ dy y' by (intro \convex ?S\[unfolded convex_contains_segment, rule_format, of y y']) (auto simp: dist_commute) finally have "y + t *\<^sub>R (y' - y) \ ?S" . } note seg = this have "\x. x \ ball y d \ Y \ onorm (blinfun_apply (fy x' x) - blinfun_apply (fy x' y)) \ e + e" by (safe intro!: onorm less_imp_le \x' \ X\ dx) (auto simp: dist_commute \0 < d\ \y \ Y\) with seg has_derivative_subset[OF assms(2)[OF \x' \ X\]] show "norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" by (rule differentiable_bound_linearization[where S="?S"]) (auto intro!: \0 < d\ \y \ Y\) qed moreover let ?le = "\x'. norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" from fx[unfolded has_derivative_within, THEN conjunct2, THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. ?le x'" by eventually_elim (simp, simp add: dist_norm field_split_simps split: if_split_asm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. ?le x'" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps) moreover have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((x', y') - (x, y)) \ 0" unfolding norm_eq_zero right_minus_eq by (auto simp: eventually_at intro!: zero_less_one) moreover from fy_cont[THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. norm (fy x' y - fy x y) < e" unfolding eventually_at using \y \ Y\ by (auto simp: dist_prod_def dist_norm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (fy x' y - fy x y) < e" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps \0 < e\) ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" apply eventually_elim proof safe fix x' y' have "norm (f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) \ norm (f x' y' - f x' y - fy x' y (y' - y)) + norm (fy x y (y' - y) - fy x' y (y' - y)) + norm (f x' y - f x y - fx (x' - x))" by norm also assume nz: "norm ((x', y') - (x, y)) \ 0" and nfy: "norm (fy x' y - fy x y) < e" assume "norm (f x' y' - f x' y - blinfun_apply (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" also assume "norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" also have "norm ((fy x y) (y' - y) - (fy x' y) (y' - y)) \ norm ((fy x y) - (fy x' y)) * norm (y' - y)" by (auto simp: blinfun.bilinear_simps[symmetric] intro!: norm_blinfun) also have "\ \ (e + e) * norm (y' - y)" using \e > 0\ nfy by (auto simp: norm_minus_commute intro!: mult_right_mono) also have "norm (x' - x) * e \ norm (x' - x) * (e + e)" using \0 < e\ by simp also have "norm (y' - y) * (e + e) + (e + e) * norm (y' - y) + norm (x' - x) * (e + e) \ (norm (y' - y) + norm (x' - x)) * (4 * e)" using \e > 0\ by (simp add: algebra_simps) also have "\ \ 2 * norm ((x', y') - (x, y)) * (4 * e)" using \0 < e\ real_sqrt_sum_squares_ge1[of "norm (x' - x)" "norm (y' - y)"] real_sqrt_sum_squares_ge2[of "norm (y' - y)" "norm (x' - x)"] by (auto intro!: mult_right_mono simp: norm_prod_def simp del: real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) also have "\ \ norm ((x', y') - (x, y)) * (8 * e)" by simp also have "\ < norm ((x', y') - (x, y)) * e'" using \0 < e'\ nz by (auto simp: e_def) finally show "norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" by (simp add: dist_norm) (auto simp add: field_split_simps) qed then show ?case by eventually_elim (auto simp: dist_norm field_simps) next from has_derivative_bounded_linear[OF fx] obtain fxb where "fx = blinfun_apply fxb" by (metis bounded_linear_Blinfun_apply) then show "bounded_linear (\(tx, ty). fx tx + blinfun_apply (fy x y) ty)" by (auto intro!: bounded_linear_intros simp: split_beta') qed subsection\<^marker>\tag unimportant\ \Differentiable case distinction\ lemma has_derivative_within_If_eq: "((\x. if P x then f x else g x) has_derivative f') (at x within S) = (bounded_linear f' \ ((\y.(if P y then (f y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x) else (g y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x))) \ 0) (at x within S))" (is "_ = (_ \ (?if \ 0) _)") proof - have "(\y. (1 / norm (y - x)) *\<^sub>R ((if P y then f y else g y) - ((if P x then f x else g x) + f' (y - x)))) = ?if" by (auto simp: inverse_eq_divide) thus ?thesis by (auto simp: has_derivative_within) qed lemma has_derivative_If_within_closures: assumes f': "x \ S \ (closure S \ closure T) \ (f has_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" assumes x_in: "x \ S \ T" shows "((\x. if x \ S then f x else g x) has_derivative (if x \ S then f' x else g' x)) (at x within (S \ T))" proof - from f' x_in interpret f': bounded_linear "if x \ S then f' x else (\x. 0)" by (auto simp add: has_derivative_within) from g' interpret g': bounded_linear "if x \ T then g' x else (\x. 0)" by (auto simp add: has_derivative_within) have bl: "bounded_linear (if x \ S then f' x else g' x)" using f'.scaleR f'.bounded f'.add g'.scaleR g'.bounded g'.add x_in by (unfold_locales; force) show ?thesis using f' g' closure_subset[of T] closure_subset[of S] unfolding has_derivative_within_If_eq by (intro conjI bl tendsto_If_within_closures x_in) (auto simp: has_derivative_within inverse_eq_divide connect connect' subsetD) qed lemma has_vector_derivative_If_within_closures: assumes x_in: "x \ S \ T" assumes "u = S \ T" assumes f': "x \ S \ (closure S \ closure T) \ (f has_vector_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_vector_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" shows "((\x. if x \ S then f x else g x) has_vector_derivative (if x \ S then f' x else g' x)) (at x within u)" unfolding has_vector_derivative_def assms using x_in apply (intro has_derivative_If_within_closures[where ?f' = "\x a. a *\<^sub>R f' x" and ?g' = "\x a. a *\<^sub>R g' x", THEN has_derivative_eq_rhs]) subgoal by (rule f'[unfolded has_vector_derivative_def]; assumption) subgoal by (rule g'[unfolded has_vector_derivative_def]; assumption) by (auto simp: assms) subsection\<^marker>\tag important\\The Inverse Function Theorem\ lemma linear_injective_contraction: assumes "linear f" "c < 1" and le: "\x. norm (f x - x) \ c * norm x" shows "inj f" unfolding linear_injective_0[OF \linear f\] proof safe fix x assume "f x = 0" with le [of x] have "norm x \ c * norm x" by simp then show "x = 0" using \c < 1\ by (simp add: mult_le_cancel_right1) qed text\From an online proof by J. Michael Boardman, Department of Mathematics, Johns Hopkins University\ lemma inverse_function_theorem_scaled: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative blinfun_apply (f' x)) (at x)" and contf: "continuous_on U f'" and "0 \ U" and [simp]: "f 0 = 0" and id: "f' 0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "0 \ U'" "open V" "0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - obtain d1 where "cball 0 d1 \ U" "d1 > 0" using \open U\ \0 \ U\ open_contains_cball by blast obtain d2 where d2: "\x. \x \ U; dist x 0 \ d2\ \ dist (f' x) (f' 0) < 1/2" "0 < d2" using continuous_onE [OF contf, of 0 "1/2"] by (metis \0 \ U\ half_gt_zero_iff zero_less_one) obtain \ where le: "\x. norm x \ \ \ dist (f' x) id_blinfun \ 1/2" and "0 < \" and subU: "cball 0 \ \ U" proof show "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) show "cball 0 (min d1 d2) \ U" using \cball 0 d1 \ U\ by auto show "dist (f' x) id_blinfun \ 1/2" if "norm x \ min d1 d2" for x using \cball 0 d1 \ U\ d2 that id by fastforce qed let ?D = "cball 0 \" define V:: "'a set" where "V \ ball 0 (\/2)" have 4: "norm (f (x + h) - f x - h) \ 1/2 * norm h" if "x \ ?D" "x+h \ ?D" for x h proof - let ?w = "\x. f x - x" have B: "\x. x \ ?D \ onorm (blinfun_apply (f' x - id_blinfun)) \ 1/2" by (metis dist_norm le mem_cball_0 norm_blinfun.rep_eq) have "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x)" by (rule derivative_eq_intros derf subsetD [OF subU] | force simp: blinfun.diff_left)+ then have Dw: "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x within ?D)" using has_derivative_at_withinI by blast have "norm (?w (x+h) - ?w x) \ (1/2) * norm h" using differentiable_bound [OF convex_cball Dw B] that by fastforce then show ?thesis by (auto simp: algebra_simps) qed have for_g: "\!x. norm x < \ \ f x = y" if y: "norm y < \/2" for y proof - let ?u = "\x. x + (y - f x)" have *: "norm (?u x) < \" if "x \ ?D" for x proof - have fxx: "norm (f x - x) \ \/2" using 4 [of 0 x] \0 < \\ \f 0 = 0\ that by auto have "norm (?u x) \ norm y + norm (f x - x)" by (metis add.commute add_diff_eq norm_minus_commute norm_triangle_ineq) also have "\ < \/2 + \/2" using fxx y by auto finally show ?thesis by simp qed have "\!x \ ?D. ?u x = x" proof (rule banach_fix) show "cball 0 \ \ {}" using \0 < \\ by auto show "(\x. x + (y - f x)) ` cball 0 \ \ cball 0 \" using * by force have "dist (x + (y - f x)) (xh + (y - f xh)) * 2 \ dist x xh" if "norm x \ \" and "norm xh \ \" for x xh using that 4 [of x "xh-x"] by (auto simp: dist_norm norm_minus_commute algebra_simps) then show "\x\cball 0 \. \ya\cball 0 \. dist (x + (y - f x)) (ya + (y - f ya)) \ (1/2) * dist x ya" by auto qed (auto simp: complete_eq_closed) then show ?thesis by (metis "*" add_cancel_right_right eq_iff_diff_eq_0 le_less mem_cball_0) qed define g where "g \ \y. THE x. norm x < \ \ f x = y" have g: "norm (g y) < \ \ f (g y) = y" if "norm y < \/2" for y unfolding g_def using that theI' [OF for_g] by meson then have fg[simp]: "f (g y) = y" if "y \ V" for y using that by (auto simp: V_def) have 5: "norm (g y' - g y) \ 2 * norm (y' - y)" if "y \ V" "y' \ V" for y y' proof - have no: "norm (g y) \ \" "norm (g y') \ \" and [simp]: "f (g y) = y" using that g unfolding V_def by force+ have "norm (g y' - g y) \ norm (g y' - g y - (y' - y)) + norm (y' - y)" by (simp add: add.commute norm_triangle_sub) also have "\ \ (1/2) * norm (g y' - g y) + norm (y' - y)" using 4 [of "g y" "g y' - g y"] that no by (simp add: g norm_minus_commute V_def) finally show ?thesis by auto qed have contg: "continuous_on V g" proof fix y::'a and e::real assume "0 < e" and y: "y \ V" show "\d>0. \x'\V. dist x' y < d \ dist (g x') (g y) \ e" proof (intro exI conjI ballI impI) show "0 < e/2" by (simp add: \0 < e\) qed (use 5 y in \force simp: dist_norm\) qed show thesis proof define U' where "U' \ (f -` V) \ ball 0 \" have contf: "continuous_on U f" using derf has_derivative_at_withinI by (fast intro: has_derivative_continuous_on) then have "continuous_on (ball 0 \) f" by (meson ball_subset_cball continuous_on_subset subU) then show "open U'" by (simp add: U'_def V_def Int_commute continuous_open_preimage) show "0 \ U'" "U' \ U" "open V" "0 \ V" using \0 < \\ subU by (auto simp: U'_def V_def) show hom: "homeomorphism U' V f g" proof show "continuous_on U' f" using \U' \ U\ contf continuous_on_subset by blast show "continuous_on V g" using contg by blast show "f ` U' \ V" using U'_def by blast show "g ` V \ U'" by (simp add: U'_def V_def g image_subset_iff) show "g (f x) = x" if "x \ U'" for x by (metis that fg Int_iff U'_def V_def for_g g mem_ball_0 vimage_eq) show "f (g y) = y" if "y \ V" for y using that by (simp add: g V_def) qed show bij: "bij (blinfun_apply (f'(g y)))" if "y \ V" for y proof - have inj: "inj (blinfun_apply (f' (g y)))" proof (rule linear_injective_contraction) show "linear (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by blast next fix x have "norm (blinfun_apply (f' (g y)) x - x) = norm (blinfun_apply (f' (g y) - id_blinfun) x)" by (simp add: blinfun.diff_left) also have "\ \ norm (f' (g y) - id_blinfun) * norm x" by (rule norm_blinfun) also have "\ \ (1/2) * norm x" proof (rule mult_right_mono) show "norm (f' (g y) - id_blinfun) \ 1/2" using that g [of y] le by (auto simp: V_def dist_norm) qed auto finally show "norm (blinfun_apply (f' (g y)) x - x) \ (1/2) * norm x" . qed auto moreover have "surj (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by (blast intro!: linear_inj_imp_surj [OF _ inj]) ultimately show ?thesis using bijI by blast qed define g' where "g' \ \y. inv (blinfun_apply (f'(g y)))" show "(g has_derivative g' y) (at y)" if "y \ V" for y proof - have gy: "g y \ U" using g subU that unfolding V_def by fastforce obtain e where e: "\h. f (g y + h) = y + blinfun_apply (f' (g y)) h + e h" and e0: "(\h. norm (e h) / norm h) \0\ 0" using iffD1 [OF has_derivative_iff_Ex derf [OF gy]] \y \ V\ by auto have [simp]: "e 0 = 0" using e [of 0] that by simp let ?INV = "inv (blinfun_apply (f' (g y)))" have inj: "inj (blinfun_apply (f' (g y)))" using bij bij_betw_def that by blast have "(g has_derivative g' y) (at y within V)" unfolding has_derivative_at_within_iff_Ex [OF \y \ V\ \open V\] proof show blinv: "bounded_linear (g' y)" unfolding g'_def using derf gy inj inj_linear_imp_inv_bounded_linear by blast define eg where "eg \ \k. - ?INV (e (g (y+k) - g y))" have "g (y+k) = g y + g' y k + eg k" if "y + k \ V" for k proof - have "?INV k = ?INV (blinfun_apply (f' (g y)) (g (y+k) - g y) + e (g (y+k) - g y))" using e [of "g(y+k) - g y"] that by simp then have "g (y+k) = g y + ?INV k - ?INV (e (g (y+k) - g y))" using inj blinv by (simp add: linear_simps g'_def) then show ?thesis by (auto simp: eg_def g'_def) qed moreover have "(\k. norm (eg k) / norm k) \0\ 0" proof (rule Lim_null_comparison) let ?g = "\k. 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" show "\\<^sub>F k in at 0. norm (norm (eg k) / norm k) \ ?g k" unfolding eventually_at_topological proof (intro exI conjI ballI impI) show "open ((+)(-y) ` V)" using \open V\ open_translation by blast show "0 \ (+)(-y) ` V" by (simp add: that) show "norm (norm (eg k) / norm k) \ 2 * onorm (inv (blinfun_apply (f' (g y)))) * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" if "k \ (+)(-y) ` V" "k \ 0" for k proof - have "y+k \ V" using that by auto have "norm (norm (eg k) / norm k) \ onorm ?INV * norm (e (g (y+k) - g y)) / norm k" using blinv g'_def onorm by (force simp: eg_def divide_simps) also have "\ = (norm (g (y+k) - g y) / norm k) * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" by (simp add: divide_simps) also have "\ \ 2 * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" apply (rule mult_right_mono) using 5 [of y "y+k"] \y \ V\ \y + k \ V\ onorm_pos_le [OF blinv] apply (auto simp: divide_simps zero_le_mult_iff zero_le_divide_iff g'_def) done finally show "norm (norm (eg k) / norm k) \ 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" by simp qed qed have 1: "(\h. norm (e h) / norm h) \0\ (norm (e 0) / norm 0)" using e0 by auto have 2: "(\k. g (y+k) - g y) \0\ 0" using contg \open V\ \y \ V\ LIM_offset_zero_iff LIM_zero_iff at_within_open continuous_on_def by fastforce from tendsto_compose [OF 1 2, simplified] have "(\k. norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)) \0\ 0" . from tendsto_mult_left [OF this] show "?g \0\ 0" by auto qed ultimately show "\e. (\k. y + k \ V \ g (y+k) = g y + g' y k + e k) \ (\k. norm (e k) / norm k) \0\ 0" by blast qed then show ?thesis by (metis \open V\ at_within_open that) qed show "g' y = inv (blinfun_apply (f' (g y)))" if "y \ V" for y by (simp add: g'_def) qed qed text\We need all this to justify the scaling and translations.\ theorem inverse_function_theorem: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative (blinfun_apply (f' x))) (at x)" and contf: "continuous_on U f'" and "x0 \ U" and invf: "invf o\<^sub>L f' x0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "x0 \ U'" "open V" "f x0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - have apply1 [simp]: "\i. blinfun_apply invf (blinfun_apply (f' x0) i) = i" by (metis blinfun_apply_blinfun_compose blinfun_apply_id_blinfun invf) have apply2 [simp]: "\i. blinfun_apply (f' x0) (blinfun_apply invf i) = i" by (metis apply1 bij_inv_eq_iff blinfun_bij1 invf) have [simp]: "(range (blinfun_apply invf)) = UNIV" using apply1 surjI by blast let ?f = "invf \ (\x. (f \ (+)x0)x - f x0)" let ?f' = "\x. invf o\<^sub>L (f' (x + x0))" obtain U' V g g' where "open U'" and U': "U' \ (+)(-x0) ` U" "0 \ U'" and "open V" "0 \ V" and hom: "homeomorphism U' V ?f g" and derg: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv (?f'(g y))" and bij: "\y. y \ V \ bij (?f'(g y))" proof (rule inverse_function_theorem_scaled [of "(+)(-x0) ` U" ?f "?f'"]) show ope: "open ((+) (- x0) ` U)" using \open U\ open_translation by blast show "(?f has_derivative blinfun_apply (?f' x)) (at x)" if "x \ (+) (- x0) ` U" for x using that apply clarify apply (rule derf derivative_eq_intros | simp add: blinfun_compose.rep_eq)+ done have YY: "(\x. f' (x + x0)) \u-x0\ f' u" if "f' \u\ f' u" "u \ U" for u using that LIM_offset [where k = x0] by (auto simp: algebra_simps) then have "continuous_on ((+) (- x0) ` U) (\x. f' (x + x0))" using contf \open U\ Lim_at_imp_Lim_at_within by (fastforce simp: continuous_on_def at_within_open_NO_MATCH ope) then show "continuous_on ((+) (- x0) ` U) ?f'" by (intro continuous_intros) simp qed (auto simp: invf \x0 \ U\) show thesis proof let ?U' = "(+)x0 ` U'" let ?V = "((+)(f x0) \ f' x0) ` V" let ?g = "(+)x0 \ g \ invf \ (+)(- f x0)" let ?g' = "\y. inv (blinfun_apply (f' (?g y)))" show oU': "open ?U'" by (simp add: \open U'\ open_translation) show subU: "?U' \ U" using ComplI \U' \ (+) (- x0) ` U\ by auto show "x0 \ ?U'" by (simp add: \0 \ U'\) show "open ?V" using blinfun_bij2 [OF invf] by (metis \open V\ bij_is_surj blinfun.bounded_linear_right bounded_linear_def image_comp open_surjective_linear_image open_translation) show "f x0 \ ?V" using \0 \ V\ image_iff by fastforce show "homeomorphism ?U' ?V f ?g" proof show "continuous_on ?U' f" by (meson subU continuous_on_eq_continuous_at derf has_derivative_continuous oU' subsetD) have "?f ` U' \ V" using hom homeomorphism_image1 by blast then show "f ` ?U' \ ?V" unfolding image_subset_iff by (clarsimp simp: image_def) (metis apply2 add.commute diff_add_cancel) show "?g ` ?V \ ?U'" using hom invf by (auto simp: image_def homeomorphism_def) show "?g (f x) = x" if "x \ ?U'" for x using that hom homeomorphism_apply1 by fastforce have "continuous_on V g" using hom homeomorphism_def by blast then show "continuous_on ?V ?g" by (intro continuous_intros) (auto elim!: continuous_on_subset) have fg: "?f (g x) = x" if "x \ V" for x using hom homeomorphism_apply2 that by blast show "f (?g y) = y" if "y \ ?V" for y using that fg by (simp add: image_iff) (metis apply2 add.commute diff_add_cancel) qed show "(?g has_derivative ?g' y) (at y)" "bij (blinfun_apply (f' (?g y)))" if "y \ ?V" for y proof - have 1: "bij (blinfun_apply invf)" using blinfun_bij1 invf by blast then have 2: "bij (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x by (metis add.commute bij bij_betw_comp_iff2 blinfun_compose.rep_eq that top_greatest) then show "bij (blinfun_apply (f' (?g y)))" using that by auto have "g' x \ blinfun_apply invf = inv (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x using that by (simp add: g' o_inv_distrib blinfun_compose.rep_eq 1 2 add.commute bij_is_inj flip: o_assoc) then show "(?g has_derivative ?g' y) (at y)" using that invf by clarsimp (rule derg derivative_eq_intros | simp flip: id_def)+ qed qed auto qed subsection\<^marker>\tag unimportant\ \Piecewise differentiable functions\ definition piecewise_differentiable_on (infixr "piecewise'_differentiable'_on" 50) where "f piecewise_differentiable_on i \ continuous_on i f \ (\S. finite S \ (\x \ i - S. f differentiable (at x within i)))" lemma piecewise_differentiable_on_imp_continuous_on: "f piecewise_differentiable_on S \ continuous_on S f" by (simp add: piecewise_differentiable_on_def) lemma piecewise_differentiable_on_subset: "f piecewise_differentiable_on S \ T \ S \ f piecewise_differentiable_on T" using continuous_on_subset unfolding piecewise_differentiable_on_def apply safe apply (blast elim: continuous_on_subset) by (meson Diff_iff differentiable_within_subset subsetCE) lemma differentiable_on_imp_piecewise_differentiable: fixes a:: "'a::{linorder_topology,real_normed_vector}" shows "f differentiable_on {a..b} \ f piecewise_differentiable_on {a..b}" apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on) apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def) done lemma differentiable_imp_piecewise_differentiable: "(\x. x \ S \ f differentiable (at x within S)) \ f piecewise_differentiable_on S" by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def intro: differentiable_within_subset) lemma piecewise_differentiable_const [iff]: "(\x. z) piecewise_differentiable_on S" by (simp add: differentiable_imp_piecewise_differentiable) lemma piecewise_differentiable_compose: "\f piecewise_differentiable_on S; g piecewise_differentiable_on (f ` S); \x. finite (S \ f-`{x})\ \ (g \ f) piecewise_differentiable_on S" apply (simp add: piecewise_differentiable_on_def, safe) apply (blast intro: continuous_on_compose2) apply (rename_tac A B) apply (rule_tac x="A \ (\x\B. S \ f-`{x})" in exI) apply (blast intro!: differentiable_chain_within) done lemma piecewise_differentiable_affine: fixes m::real assumes "f piecewise_differentiable_on ((\x. m *\<^sub>R x + c) ` S)" shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_differentiable_on S" proof (cases "m = 0") case True then show ?thesis unfolding o_def by (force intro: differentiable_imp_piecewise_differentiable differentiable_const) next case False show ?thesis apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable]) apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+ done qed lemma piecewise_differentiable_cases: fixes c::real assumes "f piecewise_differentiable_on {a..c}" "g piecewise_differentiable_on {c..b}" "a \ c" "c \ b" "f c = g c" shows "(\x. if x \ c then f x else g x) piecewise_differentiable_on {a..b}" proof - obtain S T where st: "finite S" "finite T" and fd: "\x. x \ {a..c} - S \ f differentiable at x within {a..c}" and gd: "\x. x \ {c..b} - T \ g differentiable at x within {c..b}" using assms by (auto simp: piecewise_differentiable_on_def) have finabc: "finite ({a,b,c} \ (S \ T))" by (metis \finite S\ \finite T\ finite_Un finite_insert finite.emptyI) have "continuous_on {a..c} f" "continuous_on {c..b} g" using assms piecewise_differentiable_on_def by auto then have "continuous_on {a..b} (\x. if x \ c then f x else g x)" using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], OF closed_real_atLeastAtMost [of c b], of f g "\x. x\c"] assms by (force simp: ivl_disj_un_two_touch) moreover { fix x assume x: "x \ {a..b} - ({a,b,c} \ (S \ T))" have "(\x. if x \ c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg") proof (cases x c rule: le_cases) case le show ?diff_fg proof (rule differentiable_transform_within [where d = "dist x c"]) have "f differentiable at x" using x le fd [of x] at_within_interior [of x "{a..c}"] by simp then show "f differentiable at x within {a..b}" by (simp add: differentiable_at_withinI) qed (use x le st dist_real_def in auto) next case ge show ?diff_fg proof (rule differentiable_transform_within [where d = "dist x c"]) have "g differentiable at x" using x ge gd [of x] at_within_interior [of x "{c..b}"] by simp then show "g differentiable at x within {a..b}" by (simp add: differentiable_at_withinI) qed (use x ge st dist_real_def in auto) qed } then have "\S. finite S \ (\x\{a..b} - S. (\x. if x \ c then f x else g x) differentiable at x within {a..b})" by (meson finabc) ultimately show ?thesis by (simp add: piecewise_differentiable_on_def) qed lemma piecewise_differentiable_neg: "f piecewise_differentiable_on S \ (\x. -(f x)) piecewise_differentiable_on S" by (auto simp: piecewise_differentiable_on_def continuous_on_minus) lemma piecewise_differentiable_add: assumes "f piecewise_differentiable_on i" "g piecewise_differentiable_on i" shows "(\x. f x + g x) piecewise_differentiable_on i" proof - obtain S T where st: "finite S" "finite T" "\x\i - S. f differentiable at x within i" "\x\i - T. g differentiable at x within i" using assms by (auto simp: piecewise_differentiable_on_def) then have "finite (S \ T) \ (\x\i - (S \ T). (\x. f x + g x) differentiable at x within i)" by auto moreover have "continuous_on i f" "continuous_on i g" using assms piecewise_differentiable_on_def by auto ultimately show ?thesis by (auto simp: piecewise_differentiable_on_def continuous_on_add) qed lemma piecewise_differentiable_diff: "\f piecewise_differentiable_on S; g piecewise_differentiable_on S\ \ (\x. f x - g x) piecewise_differentiable_on S" unfolding diff_conv_add_uminus by (metis piecewise_differentiable_add piecewise_differentiable_neg) subsection\The concept of continuously differentiable\ text \ John Harrison writes as follows: ``The usual assumption in complex analysis texts is that a path \\\ should be piecewise continuously differentiable, which ensures that the path integral exists at least for any continuous f, since all piecewise continuous functions are integrable. However, our notion of validity is weaker, just piecewise differentiability\ldots{} [namely] continuity plus differentiability except on a finite set\ldots{} [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this can integrate all derivatives.'' "Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ definition\<^marker>\tag important\ C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" (infix "C1'_differentiable'_on" 50) where "f C1_differentiable_on S \ (\D. (\x \ S. (f has_vector_derivative (D x)) (at x)) \ continuous_on S D)" lemma C1_differentiable_on_eq: "f C1_differentiable_on S \ (\x \ S. f differentiable at x) \ continuous_on S (\x. vector_derivative f (at x))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding C1_differentiable_on_def by (metis (no_types, lifting) continuous_on_eq differentiableI_vector vector_derivative_at) next assume ?rhs then show ?lhs using C1_differentiable_on_def vector_derivative_works by fastforce qed lemma C1_differentiable_on_subset: "f C1_differentiable_on T \ S \ T \ f C1_differentiable_on S" unfolding C1_differentiable_on_def continuous_on_eq_continuous_within by (blast intro: continuous_within_subset) lemma C1_differentiable_compose: assumes fg: "f C1_differentiable_on S" "g C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" shows "(g \ f) C1_differentiable_on S" proof - have "\x. x \ S \ g \ f differentiable at x" by (meson C1_differentiable_on_eq assms differentiable_chain_at imageI) moreover have "continuous_on S (\x. vector_derivative (g \ f) (at x))" proof (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) show "continuous_on S (\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)))" using fg apply (clarsimp simp add: C1_differentiable_on_eq) apply (rule Limits.continuous_on_scaleR, assumption) by (metis (mono_tags, lifting) continuous_at_imp_continuous_on continuous_on_compose continuous_on_cong differentiable_imp_continuous_within o_def) show "\x. x \ S \ vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)) = vector_derivative (g \ f) (at x)" by (metis (mono_tags, hide_lams) C1_differentiable_on_eq fg imageI vector_derivative_chain_at) qed ultimately show ?thesis by (simp add: C1_differentiable_on_eq) qed lemma C1_diff_imp_diff: "f C1_differentiable_on S \ f differentiable_on S" by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\x. x) C1_differentiable_on S" by (auto simp: C1_differentiable_on_eq) lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" by (auto simp: C1_differentiable_on_eq) lemma C1_differentiable_on_add [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x + g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_minus [simp, derivative_intros]: "f C1_differentiable_on S \ (\x. - f x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_diff [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x - g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_mult [simp, derivative_intros]: fixes f g :: "real \ 'a :: real_normed_algebra" shows "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x * g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) lemma C1_differentiable_on_scaleR [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x *\<^sub>R g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ definition\<^marker>\tag important\ piecewise_C1_differentiable_on (infixr "piecewise'_C1'_differentiable'_on" 50) where "f piecewise_C1_differentiable_on i \ continuous_on i f \ (\S. finite S \ (f C1_differentiable_on (i - S)))" lemma C1_differentiable_imp_piecewise: "f C1_differentiable_on S \ f piecewise_C1_differentiable_on S" by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) lemma piecewise_C1_imp_differentiable: "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def C1_differentiable_on_def differentiable_def has_vector_derivative_def intro: has_derivative_at_withinI) lemma piecewise_C1_differentiable_compose: assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" shows "(g \ f) piecewise_C1_differentiable_on S" proof - have "continuous_on S (\x. g (f x))" by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" proof - obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" using fg by (auto simp: piecewise_C1_differentiable_on_def) obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" using fg by (auto simp: piecewise_C1_differentiable_on_def) show ?thesis proof (intro exI conjI) show "finite (F \ (\x\G. S \ f-`{x}))" using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" apply (rule C1_differentiable_compose) apply (blast intro: C1_differentiable_on_subset [OF F]) apply (blast intro: C1_differentiable_on_subset [OF G]) by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) qed qed ultimately show ?thesis by (simp add: piecewise_C1_differentiable_on_def) qed lemma piecewise_C1_differentiable_on_subset: "f piecewise_C1_differentiable_on S \ T \ S \ f piecewise_C1_differentiable_on T" by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) lemma C1_differentiable_imp_continuous_on: "f C1_differentiable_on S \ continuous_on S f" unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within using differentiable_at_withinI differentiable_imp_continuous_within by blast lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" unfolding C1_differentiable_on_def by auto lemma piecewise_C1_differentiable_affine: fixes m::real assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` S)" shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on S" proof (cases "m = 0") case True then show ?thesis unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def) next case False have *: "\x. finite (S \ {y. m * y + c = x})" using False not_finite_existsD by fastforce show ?thesis apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) apply (rule * assms derivative_intros | simp add: False vimage_def)+ done qed lemma piecewise_C1_differentiable_cases: fixes c::real assumes "f piecewise_C1_differentiable_on {a..c}" "g piecewise_C1_differentiable_on {c..b}" "a \ c" "c \ b" "f c = g c" shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" proof - obtain S T where st: "f C1_differentiable_on ({a..c} - S)" "g C1_differentiable_on ({c..b} - T)" "finite S" "finite T" using assms by (force simp: piecewise_C1_differentiable_on_def) then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], OF closed_real_atLeastAtMost [of c b], of f g "\x. x\c"] assms by (force simp: ivl_disj_un_two_touch) { fix x assume x: "x \ {a..b} - insert c (S \ T)" have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") proof (cases x c rule: le_cases) case le show ?diff_fg apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) next case ge show ?diff_fg apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) qed } then have "(\x \ {a..b} - insert c (S \ T). (\x. if x \ c then f x else g x) differentiable at x)" by auto moreover { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" if "a < x" "x < c" "x \ S" for x proof - have f: "f differentiable at x" by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) show ?thesis using that apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] f) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" if "c < x" "x < b" "x \ T" for x proof - have g: "g differentiable at x" by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) show ?thesis using that apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] g) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed ultimately have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" by (rule continuous_on_subset [OF continuous_on_open_Un], auto) } note * = this have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) ultimately have "\S. finite S \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - S)" apply (rule_tac x="{a,b,c} \ S \ T" in exI) using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) with cab show ?thesis by (simp add: piecewise_C1_differentiable_on_def) qed lemma piecewise_C1_differentiable_neg: "f piecewise_C1_differentiable_on S \ (\x. -(f x)) piecewise_C1_differentiable_on S" unfolding piecewise_C1_differentiable_on_def by (auto intro!: continuous_on_minus C1_differentiable_on_minus) lemma piecewise_C1_differentiable_add: assumes "f piecewise_C1_differentiable_on i" "g piecewise_C1_differentiable_on i" shows "(\x. f x + g x) piecewise_C1_differentiable_on i" proof - obtain S t where st: "finite S" "finite t" "f C1_differentiable_on (i-S)" "g C1_differentiable_on (i-t)" using assms by (auto simp: piecewise_C1_differentiable_on_def) then have "finite (S \ t) \ (\x. f x + g x) C1_differentiable_on i - (S \ t)" by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) moreover have "continuous_on i f" "continuous_on i g" using assms piecewise_C1_differentiable_on_def by auto ultimately show ?thesis by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) qed lemma piecewise_C1_differentiable_diff: "\f piecewise_C1_differentiable_on S; g piecewise_C1_differentiable_on S\ \ (\x. f x - g x) piecewise_C1_differentiable_on S" unfolding diff_conv_add_uminus by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) end diff --git a/src/HOL/Analysis/Elementary_Metric_Spaces.thy b/src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy @@ -1,3258 +1,3258 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Elementary Metric Spaces\ theory Elementary_Metric_Spaces imports Abstract_Topology_2 Metric_Arith begin subsection \Open and closed balls\ definition\<^marker>\tag important\ ball :: "'a::metric_space \ real \ 'a set" where "ball x e = {y. dist x y < e}" definition\<^marker>\tag important\ cball :: "'a::metric_space \ real \ 'a set" where "cball x e = {y. dist x y \ e}" definition\<^marker>\tag important\ sphere :: "'a::metric_space \ real \ 'a set" where "sphere x e = {y. dist x y = e}" lemma mem_ball [simp, metric_unfold]: "y \ ball x e \ dist x y < e" by (simp add: ball_def) lemma mem_cball [simp, metric_unfold]: "y \ cball x e \ dist x y \ e" by (simp add: cball_def) lemma mem_sphere [simp]: "y \ sphere x e \ dist x y = e" by (simp add: sphere_def) lemma ball_trivial [simp]: "ball x 0 = {}" by (simp add: ball_def) lemma cball_trivial [simp]: "cball x 0 = {x}" by (simp add: cball_def) lemma sphere_trivial [simp]: "sphere x 0 = {x}" by (simp add: sphere_def) lemma disjoint_ballI: "dist x y \ r+s \ ball x r \ ball y s = {}" using dist_triangle_less_add not_le by fastforce lemma disjoint_cballI: "dist x y > r + s \ cball x r \ cball y s = {}" by (metis add_mono disjoint_iff_not_equal dist_triangle2 dual_order.trans leD mem_cball) lemma sphere_empty [simp]: "r < 0 \ sphere a r = {}" for a :: "'a::metric_space" by auto lemma centre_in_ball [simp]: "x \ ball x e \ 0 < e" by simp lemma centre_in_cball [simp]: "x \ cball x e \ 0 \ e" by simp lemma ball_subset_cball [simp, intro]: "ball x e \ cball x e" by (simp add: subset_eq) lemma mem_ball_imp_mem_cball: "x \ ball y e \ x \ cball y e" by auto lemma sphere_cball [simp,intro]: "sphere z r \ cball z r" by force lemma cball_diff_sphere: "cball a r - sphere a r = ball a r" by auto lemma subset_ball[intro]: "d \ e \ ball x d \ ball x e" by auto lemma subset_cball[intro]: "d \ e \ cball x d \ cball x e" by auto lemma mem_ball_leI: "x \ ball y e \ e \ f \ x \ ball y f" by auto lemma mem_cball_leI: "x \ cball y e \ e \ f \ x \ cball y f" by auto lemma cball_trans: "y \ cball z b \ x \ cball y a \ x \ cball z (b + a)" by metric lemma ball_max_Un: "ball a (max r s) = ball a r \ ball a s" by auto lemma ball_min_Int: "ball a (min r s) = ball a r \ ball a s" by auto lemma cball_max_Un: "cball a (max r s) = cball a r \ cball a s" by auto lemma cball_min_Int: "cball a (min r s) = cball a r \ cball a s" by auto lemma cball_diff_eq_sphere: "cball a r - ball a r = sphere a r" by auto lemma open_ball [intro, simp]: "open (ball x e)" proof - have "open (dist x -` {.. (\x\S. \e>0. ball x e \ S)" by (simp add: open_dist subset_eq Ball_def dist_commute) lemma openI [intro?]: "(\x. x\S \ \e>0. ball x e \ S) \ open S" by (auto simp: open_contains_ball) lemma openE[elim?]: assumes "open S" "x\S" obtains e where "e>0" "ball x e \ S" using assms unfolding open_contains_ball by auto lemma open_contains_ball_eq: "open S \ x\S \ (\e>0. ball x e \ S)" by (metis open_contains_ball subset_eq centre_in_ball) lemma ball_eq_empty[simp]: "ball x e = {} \ e \ 0" unfolding mem_ball set_eq_iff by (simp add: not_less) metric lemma ball_empty: "e \ 0 \ ball x e = {}" by simp lemma closed_cball [iff]: "closed (cball x e)" proof - have "closed (dist x -` {..e})" by (intro closed_vimage closed_atMost continuous_intros) also have "dist x -` {..e} = cball x e" by auto finally show ?thesis . qed lemma open_contains_cball: "open S \ (\x\S. \e>0. cball x e \ S)" proof - { fix x and e::real assume "x\S" "e>0" "ball x e \ S" then have "\d>0. cball x d \ S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) } moreover { fix x and e::real assume "x\S" "e>0" "cball x e \ S" then have "\d>0. ball x d \ S" using mem_ball_imp_mem_cball by blast } ultimately show ?thesis unfolding open_contains_ball by auto qed lemma open_contains_cball_eq: "open S \ (\x. x \ S \ (\e>0. cball x e \ S))" by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball) lemma eventually_nhds_ball: "d > 0 \ eventually (\x. x \ ball z d) (nhds z)" by (rule eventually_nhds_in_open) simp_all lemma eventually_at_ball: "d > 0 \ eventually (\t. t \ ball z d \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma eventually_at_ball': "d > 0 \ eventually (\t. t \ ball z d \ t \ z \ t \ A) (at z within A)" unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) lemma at_within_ball: "e > 0 \ dist x y < e \ at y within ball x e = at y" by (subst at_within_open) auto lemma atLeastAtMost_eq_cball: fixes a b::real shows "{a .. b} = cball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma cball_eq_atLeastAtMost: fixes a b::real shows "cball a b = {a - b .. a + b}" by (auto simp: dist_real_def) lemma greaterThanLessThan_eq_ball: fixes a b::real shows "{a <..< b} = ball ((a + b)/2) ((b - a)/2)" by (auto simp: dist_real_def field_simps) lemma ball_eq_greaterThanLessThan: fixes a b::real shows "ball a b = {a - b <..< a + b}" by (auto simp: dist_real_def) lemma interior_ball [simp]: "interior (ball x e) = ball x e" by (simp add: interior_open) lemma cball_eq_empty [simp]: "cball x e = {} \ e < 0" apply (simp add: set_eq_iff not_le) apply (metis zero_le_dist dist_self order_less_le_trans) done lemma cball_empty [simp]: "e < 0 \ cball x e = {}" by simp lemma cball_sing: fixes x :: "'a::metric_space" shows "e = 0 \ cball x e = {x}" by simp lemma ball_divide_subset: "d \ 1 \ ball x (e/d) \ ball x e" by (metis ball_eq_empty div_by_1 frac_le linear subset_ball zero_less_one) lemma ball_divide_subset_numeral: "ball x (e / numeral w) \ ball x e" using ball_divide_subset one_le_numeral by blast lemma cball_divide_subset: "d \ 1 \ cball x (e/d) \ cball x e" apply (cases "e < 0", simp add: field_split_simps) by (metis div_by_1 frac_le less_numeral_extra(1) not_le order_refl subset_cball) lemma cball_divide_subset_numeral: "cball x (e / numeral w) \ cball x e" using cball_divide_subset one_le_numeral by blast lemma cball_scale: assumes "a \ 0" shows "(\x. a *\<^sub>R x) ` cball c r = cball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" proof - have 1: "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a proof safe fix x assume x: "x \ cball c r" have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" by (auto simp: dist_norm) also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" by (simp add: algebra_simps) finally show "a *\<^sub>R x \ cball (a *\<^sub>R c) (\a\ * r)" using that x by (auto simp: dist_norm) qed have "cball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` cball (a *\<^sub>R c) (\a\ * r)" unfolding image_image using assms by simp also have "\ \ (\x. a *\<^sub>R x) ` cball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" using assms by (intro image_mono 1) auto also have "\ = (\x. a *\<^sub>R x) ` cball c r" using assms by (simp add: algebra_simps) finally have "cball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` cball c r" . moreover from assms have "(\x. a *\<^sub>R x) ` cball c r \ cball (a *\<^sub>R c) (\a\ * r)" by (intro 1) auto ultimately show ?thesis by blast qed lemma ball_scale: assumes "a \ 0" shows "(\x. a *\<^sub>R x) ` ball c r = ball (a *\<^sub>R c :: 'a :: real_normed_vector) (\a\ * r)" proof - have 1: "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" if "a \ 0" for a r and c :: 'a proof safe fix x assume x: "x \ ball c r" have "dist (a *\<^sub>R c) (a *\<^sub>R x) = norm (a *\<^sub>R c - a *\<^sub>R x)" by (auto simp: dist_norm) also have "a *\<^sub>R c - a *\<^sub>R x = a *\<^sub>R (c - x)" by (simp add: algebra_simps) finally show "a *\<^sub>R x \ ball (a *\<^sub>R c) (\a\ * r)" using that x by (auto simp: dist_norm) qed have "ball (a *\<^sub>R c) (\a\ * r) = (\x. a *\<^sub>R x) ` (\x. inverse a *\<^sub>R x) ` ball (a *\<^sub>R c) (\a\ * r)" unfolding image_image using assms by simp also have "\ \ (\x. a *\<^sub>R x) ` ball (inverse a *\<^sub>R (a *\<^sub>R c)) (\inverse a\ * (\a\ * r))" using assms by (intro image_mono 1) auto also have "\ = (\x. a *\<^sub>R x) ` ball c r" using assms by (simp add: algebra_simps) finally have "ball (a *\<^sub>R c) (\a\ * r) \ (\x. a *\<^sub>R x) ` ball c r" . moreover from assms have "(\x. a *\<^sub>R x) ` ball c r \ ball (a *\<^sub>R c) (\a\ * r)" by (intro 1) auto ultimately show ?thesis by blast qed subsection \Limit Points\ lemma islimpt_approachable: fixes x :: "'a::metric_space" shows "x islimpt S \ (\e>0. \x'\S. x' \ x \ dist x' x < e)" unfolding islimpt_iff_eventually eventually_at by fast lemma islimpt_approachable_le: "x islimpt S \ (\e>0. \x'\ S. x' \ x \ dist x' x \ e)" for x :: "'a::metric_space" unfolding islimpt_approachable using approachable_lt_le2 [where f="\y. dist y x" and P="\y. y \ S \ y = x" and Q="\x. True"] by auto lemma limpt_of_limpts: "x islimpt {y. y islimpt S} \ x islimpt S" for x :: "'a::metric_space" apply (clarsimp simp add: islimpt_approachable) apply (drule_tac x="e/2" in spec) apply (auto simp: simp del: less_divide_eq_numeral1) apply (drule_tac x="dist x' x" in spec) apply (auto simp del: less_divide_eq_numeral1) apply metric done lemma closed_limpts: "closed {x::'a::metric_space. x islimpt S}" using closed_limpt limpt_of_limpts by blast lemma limpt_of_closure: "x islimpt closure S \ x islimpt S" for x :: "'a::metric_space" by (auto simp: closure_def islimpt_Un dest: limpt_of_limpts) lemma islimpt_eq_infinite_ball: "x islimpt S \ (\e>0. infinite(S \ ball x e))" apply (simp add: islimpt_eq_acc_point, safe) apply (metis Int_commute open_ball centre_in_ball) by (metis open_contains_ball Int_mono finite_subset inf_commute subset_refl) lemma islimpt_eq_infinite_cball: "x islimpt S \ (\e>0. infinite(S \ cball x e))" apply (simp add: islimpt_eq_infinite_ball, safe) apply (meson Int_mono ball_subset_cball finite_subset order_refl) by (metis open_ball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq) subsection \Perfect Metric Spaces\ lemma perfect_choose_dist: "0 < r \ \a. a \ x \ dist a x < r" for x :: "'a::{perfect_space,metric_space}" using islimpt_UNIV [of x] by (simp add: islimpt_approachable) lemma cball_eq_sing: fixes x :: "'a::{metric_space,perfect_space}" shows "cball x e = {x} \ e = 0" proof (rule linorder_cases) assume e: "0 < e" obtain a where "a \ x" "dist a x < e" using perfect_choose_dist [OF e] by auto then have "a \ x" "dist x a \ e" by (auto simp: dist_commute) with e show ?thesis by (auto simp: set_eq_iff) qed auto subsection \?\ lemma finite_ball_include: fixes a :: "'a::metric_space" assumes "finite S" shows "\e>0. S \ ball a e" using assms proof induction case (insert x S) then obtain e0 where "e0>0" and e0:"S \ ball a e0" by auto define e where "e = max e0 (2 * dist a x)" have "e>0" unfolding e_def using \e0>0\ by auto moreover have "insert x S \ ball a e" using e0 \e>0\ unfolding e_def by auto ultimately show ?case by auto qed (auto intro: zero_less_one) lemma finite_set_avoid: fixes a :: "'a::metric_space" assumes "finite S" shows "\d>0. \x\S. x \ a \ d \ dist a x" using assms proof induction case (insert x S) then obtain d where "d > 0" and d: "\x\S. x \ a \ d \ dist a x" by blast show ?case proof (cases "x = a") case True with \d > 0 \d show ?thesis by auto next case False let ?d = "min d (dist a x)" from False \d > 0\ have dp: "?d > 0" by auto from d have d': "\x\S. x \ a \ ?d \ dist a x" by auto with dp False show ?thesis by (metis insert_iff le_less min_less_iff_conj not_less) qed qed (auto intro: zero_less_one) lemma discrete_imp_closed: fixes S :: "'a::metric_space set" assumes e: "0 < e" and d: "\x \ S. \y \ S. dist y x < e \ y = x" shows "closed S" proof - have False if C: "\e. e>0 \ \x'\S. x' \ x \ dist x' x < e" for x proof - from e have e2: "e/2 > 0" by arith from C[rule_format, OF e2] obtain y where y: "y \ S" "y \ x" "dist y x < e/2" by blast from e2 y(2) have mp: "min (e/2) (dist x y) > 0" by simp from d y C[OF mp] show ?thesis by metric qed then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a]) qed subsection \Interior\ lemma mem_interior: "x \ interior S \ (\e>0. ball x e \ S)" using open_contains_ball_eq [where S="interior S"] by (simp add: open_subset_interior) lemma mem_interior_cball: "x \ interior S \ (\e>0. cball x e \ S)" by (meson ball_subset_cball interior_subset mem_interior open_contains_cball open_interior subset_trans) subsection \Frontier\ lemma frontier_straddle: fixes a :: "'a::metric_space" shows "a \ frontier S \ (\e>0. (\x\S. dist a x < e) \ (\x. x \ S \ dist a x < e))" unfolding frontier_def closure_interior by (auto simp: mem_interior subset_eq ball_def) subsection \Limits\ proposition Lim: "(f \ l) net \ trivial_limit net \ (\e>0. eventually (\x. dist (f x) l < e) net)" by (auto simp: tendsto_iff trivial_limit_eq) text \Show that they yield usual definitions in the various cases.\ proposition Lim_within_le: "(f \ l)(at a within S) \ (\e>0. \d>0. \x\S. 0 < dist x a \ dist x a \ d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at_le) proposition Lim_within: "(f \ l) (at a within S) \ (\e >0. \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) corollary Lim_withinI [intro?]: assumes "\e. e > 0 \ \d>0. \x \ S. 0 < dist x a \ dist x a < d \ dist (f x) l \ e" shows "(f \ l) (at a within S)" apply (simp add: Lim_within, clarify) apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) done proposition Lim_at: "(f \ l) (at a) \ (\e >0. \d>0. \x. 0 < dist x a \ dist x a < d \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at) lemma Lim_transform_within_set: fixes a :: "'a::metric_space" and l :: "'b::metric_space" shows "\(f \ l) (at a within S); eventually (\x. x \ S \ x \ T) (at a)\ \ (f \ l) (at a within T)" apply (clarsimp simp: eventually_at Lim_within) apply (drule_tac x=e in spec, clarify) apply (rename_tac k) apply (rule_tac x="min d k" in exI, simp) done text \Another limit point characterization.\ lemma limpt_sequential_inj: fixes x :: "'a::metric_space" shows "x islimpt S \ (\f. (\n::nat. f n \ S - {x}) \ inj f \ (f \ x) sequentially)" (is "?lhs = ?rhs") proof assume ?lhs then have "\e>0. \x'\S. x' \ x \ dist x' x < e" by (force simp: islimpt_approachable) then obtain y where y: "\e. e>0 \ y e \ S \ y e \ x \ dist (y e) x < e" by metis define f where "f \ rec_nat (y 1) (\n fn. y (min (inverse(2 ^ (Suc n))) (dist fn x)))" have [simp]: "f 0 = y 1" "f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n by (simp_all add: f_def) have f: "f n \ S \ (f n \ x) \ dist (f n) x < inverse(2 ^ n)" for n proof (induction n) case 0 show ?case by (simp add: y) next case (Suc n) then show ?case apply (auto simp: y) by (metis half_gt_zero_iff inverse_positive_iff_positive less_divide_eq_numeral1(1) min_less_iff_conj y zero_less_dist_iff zero_less_numeral zero_less_power) qed show ?rhs proof (rule_tac x=f in exI, intro conjI allI) show "\n. f n \ S - {x}" using f by blast have "dist (f n) x < dist (f m) x" if "m < n" for m n using that proof (induction n) case 0 then show ?case by simp next case (Suc n) then consider "m < n" | "m = n" using less_Suc_eq by blast then show ?case proof cases assume "m < n" have "dist (f(Suc n)) x = dist (y (min (inverse(2 ^ (Suc n))) (dist (f n) x))) x" by simp also have "\ < dist (f n) x" by (metis dist_pos_lt f min.strict_order_iff min_less_iff_conj y) also have "\ < dist (f m) x" using Suc.IH \m < n\ by blast finally show ?thesis . next assume "m = n" then show ?case by simp (metis dist_pos_lt f half_gt_zero_iff inverse_positive_iff_positive min_less_iff_conj y zero_less_numeral zero_less_power) qed qed then show "inj f" by (metis less_irrefl linorder_injI) show "f \ x" apply (rule tendstoI) apply (rule_tac c="nat (ceiling(1/e))" in eventually_sequentiallyI) apply (rule less_trans [OF f [THEN conjunct2, THEN conjunct2]]) apply (simp add: field_simps) by (meson le_less_trans mult_less_cancel_left not_le of_nat_less_two_power) qed next assume ?rhs then show ?lhs by (fastforce simp add: islimpt_approachable lim_sequentially) qed lemma Lim_dist_ubound: assumes "\(trivial_limit net)" and "(f \ l) net" and "eventually (\x. dist a (f x) \ e) net" shows "dist a l \ e" using assms by (fast intro: tendsto_le tendsto_intros) subsection \Continuity\ text\Derive the epsilon-delta forms, which we often use as "definitions"\ proposition continuous_within_eps_delta: "continuous (at x within s) f \ (\e>0. \d>0. \x'\ s. dist x' x < d --> dist (f x') (f x) < e)" unfolding continuous_within and Lim_within by fastforce corollary continuous_at_eps_delta: "continuous (at x) f \ (\e > 0. \d > 0. \x'. dist x' x < d \ dist (f x') (f x) < e)" using continuous_within_eps_delta [of x UNIV f] by simp lemma continuous_at_right_real_increasing: fixes f :: "real \ real" assumes nondecF: "\x y. x \ y \ f x \ f y" shows "continuous (at_right a) f \ (\e>0. \d>0. f (a + d) - f a < e)" apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong, safe) apply (erule_tac x="a + d" in allE, simp) apply (simp add: nondecF field_simps) apply (drule nondecF, simp) done lemma continuous_at_left_real_increasing: assumes nondecF: "\ x y. x \ y \ f x \ ((f y) :: real)" shows "(continuous (at_left (a :: real)) f) = (\e > 0. \delta > 0. f a - f (a - delta) < e)" apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le) apply (intro all_cong ex_cong, safe) apply (erule_tac x="a - d" in allE, simp) apply (simp add: nondecF field_simps) apply (cut_tac x="a - d" and y=x in nondecF, simp_all) done text\Versions in terms of open balls.\ lemma continuous_within_ball: "continuous (at x within s) f \ (\e > 0. \d > 0. f ` (ball x d \ s) \ ball (f x) e)" (is "?lhs = ?rhs") proof assume ?lhs { fix e :: real assume "e > 0" then obtain d where d: "d>0" "\xa\s. 0 < dist xa x \ dist xa x < d \ dist (f xa) (f x) < e" using \?lhs\[unfolded continuous_within Lim_within] by auto { fix y assume "y \ f ` (ball x d \ s)" then have "y \ ball (f x) e" using d(2) using \e > 0\ by (auto simp: dist_commute) } then have "\d>0. f ` (ball x d \ s) \ ball (f x) e" using \d > 0\ unfolding subset_eq ball_def by (auto simp: dist_commute) } then show ?rhs by auto next assume ?rhs then show ?lhs unfolding continuous_within Lim_within ball_def subset_eq apply (auto simp: dist_commute) apply (erule_tac x=e in allE, auto) done qed lemma continuous_at_ball: "continuous (at x) f \ (\e>0. \d>0. f ` (ball x d) \ ball (f x) e)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball by (metis dist_commute dist_pos_lt dist_self) next assume ?rhs then show ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball by (metis dist_commute) qed text\Define setwise continuity in terms of limits within the set.\ lemma continuous_on_iff: "continuous_on s f \ (\x\s. \e>0. \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" unfolding continuous_on_def Lim_within by (metis dist_pos_lt dist_self) lemma continuous_within_E: assumes "continuous (at x within s) f" "e>0" obtains d where "d>0" "\x'. \x'\ s; dist x' x \ d\ \ dist (f x') (f x) < e" using assms apply (simp add: continuous_within_eps_delta) apply (drule spec [of _ e], clarify) apply (rule_tac d="d/2" in that, auto) done lemma continuous_onI [intro?]: assumes "\x e. \e > 0; x \ s\ \ \d>0. \x'\s. dist x' x < d \ dist (f x') (f x) \ e" shows "continuous_on s f" apply (simp add: continuous_on_iff, clarify) apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) done text\Some simple consequential lemmas.\ lemma continuous_onE: assumes "continuous_on s f" "x\s" "e>0" obtains d where "d>0" "\x'. \x' \ s; dist x' x \ d\ \ dist (f x') (f x) < e" using assms apply (simp add: continuous_on_iff) apply (elim ballE allE) apply (auto intro: that [where d="d/2" for d]) done text\The usual transformation theorems.\ lemma continuous_transform_within: fixes f g :: "'a::metric_space \ 'b::topological_space" assumes "continuous (at x within s) f" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "continuous (at x within s) g" using assms unfolding continuous_within by (force intro: Lim_transform_within) subsection \Closure and Limit Characterization\ lemma closure_approachable: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x < e)" apply (auto simp: closure_def islimpt_approachable) apply (metis dist_self) done lemma closure_approachable_le: fixes S :: "'a::metric_space set" shows "x \ closure S \ (\e>0. \y\S. dist y x \ e)" unfolding closure_approachable using dense by force lemma closure_approachableD: assumes "x \ closure S" "e>0" shows "\y\S. dist x y < e" using assms unfolding closure_approachable by (auto simp: dist_commute) lemma closed_approachable: fixes S :: "'a::metric_space set" shows "closed S \ (\e>0. \y\S. dist y x < e) \ x \ S" by (metis closure_closed closure_approachable) lemma closure_contains_Inf: fixes S :: "real set" assumes "S \ {}" "bdd_below S" shows "Inf S \ closure S" proof - have *: "\x\S. Inf S \ x" using cInf_lower[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Inf S < Inf S + e" by simp with assms obtain x where "x \ S" "x < Inf S + e" by (subst (asm) cInf_less_iff) auto with * have "\x\S. dist x (Inf S) < e" by (intro bexI[of _ x]) (auto simp: dist_real_def) } then show ?thesis unfolding closure_approachable by auto qed lemma closure_contains_Sup: fixes S :: "real set" assumes "S \ {}" "bdd_above S" shows "Sup S \ closure S" proof - have *: "\x\S. x \ Sup S" using cSup_upper[of _ S] assms by metis { fix e :: real assume "e > 0" then have "Sup S - e < Sup S" by simp with assms obtain x where "x \ S" "Sup S - e < x" by (subst (asm) less_cSup_iff) auto with * have "\x\S. dist x (Sup S) < e" by (intro bexI[of _ x]) (auto simp: dist_real_def) } then show ?thesis unfolding closure_approachable by auto qed lemma not_trivial_limit_within_ball: "\ trivial_limit (at x within S) \ (\e>0. S \ ball x e - {x} \ {})" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S - {x}" and "dist y x < e" using \?lhs\ not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto then have "y \ S \ ball x e - {x}" unfolding ball_def by (simp add: dist_commute) then have "S \ ball x e - {x} \ {}" by blast } then show ?thesis by auto qed show ?lhs if ?rhs proof - { fix e :: real assume "e > 0" then obtain y where "y \ S \ ball x e - {x}" using \?rhs\ by blast then have "y \ S - {x}" and "dist y x < e" unfolding ball_def by (simp_all add: dist_commute) then have "\y \ S - {x}. dist y x < e" by auto } then show ?thesis using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto qed qed subsection \Boundedness\ (* FIXME: This has to be unified with BSEQ!! *) definition\<^marker>\tag important\ (in metric_space) bounded :: "'a set \ bool" where "bounded S \ (\x e. \y\S. dist x y \ e)" lemma bounded_subset_cball: "bounded S \ (\e x. S \ cball x e \ 0 \ e)" unfolding bounded_def subset_eq by auto (meson order_trans zero_le_dist) lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" unfolding bounded_def by auto (metis add.commute add_le_cancel_right dist_commute dist_triangle_le) lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" unfolding bounded_any_center [where a=0] by (simp add: dist_norm) lemma bdd_above_norm: "bdd_above (norm ` X) \ bounded X" by (simp add: bounded_iff bdd_above_def) lemma bounded_norm_comp: "bounded ((\x. norm (f x)) ` S) = bounded (f ` S)" by (simp add: bounded_iff) lemma boundedI: assumes "\x. x \ S \ norm x \ B" shows "bounded S" using assms bounded_iff by blast lemma bounded_empty [simp]: "bounded {}" by (simp add: bounded_def) lemma bounded_subset: "bounded T \ S \ T \ bounded S" by (metis bounded_def subset_eq) lemma bounded_interior[intro]: "bounded S \ bounded(interior S)" by (metis bounded_subset interior_subset) lemma bounded_closure[intro]: assumes "bounded S" shows "bounded (closure S)" proof - from assms obtain x and a where a: "\y\S. dist x y \ a" unfolding bounded_def by auto { fix y assume "y \ closure S" then obtain f where f: "\n. f n \ S" "(f \ y) sequentially" unfolding closure_sequential by auto have "\n. f n \ S \ dist x (f n) \ a" using a by simp then have "eventually (\n. dist x (f n) \ a) sequentially" by (simp add: f(1)) then have "dist x y \ a" using Lim_dist_ubound f(2) trivial_limit_sequentially by blast } then show ?thesis unfolding bounded_def by auto qed lemma bounded_closure_image: "bounded (f ` closure S) \ bounded (f ` S)" by (simp add: bounded_subset closure_subset image_mono) lemma bounded_cball[simp,intro]: "bounded (cball x e)" unfolding bounded_def using mem_cball by blast lemma bounded_ball[simp,intro]: "bounded (ball x e)" by (metis ball_subset_cball bounded_cball bounded_subset) lemma bounded_Un[simp]: "bounded (S \ T) \ bounded S \ bounded T" by (auto simp: bounded_def) (metis Un_iff bounded_any_center le_max_iff_disj) lemma bounded_Union[intro]: "finite F \ \S\F. bounded S \ bounded (\F)" by (induct rule: finite_induct[of F]) auto lemma bounded_UN [intro]: "finite A \ \x\A. bounded (B x) \ bounded (\x\A. B x)" by auto lemma bounded_insert [simp]: "bounded (insert x S) \ bounded S" proof - have "\y\{x}. dist x y \ 0" by simp then have "bounded {x}" unfolding bounded_def by fast then show ?thesis by (metis insert_is_Un bounded_Un) qed lemma bounded_subset_ballI: "S \ ball x r \ bounded S" by (meson bounded_ball bounded_subset) lemma bounded_subset_ballD: assumes "bounded S" shows "\r. 0 < r \ S \ ball x r" proof - obtain e::real and y where "S \ cball y e" "0 \ e" using assms by (auto simp: bounded_subset_cball) then show ?thesis by (intro exI[where x="dist x y + e + 1"]) metric qed lemma finite_imp_bounded [intro]: "finite S \ bounded S" by (induct set: finite) simp_all lemma bounded_Int[intro]: "bounded S \ bounded T \ bounded (S \ T)" by (metis Int_lower1 Int_lower2 bounded_subset) lemma bounded_diff[intro]: "bounded S \ bounded (S - T)" by (metis Diff_subset bounded_subset) lemma bounded_dist_comp: assumes "bounded (f ` S)" "bounded (g ` S)" shows "bounded ((\x. dist (f x) (g x)) ` S)" proof - from assms obtain M1 M2 where *: "dist (f x) undefined \ M1" "dist undefined (g x) \ M2" if "x \ S" for x by (auto simp: bounded_any_center[of _ undefined] dist_commute) have "dist (f x) (g x) \ M1 + M2" if "x \ S" for x using *[OF that] by metric then show ?thesis by (auto intro!: boundedI) qed lemma bounded_Times: assumes "bounded s" "bounded t" shows "bounded (s \ t)" proof - obtain x y a b where "\z\s. dist x z \ a" "\z\t. dist y z \ b" using assms [unfolded bounded_def] by auto then have "\z\s \ t. dist (x, y) z \ sqrt (a\<^sup>2 + b\<^sup>2)" by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto qed subsection \Compactness\ lemma compact_imp_bounded: assumes "compact U" shows "bounded U" proof - have "compact U" "\x\U. open (ball x 1)" "U \ (\x\U. ball x 1)" using assms by auto then obtain D where D: "D \ U" "finite D" "U \ (\x\D. ball x 1)" by (metis compactE_image) from \finite D\ have "bounded (\x\D. ball x 1)" by (simp add: bounded_UN) then show "bounded U" using \U \ (\x\D. ball x 1)\ by (rule bounded_subset) qed lemma closure_Int_ball_not_empty: assumes "S \ closure T" "x \ S" "r > 0" shows "T \ ball x r \ {}" using assms centre_in_ball closure_iff_nhds_not_empty by blast lemma compact_sup_maxdistance: fixes S :: "'a::metric_space set" assumes "compact S" and "S \ {}" shows "\x\S. \y\S. \u\S. \v\S. dist u v \ dist x y" proof - have "compact (S \ S)" using \compact S\ by (intro compact_Times) moreover have "S \ S \ {}" using \S \ {}\ by auto moreover have "continuous_on (S \ S) (\x. dist (fst x) (snd x))" by (intro continuous_at_imp_continuous_on ballI continuous_intros) ultimately show ?thesis using continuous_attains_sup[of "S \ S" "\x. dist (fst x) (snd x)"] by auto qed subsubsection\Totally bounded\ lemma cauchy_def: "Cauchy S \ (\e>0. \N. \m n. m \ N \ n \ N \ dist (S m) (S n) < e)" unfolding Cauchy_def by metis proposition seq_compact_imp_totally_bounded: assumes "seq_compact S" shows "\e>0. \k. finite k \ k \ S \ S \ (\x\k. ball x e)" proof - { fix e::real assume "e > 0" assume *: "\k. finite k \ k \ S \ \ S \ (\x\k. ball x e)" let ?Q = "\x n r. r \ S \ (\m < (n::nat). \ (dist (x m) r < e))" have "\x. \n::nat. ?Q x n (x n)" proof (rule dependent_wellorder_choice) fix n x assume "\y. y < n \ ?Q x y (x y)" then have "\ S \ (\x\x ` {0..S" "z \ (\x\x ` {0..r. ?Q x n r" using z by auto qed simp then obtain x where "\n::nat. x n \ S" and x:"\n m. m < n \ \ (dist (x m) (x n) < e)" by blast then obtain l r where "l \ S" and r:"strict_mono r" and "((x \ r) \ l) sequentially" using assms by (metis seq_compact_def) then have "Cauchy (x \ r)" using LIMSEQ_imp_Cauchy by auto then obtain N::nat where "\m n. N \ m \ N \ n \ dist ((x \ r) m) ((x \ r) n) < e" unfolding cauchy_def using \e > 0\ by blast then have False using x[of "r N" "r (N+1)"] r by (auto simp: strict_mono_def) } then show ?thesis by metis qed subsubsection\Heine-Borel theorem\ proposition seq_compact_imp_Heine_Borel: fixes S :: "'a :: metric_space set" assumes "seq_compact S" shows "compact S" proof - from seq_compact_imp_totally_bounded[OF \seq_compact S\] obtain f where f: "\e>0. finite (f e) \ f e \ S \ S \ (\x\f e. ball x e)" unfolding choice_iff' .. define K where "K = (\(x, r). ball x r) ` ((\e \ \ \ {0 <..}. f e) \ \)" have "countably_compact S" using \seq_compact S\ by (rule seq_compact_imp_countably_compact) then show "compact S" proof (rule countably_compact_imp_compact) show "countable K" unfolding K_def using f by (auto intro: countable_finite countable_subset countable_rat intro!: countable_image countable_SIGMA countable_UN) show "\b\K. open b" by (auto simp: K_def) next fix T x assume T: "open T" "x \ T" and x: "x \ S" from openE[OF T] obtain e where "0 < e" "ball x e \ T" by auto then have "0 < e/2" "ball x (e/2) \ T" by auto from Rats_dense_in_real[OF \0 < e/2\] obtain r where "r \ \" "0 < r" "r < e/2" by auto from f[rule_format, of r] \0 < r\ \x \ S\ obtain k where "k \ f r" "x \ ball k r" by auto from \r \ \\ \0 < r\ \k \ f r\ have "ball k r \ K" by (auto simp: K_def) then show "\b\K. x \ b \ b \ S \ T" proof (rule bexI[rotated], safe) fix y assume "y \ ball k r" with \r < e/2\ \x \ ball k r\ have "dist x y < e" by (intro dist_triangle_half_r [of k _ e]) (auto simp: dist_commute) with \ball x e \ T\ show "y \ T" by auto next show "x \ ball k r" by fact qed qed qed proposition compact_eq_seq_compact_metric: "compact (S :: 'a::metric_space set) \ seq_compact S" using compact_imp_seq_compact seq_compact_imp_Heine_Borel by blast proposition compact_def: \ \this is the definition of compactness in HOL Light\ "compact (S :: 'a::metric_space set) \ (\f. (\n. f n \ S) \ (\l\S. \r::nat\nat. strict_mono r \ (f \ r) \ l))" unfolding compact_eq_seq_compact_metric seq_compact_def by auto subsubsection \Complete the chain of compactness variants\ proposition compact_eq_Bolzano_Weierstrass: fixes S :: "'a::metric_space set" shows "compact S \ (\T. infinite T \ T \ S \ (\x \ S. x islimpt T))" using Bolzano_Weierstrass_imp_seq_compact Heine_Borel_imp_Bolzano_Weierstrass compact_eq_seq_compact_metric by blast proposition Bolzano_Weierstrass_imp_bounded: "(\T. \infinite T; T \ S\ \ (\x \ S. x islimpt T)) \ bounded S" using compact_imp_bounded unfolding compact_eq_Bolzano_Weierstrass by metis subsection \Banach fixed point theorem\ theorem banach_fix:\ \TODO: rename to \Banach_fix\\ assumes s: "complete s" "s \ {}" and c: "0 \ c" "c < 1" and f: "f ` s \ s" and lipschitz: "\x\s. \y\s. dist (f x) (f y) \ c * dist x y" shows "\!x\s. f x = x" proof - from c have "1 - c > 0" by simp from s(2) obtain z0 where z0: "z0 \ s" by blast define z where "z n = (f ^^ n) z0" for n with f z0 have z_in_s: "z n \ s" for n :: nat by (induct n) auto define d where "d = dist (z 0) (z 1)" have fzn: "f (z n) = z (Suc n)" for n by (simp add: z_def) have cf_z: "dist (z n) (z (Suc n)) \ (c ^ n) * d" for n :: nat proof (induct n) case 0 then show ?case by (simp add: d_def) next case (Suc m) with \0 \ c\ have "c * dist (z m) (z (Suc m)) \ c ^ Suc m * d" using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp then show ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] by (simp add: fzn mult_le_cancel_left) qed have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \ (c ^ m) * d * (1 - c ^ n)" for n m :: nat proof (induct n) case 0 show ?case by simp next case (Suc k) from c have "(1 - c) * dist (z m) (z (m + Suc k)) \ (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" by (simp add: dist_triangle) also from c cf_z[of "m + k"] have "\ \ (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" by simp also from Suc have "\ \ c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" by (simp add: field_simps) also have "\ = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" by (simp add: power_add field_simps) also from c have "\ \ (c ^ m) * d * (1 - c ^ Suc k)" by (simp add: field_simps) finally show ?case by simp qed have "\N. \m n. N \ m \ N \ n \ dist (z m) (z n) < e" if "e > 0" for e proof (cases "d = 0") case True from \1 - c > 0\ have "(1 - c) * x \ 0 \ x \ 0" for x - by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1) + by (simp add: mult_le_0_iff) with c cf_z2[of 0] True have "z n = z0" for n by (simp add: z_def) with \e > 0\ show ?thesis by simp next case False with zero_le_dist[of "z 0" "z 1"] have "d > 0" by (metis d_def less_le) with \1 - c > 0\ \e > 0\ have "0 < e * (1 - c) / d" by simp with c obtain N where N: "c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto have *: "dist (z m) (z n) < e" if "m > n" and as: "m \ N" "n \ N" for m n :: nat proof - from c \n \ N\ have *: "c ^ n \ c ^ N" using power_decreasing[OF \n\N\, of c] by simp from c \m > n\ have "1 - c ^ (m - n) > 0" using power_strict_mono[of c 1 "m - n"] by simp with \d > 0\ \0 < 1 - c\ have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" by simp from cf_z2[of n "m - n"] \m > n\ have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" by (simp add: pos_le_divide_eq[OF \1 - c > 0\] mult.commute dist_commute) also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" using mult_right_mono[OF * order_less_imp_le[OF **]] by (simp add: mult.assoc) also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc) also from c \d > 0\ \1 - c > 0\ have "\ = e * (1 - c ^ (m - n))" by simp also from c \1 - c ^ (m - n) > 0\ \e > 0\ have "\ \ e" using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto finally show ?thesis by simp qed have "dist (z n) (z m) < e" if "N \ m" "N \ n" for m n :: nat proof (cases "n = m") case True with \e > 0\ show ?thesis by simp next case False with *[of n m] *[of m n] and that show ?thesis by (auto simp: dist_commute nat_neq_iff) qed then show ?thesis by auto qed then have "Cauchy z" by (simp add: cauchy_def) then obtain x where "x\s" and x:"(z \ x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto define e where "e = dist (f x) x" have "e = 0" proof (rule ccontr) assume "e \ 0" then have "e > 0" unfolding e_def using zero_le_dist[of "f x" x] by (metis dist_eq_0_iff dist_nz e_def) then obtain N where N:"\n\N. dist (z n) x < e/2" using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto then have N':"dist (z N) x < e/2" by auto have *: "c * dist (z N) x \ dist (z N) x" unfolding mult_le_cancel_right2 using zero_le_dist[of "z N" x] and c by (metis dist_eq_0_iff dist_nz order_less_asym less_le) have "dist (f (z N)) (f x) \ c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] using z_in_s[of N] \x\s\ using c by auto also have "\ < e/2" using N' and c using * by auto finally show False unfolding fzn using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] unfolding e_def by auto qed then have "f x = x" by (auto simp: e_def) moreover have "y = x" if "f y = y" "y \ s" for y proof - from \x \ s\ \f x = x\ that have "dist x y \ c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp with c and zero_le_dist[of x y] have "dist x y = 0" by (simp add: mult_le_cancel_right1) then show ?thesis by simp qed ultimately show ?thesis using \x\s\ by blast qed subsection \Edelstein fixed point theorem\ theorem Edelstein_fix: fixes S :: "'a::metric_space set" assumes S: "compact S" "S \ {}" and gs: "(g ` S) \ S" and dist: "\x\S. \y\S. x \ y \ dist (g x) (g y) < dist x y" shows "\!x\S. g x = x" proof - let ?D = "(\x. (x, x)) ` S" have D: "compact ?D" "?D \ {}" by (rule compact_continuous_image) (auto intro!: S continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within) have "\x y e. x \ S \ y \ S \ 0 < e \ dist y x < e \ dist (g y) (g x) < e" using dist by fastforce then have "continuous_on S g" by (auto simp: continuous_on_iff) then have cont: "continuous_on ?D (\x. dist ((g \ fst) x) (snd x))" unfolding continuous_on_eq_continuous_within by (intro continuous_dist ballI continuous_within_compose) (auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image) obtain a where "a \ S" and le: "\x. x \ S \ dist (g a) a \ dist (g x) x" using continuous_attains_inf[OF D cont] by auto have "g a = a" proof (rule ccontr) assume "g a \ a" with \a \ S\ gs have "dist (g (g a)) (g a) < dist (g a) a" by (intro dist[rule_format]) auto moreover have "dist (g a) a \ dist (g (g a)) (g a)" using \a \ S\ gs by (intro le) auto ultimately show False by auto qed moreover have "\x. x \ S \ g x = x \ x = a" using dist[THEN bspec[where x=a]] \g a = a\ and \a\S\ by auto ultimately show "\!x\S. g x = x" using \a \ S\ by blast qed subsection \The diameter of a set\ definition\<^marker>\tag important\ diameter :: "'a::metric_space set \ real" where "diameter S = (if S = {} then 0 else SUP (x,y)\S\S. dist x y)" lemma diameter_empty [simp]: "diameter{} = 0" by (auto simp: diameter_def) lemma diameter_singleton [simp]: "diameter{x} = 0" by (auto simp: diameter_def) lemma diameter_le: assumes "S \ {} \ 0 \ d" and no: "\x y. \x \ S; y \ S\ \ norm(x - y) \ d" shows "diameter S \ d" using assms by (auto simp: dist_norm diameter_def intro: cSUP_least) lemma diameter_bounded_bound: fixes S :: "'a :: metric_space set" assumes S: "bounded S" "x \ S" "y \ S" shows "dist x y \ diameter S" proof - from S obtain z d where z: "\x. x \ S \ dist z x \ d" unfolding bounded_def by auto have "bdd_above (case_prod dist ` (S\S))" proof (intro bdd_aboveI, safe) fix a b assume "a \ S" "b \ S" with z[of a] z[of b] dist_triangle[of a b z] show "dist a b \ 2 * d" by (simp add: dist_commute) qed moreover have "(x,y) \ S\S" using S by auto ultimately have "dist x y \ (SUP (x,y)\S\S. dist x y)" by (rule cSUP_upper2) simp with \x \ S\ show ?thesis by (auto simp: diameter_def) qed lemma diameter_lower_bounded: fixes S :: "'a :: metric_space set" assumes S: "bounded S" and d: "0 < d" "d < diameter S" shows "\x\S. \y\S. d < dist x y" proof (rule ccontr) assume contr: "\ ?thesis" moreover have "S \ {}" using d by (auto simp: diameter_def) ultimately have "diameter S \ d" by (auto simp: not_less diameter_def intro!: cSUP_least) with \d < diameter S\ show False by auto qed lemma diameter_bounded: assumes "bounded S" shows "\x\S. \y\S. dist x y \ diameter S" and "\d>0. d < diameter S \ (\x\S. \y\S. dist x y > d)" using diameter_bounded_bound[of S] diameter_lower_bounded[of S] assms by auto lemma bounded_two_points: "bounded S \ (\e. \x\S. \y\S. dist x y \ e)" by (meson bounded_def diameter_bounded(1)) lemma diameter_compact_attained: assumes "compact S" and "S \ {}" shows "\x\S. \y\S. dist x y = diameter S" proof - have b: "bounded S" using assms(1) by (rule compact_imp_bounded) then obtain x y where xys: "x\S" "y\S" and xy: "\u\S. \v\S. dist u v \ dist x y" using compact_sup_maxdistance[OF assms] by auto then have "diameter S \ dist x y" unfolding diameter_def apply clarsimp apply (rule cSUP_least, fast+) done then show ?thesis by (metis b diameter_bounded_bound order_antisym xys) qed lemma diameter_ge_0: assumes "bounded S" shows "0 \ diameter S" by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl) lemma diameter_subset: assumes "S \ T" "bounded T" shows "diameter S \ diameter T" proof (cases "S = {} \ T = {}") case True with assms show ?thesis by (force simp: diameter_ge_0) next case False then have "bdd_above ((\x. case x of (x, xa) \ dist x xa) ` (T \ T))" using \bounded T\ diameter_bounded_bound by (force simp: bdd_above_def) with False \S \ T\ show ?thesis apply (simp add: diameter_def) apply (rule cSUP_subset_mono, auto) done qed lemma diameter_closure: assumes "bounded S" shows "diameter(closure S) = diameter S" proof (rule order_antisym) have "False" if "diameter S < diameter (closure S)" proof - define d where "d = diameter(closure S) - diameter(S)" have "d > 0" using that by (simp add: d_def) then have "diameter(closure(S)) - d / 2 < diameter(closure(S))" by simp have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2" by (simp add: d_def field_split_simps) have bocl: "bounded (closure S)" using assms by blast moreover have "0 \ diameter S" using assms diameter_ge_0 by blast ultimately obtain x y where "x \ closure S" "y \ closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y" using diameter_bounded(2) [OF bocl, rule_format, of "diameter(closure(S)) - d / 2"] \d > 0\ d_def by auto then obtain x' y' where x'y': "x' \ S" "dist x' x < d/4" "y' \ S" "dist y' y < d/4" using closure_approachable by (metis \0 < d\ zero_less_divide_iff zero_less_numeral) then have "dist x' y' \ diameter S" using assms diameter_bounded_bound by blast with x'y' have "dist x y \ d / 4 + diameter S + d / 4" by (meson add_mono_thms_linordered_semiring(1) dist_triangle dist_triangle3 less_eq_real_def order_trans) then show ?thesis using xy d_def by linarith qed then show "diameter (closure S) \ diameter S" by fastforce next show "diameter S \ diameter (closure S)" by (simp add: assms bounded_closure closure_subset diameter_subset) qed proposition Lebesgue_number_lemma: assumes "compact S" "\ \ {}" "S \ \\" and ope: "\B. B \ \ \ open B" obtains \ where "0 < \" "\T. \T \ S; diameter T < \\ \ \B \ \. T \ B" proof (cases "S = {}") case True then show ?thesis by (metis \\ \ {}\ zero_less_one empty_subsetI equals0I subset_trans that) next case False { fix x assume "x \ S" then obtain C where C: "x \ C" "C \ \" using \S \ \\\ by blast then obtain r where r: "r>0" "ball x (2*r) \ C" by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff) then have "\r C. r > 0 \ ball x (2*r) \ C \ C \ \" using C by blast } then obtain r where r: "\x. x \ S \ r x > 0 \ (\C \ \. ball x (2*r x) \ C)" by metis then have "S \ (\x \ S. ball x (r x))" by auto then obtain \ where "finite \" "S \ \\" and \: "\ \ (\x. ball x (r x)) ` S" by (rule compactE [OF \compact S\]) auto then obtain S0 where "S0 \ S" "finite S0" and S0: "\ = (\x. ball x (r x)) ` S0" by (meson finite_subset_image) then have "S0 \ {}" using False \S \ \\\ by auto define \ where "\ = Inf (r ` S0)" have "\ > 0" using \finite S0\ \S0 \ S\ \S0 \ {}\ r by (auto simp: \_def finite_less_Inf_iff) show ?thesis proof show "0 < \" by (simp add: \0 < \\) show "\B \ \. T \ B" if "T \ S" and dia: "diameter T < \" for T proof (cases "T = {}") case True then show ?thesis using \\ \ {}\ by blast next case False then obtain y where "y \ T" by blast then have "y \ S" using \T \ S\ by auto then obtain x where "x \ S0" and x: "y \ ball x (r x)" using \S \ \\\ S0 that by blast have "ball y \ \ ball y (r x)" by (metis \_def \S0 \ {}\ \finite S0\ \x \ S0\ empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball) also have "... \ ball x (2*r x)" using x by metric finally obtain C where "C \ \" "ball y \ \ C" by (meson r \S0 \ S\ \x \ S0\ dual_order.trans subsetCE) have "bounded T" using \compact S\ bounded_subset compact_imp_bounded \T \ S\ by blast then have "T \ ball y \" using \y \ T\ dia diameter_bounded_bound by fastforce then show ?thesis apply (rule_tac x=C in bexI) using \ball y \ \ C\ \C \ \\ by auto qed qed qed subsection \Metric spaces with the Heine-Borel property\ text \ A metric space (or topological vector space) is said to have the Heine-Borel property if every closed and bounded subset is compact. \ class heine_borel = metric_space + assumes bounded_imp_convergent_subsequence: "bounded (range f) \ \l r. strict_mono (r::nat\nat) \ ((f \ r) \ l) sequentially" proposition bounded_closed_imp_seq_compact: fixes S::"'a::heine_borel set" assumes "bounded S" and "closed S" shows "seq_compact S" proof (unfold seq_compact_def, clarify) fix f :: "nat \ 'a" assume f: "\n. f n \ S" with \bounded S\ have "bounded (range f)" by (auto intro: bounded_subset) obtain l r where r: "strict_mono (r :: nat \ nat)" and l: "((f \ r) \ l) sequentially" using bounded_imp_convergent_subsequence [OF \bounded (range f)\] by auto from f have fr: "\n. (f \ r) n \ S" by simp have "l \ S" using \closed S\ fr l by (rule closed_sequentially) show "\l\S. \r. strict_mono r \ ((f \ r) \ l) sequentially" using \l \ S\ r l by blast qed lemma compact_eq_bounded_closed: fixes S :: "'a::heine_borel set" shows "compact S \ bounded S \ closed S" using bounded_closed_imp_seq_compact compact_eq_seq_compact_metric compact_imp_bounded compact_imp_closed by auto lemma compact_Inter: fixes \ :: "'a :: heine_borel set set" assumes com: "\S. S \ \ \ compact S" and "\ \ {}" shows "compact(\ \)" using assms by (meson Inf_lower all_not_in_conv bounded_subset closed_Inter compact_eq_bounded_closed) lemma compact_closure [simp]: fixes S :: "'a::heine_borel set" shows "compact(closure S) \ bounded S" by (meson bounded_closure bounded_subset closed_closure closure_subset compact_eq_bounded_closed) instance\<^marker>\tag important\ real :: heine_borel proof fix f :: "nat \ real" assume f: "bounded (range f)" obtain r :: "nat \ nat" where r: "strict_mono r" "monoseq (f \ r)" unfolding comp_def by (metis seq_monosub) then have "Bseq (f \ r)" unfolding Bseq_eq_bounded using f by (metis BseqI' bounded_iff comp_apply rangeI) with r show "\l r. strict_mono r \ (f \ r) \ l" using Bseq_monoseq_convergent[of "f \ r"] by (auto simp: convergent_def) qed lemma compact_lemma_general: fixes f :: "nat \ 'a" fixes proj::"'a \ 'b \ 'c::heine_borel" (infixl "proj" 60) fixes unproj:: "('b \ 'c) \ 'a" assumes finite_basis: "finite basis" assumes bounded_proj: "\k. k \ basis \ bounded ((\x. x proj k) ` range f)" assumes proj_unproj: "\e k. k \ basis \ (unproj e) proj k = e k" assumes unproj_proj: "\x. unproj (\k. x proj k) = x" shows "\d\basis. \l::'a. \ r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof safe fix d :: "'b set" assume d: "d \ basis" with finite_basis have "finite d" by (blast intro: finite_subset) from this d show "\l::'a. \r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) proj i) (l proj i) < e) sequentially)" proof (induct d) case empty then show ?case unfolding strict_mono_def by auto next case (insert k d) have k[intro]: "k \ basis" using insert by auto have s': "bounded ((\x. x proj k) ` range f)" using k by (rule bounded_proj) obtain l1::"'a" and r1 where r1: "strict_mono r1" and lr1: "\e > 0. eventually (\n. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially" using insert(3) using insert(4) by auto have f': "\n. f (r1 n) proj k \ (\x. x proj k) ` range f" by simp have "bounded (range (\i. f (r1 i) proj k))" by (metis (lifting) bounded_subset f' image_subsetI s') then obtain l2 r2 where r2:"strict_mono r2" and lr2:"((\i. f (r1 (r2 i)) proj k) \ l2) sequentially" using bounded_imp_convergent_subsequence[of "\i. f (r1 i) proj k"] by (auto simp: o_def) define r where "r = r1 \ r2" have r:"strict_mono r" using r1 and r2 unfolding r_def o_def strict_mono_def by auto moreover define l where "l = unproj (\i. if i = k then l2 else l1 proj i)" { fix e::real assume "e > 0" from lr1 \e > 0\ have N1: "eventually (\n. \i\d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially" by blast from lr2 \e > 0\ have N2:"eventually (\n. dist (f (r1 (r2 n)) proj k) l2 < e) sequentially" by (rule tendstoD) from r2 N1 have N1': "eventually (\n. \i\d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e) sequentially" by (rule eventually_subseq) have "eventually (\n. \i\(insert k d). dist (f (r n) proj i) (l proj i) < e) sequentially" using N1' N2 by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def proj_unproj) } ultimately show ?case by auto qed qed lemma bounded_fst: "bounded s \ bounded (fst ` s)" unfolding bounded_def by (metis (erased, hide_lams) dist_fst_le image_iff order_trans) lemma bounded_snd: "bounded s \ bounded (snd ` s)" unfolding bounded_def by (metis (no_types, hide_lams) dist_snd_le image_iff order.trans) instance\<^marker>\tag important\ prod :: (heine_borel, heine_borel) heine_borel proof fix f :: "nat \ 'a \ 'b" assume f: "bounded (range f)" then have "bounded (fst ` range f)" by (rule bounded_fst) then have s1: "bounded (range (fst \ f))" by (simp add: image_comp) obtain l1 r1 where r1: "strict_mono r1" and l1: "(\n. fst (f (r1 n))) \ l1" using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast from f have s2: "bounded (range (snd \ f \ r1))" by (auto simp: image_comp intro: bounded_snd bounded_subset) obtain l2 r2 where r2: "strict_mono r2" and l2: "((\n. snd (f (r1 (r2 n)))) \ l2) sequentially" using bounded_imp_convergent_subsequence [OF s2] unfolding o_def by fast have l1': "((\n. fst (f (r1 (r2 n)))) \ l1) sequentially" using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def . have l: "((f \ (r1 \ r2)) \ (l1, l2)) sequentially" using tendsto_Pair [OF l1' l2] unfolding o_def by simp have r: "strict_mono (r1 \ r2)" using r1 r2 unfolding strict_mono_def by simp show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" using l r by fast qed subsection \Completeness\ proposition (in metric_space) completeI: assumes "\f. \n. f n \ s \ Cauchy f \ \l\s. f \ l" shows "complete s" using assms unfolding complete_def by fast proposition (in metric_space) completeE: assumes "complete s" and "\n. f n \ s" and "Cauchy f" obtains l where "l \ s" and "f \ l" using assms unfolding complete_def by fast (* TODO: generalize to uniform spaces *) lemma compact_imp_complete: fixes s :: "'a::metric_space set" assumes "compact s" shows "complete s" proof - { fix f assume as: "(\n::nat. f n \ s)" "Cauchy f" from as(1) obtain l r where lr: "l\s" "strict_mono r" "(f \ r) \ l" using assms unfolding compact_def by blast note lr' = seq_suble [OF lr(2)] { fix e :: real assume "e > 0" from as(2) obtain N where N:"\m n. N \ m \ N \ n \ dist (f m) (f n) < e/2" unfolding cauchy_def using \e > 0\ apply (erule_tac x="e/2" in allE, auto) done from lr(3)[unfolded lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\n\M. dist ((f \ r) n) l < e/2" using \e > 0\ by auto { fix n :: nat assume n: "n \ max N M" have "dist ((f \ r) n) l < e/2" using n M by auto moreover have "r n \ N" using lr'[of n] n by auto then have "dist (f n) ((f \ r) n) < e/2" using N and n by auto ultimately have "dist (f n) l < e" using n M by metric } then have "\N. \n\N. dist (f n) l < e" by blast } then have "\l\s. (f \ l) sequentially" using \l\s\ unfolding lim_sequentially by auto } then show ?thesis unfolding complete_def by auto qed proposition compact_eq_totally_bounded: "compact s \ complete s \ (\e>0. \k. finite k \ s \ (\x\k. ball x e))" (is "_ \ ?rhs") proof assume assms: "?rhs" then obtain k where k: "\e. 0 < e \ finite (k e)" "\e. 0 < e \ s \ (\x\k e. ball x e)" by (auto simp: choice_iff') show "compact s" proof cases assume "s = {}" then show "compact s" by (simp add: compact_def) next assume "s \ {}" show ?thesis unfolding compact_def proof safe fix f :: "nat \ 'a" assume f: "\n. f n \ s" define e where "e n = 1 / (2 * Suc n)" for n then have [simp]: "\n. 0 < e n" by auto define B where "B n U = (SOME b. infinite {n. f n \ b} \ (\x. b \ ball x (e n) \ U))" for n U { fix n U assume "infinite {n. f n \ U}" then have "\b\k (e n). infinite {i\{n. f n \ U}. f i \ ball b (e n)}" using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq) then obtain a where "a \ k (e n)" "infinite {i \ {n. f n \ U}. f i \ ball a (e n)}" .. then have "\b. infinite {i. f i \ b} \ (\x. b \ ball x (e n) \ U)" by (intro exI[of _ "ball a (e n) \ U"] exI[of _ a]) (auto simp: ac_simps) from someI_ex[OF this] have "infinite {i. f i \ B n U}" "\x. B n U \ ball x (e n) \ U" unfolding B_def by auto } note B = this define F where "F = rec_nat (B 0 UNIV) B" { fix n have "infinite {i. f i \ F n}" by (induct n) (auto simp: F_def B) } then have F: "\n. \x. F (Suc n) \ ball x (e n) \ F n" using B by (simp add: F_def) then have F_dec: "\m n. m \ n \ F n \ F m" using decseq_SucI[of F] by (auto simp: decseq_def) obtain sel where sel: "\k i. i < sel k i" "\k i. f (sel k i) \ F k" proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI) fix k i have "infinite ({n. f n \ F k} - {.. i})" using \infinite {n. f n \ F k}\ by auto from infinite_imp_nonempty[OF this] show "\x>i. f x \ F k" by (simp add: set_eq_iff not_le conj_commute) qed define t where "t = rec_nat (sel 0 0) (\n i. sel (Suc n) i)" have "strict_mono t" unfolding strict_mono_Suc_iff by (simp add: t_def sel) moreover have "\i. (f \ t) i \ s" using f by auto moreover have t: "(f \ t) n \ F n" for n by (cases n) (simp_all add: t_def sel) have "Cauchy (f \ t)" proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE) fix r :: real and N n m assume "1 / Suc N < r" "Suc N \ n" "Suc N \ m" then have "(f \ t) n \ F (Suc N)" "(f \ t) m \ F (Suc N)" "2 * e N < r" using F_dec t by (auto simp: e_def field_simps) with F[of N] obtain x where "dist x ((f \ t) n) < e N" "dist x ((f \ t) m) < e N" by (auto simp: subset_eq) with \2 * e N < r\ show "dist ((f \ t) m) ((f \ t) n) < r" by metric qed ultimately show "\l\s. \r. strict_mono r \ (f \ r) \ l" using assms unfolding complete_def by blast qed qed qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded) lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)" proof - from assms obtain N :: nat where "\m n. N \ m \ N \ n \ dist (s m) (s n) < 1" unfolding cauchy_def by force then have N:"\n. N \ n \ dist (s N) (s n) < 1" by auto moreover have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto then obtain a where a:"\x\s ` {0..N}. dist (s N) x \ a" unfolding bounded_any_center [where a="s N"] by auto ultimately show "?thesis" unfolding bounded_any_center [where a="s N"] apply (rule_tac x="max a 1" in exI, auto) apply (erule_tac x=y in allE) apply (erule_tac x=y in ballE, auto) done qed instance heine_borel < complete_space proof fix f :: "nat \ 'a" assume "Cauchy f" then have "bounded (range f)" by (rule cauchy_imp_bounded) then have "compact (closure (range f))" unfolding compact_eq_bounded_closed by auto then have "complete (closure (range f))" by (rule compact_imp_complete) moreover have "\n. f n \ closure (range f)" using closure_subset [of "range f"] by auto ultimately have "\l\closure (range f). (f \ l) sequentially" using \Cauchy f\ unfolding complete_def by auto then show "convergent f" unfolding convergent_def by auto qed lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)" proof (rule completeI) fix f :: "nat \ 'a" assume "Cauchy f" then have "convergent f" by (rule Cauchy_convergent) then show "\l\UNIV. f \ l" unfolding convergent_def by simp qed lemma complete_imp_closed: fixes S :: "'a::metric_space set" assumes "complete S" shows "closed S" proof (unfold closed_sequential_limits, clarify) fix f x assume "\n. f n \ S" and "f \ x" from \f \ x\ have "Cauchy f" by (rule LIMSEQ_imp_Cauchy) with \complete S\ and \\n. f n \ S\ obtain l where "l \ S" and "f \ l" by (rule completeE) from \f \ x\ and \f \ l\ have "x = l" by (rule LIMSEQ_unique) with \l \ S\ show "x \ S" by simp qed lemma complete_Int_closed: fixes S :: "'a::metric_space set" assumes "complete S" and "closed t" shows "complete (S \ t)" proof (rule completeI) fix f assume "\n. f n \ S \ t" and "Cauchy f" then have "\n. f n \ S" and "\n. f n \ t" by simp_all from \complete S\ obtain l where "l \ S" and "f \ l" using \\n. f n \ S\ and \Cauchy f\ by (rule completeE) from \closed t\ and \\n. f n \ t\ and \f \ l\ have "l \ t" by (rule closed_sequentially) with \l \ S\ and \f \ l\ show "\l\S \ t. f \ l" by fast qed lemma complete_closed_subset: fixes S :: "'a::metric_space set" assumes "closed S" and "S \ t" and "complete t" shows "complete S" using assms complete_Int_closed [of t S] by (simp add: Int_absorb1) lemma complete_eq_closed: fixes S :: "('a::complete_space) set" shows "complete S \ closed S" proof assume "closed S" then show "complete S" using subset_UNIV complete_UNIV by (rule complete_closed_subset) next assume "complete S" then show "closed S" by (rule complete_imp_closed) qed lemma convergent_eq_Cauchy: fixes S :: "nat \ 'a::complete_space" shows "(\l. (S \ l) sequentially) \ Cauchy S" unfolding Cauchy_convergent_iff convergent_def .. lemma convergent_imp_bounded: fixes S :: "nat \ 'a::metric_space" shows "(S \ l) sequentially \ bounded (range S)" by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy) lemma frontier_subset_compact: fixes S :: "'a::heine_borel set" shows "compact S \ frontier S \ S" using frontier_subset_closed compact_eq_bounded_closed by blast lemma continuous_closed_imp_Cauchy_continuous: fixes S :: "('a::complete_space) set" shows "\continuous_on S f; closed S; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" apply (simp add: complete_eq_closed [symmetric] continuous_on_sequentially) by (meson LIMSEQ_imp_Cauchy complete_def) lemma banach_fix_type: fixes f::"'a::complete_space\'a" assumes c:"0 \ c" "c < 1" and lipschitz:"\x. \y. dist (f x) (f y) \ c * dist x y" shows "\!x. (f x = x)" using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f] by auto subsection\<^marker>\tag unimportant\\ Finite intersection property\ text\Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\ lemma closed_imp_fip: fixes S :: "'a::heine_borel set" assumes "closed S" and T: "T \ \" "bounded T" and clof: "\T. T \ \ \ closed T" and none: "\\'. \finite \'; \' \ \\ \ S \ \\' \ {}" shows "S \ \\ \ {}" proof - have "compact (S \ T)" using \closed S\ clof compact_eq_bounded_closed T by blast then have "(S \ T) \ \\ \ {}" apply (rule compact_imp_fip) apply (simp add: clof) by (metis Int_assoc complete_lattice_class.Inf_insert finite_insert insert_subset none \T \ \\) then show ?thesis by blast qed lemma closed_imp_fip_compact: fixes S :: "'a::heine_borel set" shows "\closed S; \T. T \ \ \ compact T; \\'. \finite \'; \' \ \\ \ S \ \\' \ {}\ \ S \ \\ \ {}" by (metis Inf_greatest closed_imp_fip compact_eq_bounded_closed empty_subsetI finite.emptyI inf.orderE) lemma closed_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes "closed S" "T \ \" "bounded T" and "\T. T \ \ \ closed T" and "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" proof - have "UNIV \ \\ \ {}" using assms closed_imp_fip [OF closed_UNIV] by auto then show ?thesis by simp qed lemma compact_fip_Heine_Borel: fixes \ :: "'a::heine_borel set set" assumes clof: "\T. T \ \ \ compact T" and none: "\\'. \finite \'; \' \ \\ \ \\' \ {}" shows "\\ \ {}" by (metis InterI all_not_in_conv clof closed_fip_Heine_Borel compact_eq_bounded_closed none) lemma compact_sequence_with_limit: fixes f :: "nat \ 'a::heine_borel" shows "(f \ l) sequentially \ compact (insert l (range f))" apply (simp add: compact_eq_bounded_closed, auto) apply (simp add: convergent_imp_bounded) by (simp add: closed_limpt islimpt_insert sequence_unique_limpt) subsection \Properties of Balls and Spheres\ lemma compact_cball[simp]: fixes x :: "'a::heine_borel" shows "compact (cball x e)" using compact_eq_bounded_closed bounded_cball closed_cball by blast lemma compact_frontier_bounded[intro]: fixes S :: "'a::heine_borel set" shows "bounded S \ compact (frontier S)" unfolding frontier_def using compact_eq_bounded_closed by blast lemma compact_frontier[intro]: fixes S :: "'a::heine_borel set" shows "compact S \ compact (frontier S)" using compact_eq_bounded_closed compact_frontier_bounded by blast subsection \Distance from a Set\ lemma distance_attains_sup: assumes "compact s" "s \ {}" shows "\x\s. \y\s. dist a y \ dist a x" proof (rule continuous_attains_sup [OF assms]) { fix x assume "x\s" have "(dist a \ dist a x) (at x within s)" by (intro tendsto_dist tendsto_const tendsto_ident_at) } then show "continuous_on s (dist a)" unfolding continuous_on .. qed text \For \emph{minimal} distance, we only need closure, not compactness.\ lemma distance_attains_inf: fixes a :: "'a::heine_borel" assumes "closed s" and "s \ {}" obtains x where "x\s" "\y. y \ s \ dist a x \ dist a y" proof - from assms obtain b where "b \ s" by auto let ?B = "s \ cball a (dist b a)" have "?B \ {}" using \b \ s\ by (auto simp: dist_commute) moreover have "continuous_on ?B (dist a)" by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_ident continuous_const) moreover have "compact ?B" by (intro closed_Int_compact \closed s\ compact_cball) ultimately obtain x where "x \ ?B" "\y\?B. dist a x \ dist a y" by (metis continuous_attains_inf) with that show ?thesis by fastforce qed subsection \Infimum Distance\ definition\<^marker>\tag important\ "infdist x A = (if A = {} then 0 else INF a\A. dist x a)" lemma bdd_below_image_dist[intro, simp]: "bdd_below (dist x ` A)" by (auto intro!: zero_le_dist) lemma infdist_notempty: "A \ {} \ infdist x A = (INF a\A. dist x a)" by (simp add: infdist_def) lemma infdist_nonneg: "0 \ infdist x A" by (auto simp: infdist_def intro: cINF_greatest) lemma infdist_le: "a \ A \ infdist x A \ dist x a" by (auto intro: cINF_lower simp add: infdist_def) lemma infdist_le2: "a \ A \ dist x a \ d \ infdist x A \ d" by (auto intro!: cINF_lower2 simp add: infdist_def) lemma infdist_zero[simp]: "a \ A \ infdist a A = 0" by (auto intro!: antisym infdist_nonneg infdist_le2) lemma infdist_Un_min: assumes "A \ {}" "B \ {}" shows "infdist x (A \ B) = min (infdist x A) (infdist x B)" using assms by (simp add: infdist_def cINF_union inf_real_def) lemma infdist_triangle: "infdist x A \ infdist y A + dist x y" proof (cases "A = {}") case True then show ?thesis by (simp add: infdist_def) next case False then obtain a where "a \ A" by auto have "infdist x A \ Inf {dist x y + dist y a |a. a \ A}" proof (rule cInf_greatest) from \A \ {}\ show "{dist x y + dist y a |a. a \ A} \ {}" by simp fix d assume "d \ {dist x y + dist y a |a. a \ A}" then obtain a where d: "d = dist x y + dist y a" "a \ A" by auto show "infdist x A \ d" unfolding infdist_notempty[OF \A \ {}\] proof (rule cINF_lower2) show "a \ A" by fact show "dist x a \ d" unfolding d by (rule dist_triangle) qed simp qed also have "\ = dist x y + infdist y A" proof (rule cInf_eq, safe) fix a assume "a \ A" then show "dist x y + infdist y A \ dist x y + dist y a" by (auto intro: infdist_le) next fix i assume inf: "\d. d \ {dist x y + dist y a |a. a \ A} \ i \ d" then have "i - dist x y \ infdist y A" unfolding infdist_notempty[OF \A \ {}\] using \a \ A\ by (intro cINF_greatest) (auto simp: field_simps) then show "i \ dist x y + infdist y A" by simp qed finally show ?thesis by simp qed lemma infdist_triangle_abs: "\infdist x A - infdist y A\ \ dist x y" by (metis (full_types) abs_diff_le_iff diff_le_eq dist_commute infdist_triangle) lemma in_closure_iff_infdist_zero: assumes "A \ {}" shows "x \ closure A \ infdist x A = 0" proof assume "x \ closure A" show "infdist x A = 0" proof (rule ccontr) assume "infdist x A \ 0" with infdist_nonneg[of x A] have "infdist x A > 0" by auto then have "ball x (infdist x A) \ closure A = {}" apply auto apply (metis \x \ closure A\ closure_approachable dist_commute infdist_le not_less) done then have "x \ closure A" by (metis \0 < infdist x A\ centre_in_ball disjoint_iff_not_equal) then show False using \x \ closure A\ by simp qed next assume x: "infdist x A = 0" then obtain a where "a \ A" by atomize_elim (metis all_not_in_conv assms) show "x \ closure A" unfolding closure_approachable apply safe proof (rule ccontr) fix e :: real assume "e > 0" assume "\ (\y\A. dist y x < e)" then have "infdist x A \ e" using \a \ A\ unfolding infdist_def by (force simp: dist_commute intro: cINF_greatest) with x \e > 0\ show False by auto qed qed lemma in_closed_iff_infdist_zero: assumes "closed A" "A \ {}" shows "x \ A \ infdist x A = 0" proof - have "x \ closure A \ infdist x A = 0" by (rule in_closure_iff_infdist_zero) fact with assms show ?thesis by simp qed lemma infdist_pos_not_in_closed: assumes "closed S" "S \ {}" "x \ S" shows "infdist x S > 0" using in_closed_iff_infdist_zero[OF assms(1) assms(2), of x] assms(3) infdist_nonneg le_less by fastforce lemma infdist_attains_inf: fixes X::"'a::heine_borel set" assumes "closed X" assumes "X \ {}" obtains x where "x \ X" "infdist y X = dist y x" proof - have "bdd_below (dist y ` X)" by auto from distance_attains_inf[OF assms, of y] obtain x where INF: "x \ X" "\z. z \ X \ dist y x \ dist y z" by auto have "infdist y X = dist y x" by (auto simp: infdist_def assms intro!: antisym cINF_lower[OF _ \x \ X\] cINF_greatest[OF assms(2) INF(2)]) with \x \ X\ show ?thesis .. qed text \Every metric space is a T4 space:\ instance metric_space \ t4_space proof fix S T::"'a set" assume H: "closed S" "closed T" "S \ T = {}" consider "S = {}" | "T = {}" | "S \ {} \ T \ {}" by auto then show "\U V. open U \ open V \ S \ U \ T \ V \ U \ V = {}" proof (cases) case 1 show ?thesis apply (rule exI[of _ "{}"], rule exI[of _ UNIV]) using 1 by auto next case 2 show ?thesis apply (rule exI[of _ UNIV], rule exI[of _ "{}"]) using 2 by auto next case 3 define U where "U = (\x\S. ball x ((infdist x T)/2))" have A: "open U" unfolding U_def by auto have "infdist x T > 0" if "x \ S" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have B: "S \ U" unfolding U_def by auto define V where "V = (\x\T. ball x ((infdist x S)/2))" have C: "open V" unfolding V_def by auto have "infdist x S > 0" if "x \ T" for x using H that 3 by (auto intro!: infdist_pos_not_in_closed) then have D: "T \ V" unfolding V_def by auto have "(ball x ((infdist x T)/2)) \ (ball y ((infdist y S)/2)) = {}" if "x \ S" "y \ T" for x y proof auto fix z assume H: "dist x z * 2 < infdist x T" "dist y z * 2 < infdist y S" have "2 * dist x y \ 2 * dist x z + 2 * dist y z" by metric also have "... < infdist x T + infdist y S" using H by auto finally have "dist x y < infdist x T \ dist x y < infdist y S" by auto then show False using infdist_le[OF \x \ S\, of y] infdist_le[OF \y \ T\, of x] by (auto simp add: dist_commute) qed then have E: "U \ V = {}" unfolding U_def V_def by auto show ?thesis apply (rule exI[of _ U], rule exI[of _ V]) using A B C D E by auto qed qed lemma tendsto_infdist [tendsto_intros]: assumes f: "(f \ l) F" shows "((\x. infdist (f x) A) \ infdist l A) F" proof (rule tendstoI) fix e ::real assume "e > 0" from tendstoD[OF f this] show "eventually (\x. dist (infdist (f x) A) (infdist l A) < e) F" proof (eventually_elim) fix x from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l] have "dist (infdist (f x) A) (infdist l A) \ dist (f x) l" by (simp add: dist_commute dist_real_def) also assume "dist (f x) l < e" finally show "dist (infdist (f x) A) (infdist l A) < e" . qed qed lemma continuous_infdist[continuous_intros]: assumes "continuous F f" shows "continuous F (\x. infdist (f x) A)" using assms unfolding continuous_def by (rule tendsto_infdist) lemma continuous_on_infdist [continuous_intros]: assumes "continuous_on S f" shows "continuous_on S (\x. infdist (f x) A)" using assms unfolding continuous_on by (auto intro: tendsto_infdist) lemma compact_infdist_le: fixes A::"'a::heine_borel set" assumes "A \ {}" assumes "compact A" assumes "e > 0" shows "compact {x. infdist x A \ e}" proof - from continuous_closed_vimage[of "{0..e}" "\x. infdist x A"] continuous_infdist[OF continuous_ident, of _ UNIV A] have "closed {x. infdist x A \ e}" by (auto simp: vimage_def infdist_nonneg) moreover from assms obtain x0 b where b: "\x. x \ A \ dist x0 x \ b" "closed A" by (auto simp: compact_eq_bounded_closed bounded_def) { fix y assume "infdist y A \ e" moreover from infdist_attains_inf[OF \closed A\ \A \ {}\, of y] obtain z where "z \ A" "infdist y A = dist y z" by blast ultimately have "dist x0 y \ b + e" using b by metric } then have "bounded {x. infdist x A \ e}" by (auto simp: bounded_any_center[where a=x0] intro!: exI[where x="b + e"]) ultimately show "compact {x. infdist x A \ e}" by (simp add: compact_eq_bounded_closed) qed subsection \Separation between Points and Sets\ proposition separate_point_closed: fixes s :: "'a::heine_borel set" assumes "closed s" and "a \ s" shows "\d>0. \x\s. d \ dist a x" proof (cases "s = {}") case True then show ?thesis by(auto intro!: exI[where x=1]) next case False from assms obtain x where "x\s" "\y\s. dist a x \ dist a y" using \s \ {}\ by (blast intro: distance_attains_inf [of s a]) with \x\s\ show ?thesis using dist_pos_lt[of a x] and\a \ s\ by blast qed proposition separate_compact_closed: fixes s t :: "'a::heine_borel set" assumes "compact s" and t: "closed t" "s \ t = {}" shows "\d>0. \x\s. \y\t. d \ dist x y" proof cases assume "s \ {} \ t \ {}" then have "s \ {}" "t \ {}" by auto let ?inf = "\x. infdist x t" have "continuous_on s ?inf" by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident) then obtain x where x: "x \ s" "\y\s. ?inf x \ ?inf y" using continuous_attains_inf[OF \compact s\ \s \ {}\] by auto then have "0 < ?inf x" using t \t \ {}\ in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg) moreover have "\x'\s. \y\t. ?inf x \ dist x' y" using x by (auto intro: order_trans infdist_le) ultimately show ?thesis by auto qed (auto intro!: exI[of _ 1]) proposition separate_closed_compact: fixes s t :: "'a::heine_borel set" assumes "closed s" and "compact t" and "s \ t = {}" shows "\d>0. \x\s. \y\t. d \ dist x y" proof - have *: "t \ s = {}" using assms(3) by auto show ?thesis using separate_compact_closed[OF assms(2,1) *] by (force simp: dist_commute) qed proposition compact_in_open_separated: fixes A::"'a::heine_borel set" assumes "A \ {}" assumes "compact A" assumes "open B" assumes "A \ B" obtains e where "e > 0" "{x. infdist x A \ e} \ B" proof atomize_elim have "closed (- B)" "compact A" "- B \ A = {}" using assms by (auto simp: open_Diff compact_eq_bounded_closed) from separate_closed_compact[OF this] obtain d'::real where d': "d'>0" "\x y. x \ B \ y \ A \ d' \ dist x y" by auto define d where "d = d' / 2" hence "d>0" "d < d'" using d' by auto with d' have d: "\x y. x \ B \ y \ A \ d < dist x y" by force show "\e>0. {x. infdist x A \ e} \ B" proof (rule ccontr) assume "\e. 0 < e \ {x. infdist x A \ e} \ B" with \d > 0\ obtain x where x: "infdist x A \ d" "x \ B" by auto from assms have "closed A" "A \ {}" by (auto simp: compact_eq_bounded_closed) from infdist_attains_inf[OF this] obtain y where y: "y \ A" "infdist x A = dist x y" by auto have "dist x y \ d" using x y by simp also have "\ < dist x y" using y d x by auto finally show False by simp qed qed subsection \Uniform Continuity\ lemma uniformly_continuous_onE: assumes "uniformly_continuous_on s f" "0 < e" obtains d where "d>0" "\x x'. \x\s; x'\s; dist x' x < d\ \ dist (f x') (f x) < e" using assms by (auto simp: uniformly_continuous_on_def) lemma uniformly_continuous_on_sequentially: "uniformly_continuous_on s f \ (\x y. (\n. x n \ s) \ (\n. y n \ s) \ (\n. dist (x n) (y n)) \ 0 \ (\n. dist (f(x n)) (f(y n))) \ 0)" (is "?lhs = ?rhs") proof assume ?lhs { fix x y assume x: "\n. x n \ s" and y: "\n. y n \ s" and xy: "((\n. dist (x n) (y n)) \ 0) sequentially" { fix e :: real assume "e > 0" then obtain d where "d > 0" and d: "\x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" using \?lhs\[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto obtain N where N: "\n\N. dist (x n) (y n) < d" using xy[unfolded lim_sequentially dist_norm] and \d>0\ by auto { fix n assume "n\N" then have "dist (f (x n)) (f (y n)) < e" using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y by (simp add: dist_commute) } then have "\N. \n\N. dist (f (x n)) (f (y n)) < e" by auto } then have "((\n. dist (f(x n)) (f(y n))) \ 0) sequentially" unfolding lim_sequentially and dist_real_def by auto } then show ?rhs by auto next assume ?rhs { assume "\ ?lhs" then obtain e where "e > 0" "\d>0. \x\s. \x'\s. dist x' x < d \ \ dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto then obtain fa where fa: "\x. 0 < x \ fst (fa x) \ s \ snd (fa x) \ s \ dist (fst (fa x)) (snd (fa x)) < x \ \ dist (f (fst (fa x))) (f (snd (fa x))) < e" using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def by (auto simp: dist_commute) define x where "x n = fst (fa (inverse (real n + 1)))" for n define y where "y n = snd (fa (inverse (real n + 1)))" for n have xyn: "\n. x n \ s \ y n \ s" and xy0: "\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" unfolding x_def and y_def using fa by auto { fix e :: real assume "e > 0" then obtain N :: nat where "N \ 0" and N: "0 < inverse (real N) \ inverse (real N) < e" unfolding real_arch_inverse[of e] by auto { fix n :: nat assume "n \ N" then have "inverse (real n + 1) < inverse (real N)" using of_nat_0_le_iff and \N\0\ by auto also have "\ < e" using N by auto finally have "inverse (real n + 1) < e" by auto then have "dist (x n) (y n) < e" using xy0[THEN spec[where x=n]] by auto } then have "\N. \n\N. dist (x n) (y n) < e" by auto } then have "\e>0. \N. \n\N. dist (f (x n)) (f (y n)) < e" using \?rhs\[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding lim_sequentially dist_real_def by auto then have False using fxy and \e>0\ by auto } then show ?lhs unfolding uniformly_continuous_on_def by blast qed subsection \Continuity on a Compact Domain Implies Uniform Continuity\ text\From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\ lemma Heine_Borel_lemma: assumes "compact S" and Ssub: "S \ \\" and opn: "\G. G \ \ \ open G" obtains e where "0 < e" "\x. x \ S \ \G \ \. ball x e \ G" proof - have False if neg: "\e. 0 < e \ \x \ S. \G \ \. \ ball x e \ G" proof - have "\x \ S. \G \ \. \ ball x (1 / Suc n) \ G" for n using neg by simp then obtain f where "\n. f n \ S" and fG: "\G n. G \ \ \ \ ball (f n) (1 / Suc n) \ G" by metis then obtain l r where "l \ S" "strict_mono r" and to_l: "(f \ r) \ l" using \compact S\ compact_def that by metis then obtain G where "l \ G" "G \ \" using Ssub by auto then obtain e where "0 < e" and e: "\z. dist z l < e \ z \ G" using opn open_dist by blast obtain N1 where N1: "\n. n \ N1 \ dist (f (r n)) l < e/2" using to_l apply (simp add: lim_sequentially) using \0 < e\ half_gt_zero that by blast obtain N2 where N2: "of_nat N2 > 2/e" using reals_Archimedean2 by blast obtain x where "x \ ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \ G" using fG [OF \G \ \\, of "r (max N1 N2)"] by blast then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" by simp also have "... \ 1 / real (Suc (max N1 N2))" apply (simp add: field_split_simps del: max.bounded_iff) using \strict_mono r\ seq_suble by blast also have "... \ 1 / real (Suc N2)" by (simp add: field_simps) also have "... < e/2" using N2 \0 < e\ by (simp add: field_simps) finally have "dist (f (r (max N1 N2))) x < e/2" . moreover have "dist (f (r (max N1 N2))) l < e/2" using N1 max.cobounded1 by blast ultimately have "dist x l < e" by metric then show ?thesis using e \x \ G\ by blast qed then show ?thesis by (meson that) qed lemma compact_uniformly_equicontinuous: assumes "compact S" and cont: "\x e. \x \ S; 0 < e\ \ \d. 0 < d \ (\f \ \. \x' \ S. dist x' x < d \ dist (f x') (f x) < e)" and "0 < e" obtains d where "0 < d" "\f x x'. \f \ \; x \ S; x' \ S; dist x' x < d\ \ dist (f x') (f x) < e" proof - obtain d where d_pos: "\x e. \x \ S; 0 < e\ \ 0 < d x e" and d_dist : "\x x' e f. \dist x' x < d x e; x \ S; x' \ S; 0 < e; f \ \\ \ dist (f x') (f x) < e" using cont by metis let ?\ = "((\x. ball x (d x (e/2))) ` S)" have Ssub: "S \ \ ?\" by clarsimp (metis d_pos \0 < e\ dist_self half_gt_zero_iff) then obtain k where "0 < k" and k: "\x. x \ S \ \G \ ?\. ball x k \ G" by (rule Heine_Borel_lemma [OF \compact S\]) auto moreover have "dist (f v) (f u) < e" if "f \ \" "u \ S" "v \ S" "dist v u < k" for f u v proof - obtain G where "G \ ?\" "u \ G" "v \ G" using k that by (metis \dist v u < k\ \u \ S\ \0 < k\ centre_in_ball subsetD dist_commute mem_ball) then obtain w where w: "dist w u < d w (e/2)" "dist w v < d w (e/2)" "w \ S" by auto with that d_dist have "dist (f w) (f v) < e/2" by (metis \0 < e\ dist_commute half_gt_zero) moreover have "dist (f w) (f u) < e/2" using that d_dist w by (metis \0 < e\ dist_commute divide_pos_pos zero_less_numeral) ultimately show ?thesis using dist_triangle_half_r by blast qed ultimately show ?thesis using that by blast qed corollary compact_uniformly_continuous: fixes f :: "'a :: metric_space \ 'b :: metric_space" assumes f: "continuous_on S f" and S: "compact S" shows "uniformly_continuous_on S f" using f unfolding continuous_on_iff uniformly_continuous_on_def by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"]) subsection\<^marker>\tag unimportant\\ Theorems relating continuity and uniform continuity to closures\ lemma continuous_on_closure: "continuous_on (closure S) f \ (\x e. x \ closure S \ 0 < e \ (\d. 0 < d \ (\y. y \ S \ dist y x < d \ dist (f y) (f x) < e)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding continuous_on_iff by (metis Un_iff closure_def) next assume R [rule_format]: ?rhs show ?lhs proof fix x and e::real assume "0 < e" and x: "x \ closure S" obtain \::real where "\ > 0" and \: "\y. \y \ S; dist y x < \\ \ dist (f y) (f x) < e/2" using R [of x "e/2"] \0 < e\ x by auto have "dist (f y) (f x) \ e" if y: "y \ closure S" and dyx: "dist y x < \/2" for y proof - obtain \'::real where "\' > 0" and \': "\z. \z \ S; dist z y < \'\ \ dist (f z) (f y) < e/2" using R [of y "e/2"] \0 < e\ y by auto obtain z where "z \ S" and z: "dist z y < min \' \ / 2" using closure_approachable y by (metis \0 < \'\ \0 < \\ divide_pos_pos min_less_iff_conj zero_less_numeral) have "dist (f z) (f y) < e/2" using \' [OF \z \ S\] z \0 < \'\ by metric moreover have "dist (f z) (f x) < e/2" using \[OF \z \ S\] z dyx by metric ultimately show ?thesis by metric qed then show "\d>0. \x'\closure S. dist x' x < d \ dist (f x') (f x) \ e" by (rule_tac x="\/2" in exI) (simp add: \\ > 0\) qed qed lemma continuous_on_closure_sequentially: fixes f :: "'a::metric_space \ 'b :: metric_space" shows "continuous_on (closure S) f \ (\x a. a \ closure S \ (\n. x n \ S) \ x \ a \ (f \ x) \ f a)" (is "?lhs = ?rhs") proof - have "continuous_on (closure S) f \ (\x \ closure S. continuous (at x within S) f)" by (force simp: continuous_on_closure continuous_within_eps_delta) also have "... = ?rhs" by (force simp: continuous_within_sequentially) finally show ?thesis . qed lemma uniformly_continuous_on_closure: fixes f :: "'a::metric_space \ 'b::metric_space" assumes ucont: "uniformly_continuous_on S f" and cont: "continuous_on (closure S) f" shows "uniformly_continuous_on (closure S) f" unfolding uniformly_continuous_on_def proof (intro allI impI) fix e::real assume "0 < e" then obtain d::real where "d>0" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (f x') (f x) < e/3" using ucont [unfolded uniformly_continuous_on_def, rule_format, of "e/3"] by auto show "\d>0. \x\closure S. \x'\closure S. dist x' x < d \ dist (f x') (f x) < e" proof (rule exI [where x="d/3"], clarsimp simp: \d > 0\) fix x y assume x: "x \ closure S" and y: "y \ closure S" and dyx: "dist y x * 3 < d" obtain d1::real where "d1 > 0" and d1: "\w. \w \ closure S; dist w x < d1\ \ dist (f w) (f x) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "x" "e/3"] \0 < e\ x by auto obtain x' where "x' \ S" and x': "dist x' x < min d1 (d / 3)" using closure_approachable [of x S] by (metis \0 < d1\ \0 < d\ divide_pos_pos min_less_iff_conj x zero_less_numeral) obtain d2::real where "d2 > 0" and d2: "\w \ closure S. dist w y < d2 \ dist (f w) (f y) < e/3" using cont [unfolded continuous_on_iff, rule_format, of "y" "e/3"] \0 < e\ y by auto obtain y' where "y' \ S" and y': "dist y' y < min d2 (d / 3)" using closure_approachable [of y S] by (metis \0 < d2\ \0 < d\ divide_pos_pos min_less_iff_conj y zero_less_numeral) have "dist x' x < d/3" using x' by auto then have "dist x' y' < d" using dyx y' by metric then have "dist (f x') (f y') < e/3" by (rule d [OF \y' \ S\ \x' \ S\]) moreover have "dist (f x') (f x) < e/3" using \x' \ S\ closure_subset x' d1 by (simp add: closure_def) moreover have "dist (f y') (f y) < e/3" using \y' \ S\ closure_subset y' d2 by (simp add: closure_def) ultimately show "dist (f y) (f x) < e" by metric qed qed lemma uniformly_continuous_on_extension_at_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" assumes "x \ closure X" obtains l where "(f \ l) (at x within X)" proof - from assms obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF xs] obtain l where l: "(\n. f (xs n)) \ l" by atomize_elim (simp only: convergent_eq_Cauchy) have "(f \ l) (at x within X)" proof (safe intro!: Lim_within_LIMSEQ) fix xs' assume "\n. xs' n \ x \ xs' n \ X" and xs': "xs' \ x" then have "xs' n \ x" "xs' n \ X" for n by auto from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF \xs' \ x\ \xs' _ \ X\] obtain l' where l': "(\n. f (xs' n)) \ l'" by atomize_elim (simp only: convergent_eq_Cauchy) show "(\n. f (xs' n)) \ l" proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' \ e/2" have "e' > 0" using \e > 0\ by (simp add: e'_def) have "\\<^sub>F n in sequentially. dist (f (xs n)) l < e'" by (simp add: \0 < e'\ l tendstoD) moreover from uc[unfolded uniformly_continuous_on_def, rule_format, OF \e' > 0\] obtain d where d: "d > 0" "\x x'. x \ X \ x' \ X \ dist x x' < d \ dist (f x) (f x') < e'" by auto have "\\<^sub>F n in sequentially. dist (xs n) (xs' n) < d" by (auto intro!: \0 < d\ order_tendstoD tendsto_eq_intros xs xs') ultimately show "\\<^sub>F n in sequentially. dist (f (xs' n)) l < e" proof eventually_elim case (elim n) have "dist (f (xs' n)) l \ dist (f (xs n)) (f (xs' n)) + dist (f (xs n)) l" by metric also have "dist (f (xs n)) (f (xs' n)) < e'" by (auto intro!: d xs \xs' _ \ _\ elim) also note \dist (f (xs n)) l < e'\ also have "e' + e' = e" by (simp add: e'_def) finally show ?case by simp qed qed qed thus ?thesis .. qed lemma uniformly_continuous_on_extension_on_closure: fixes f::"'a::metric_space \ 'b::complete_space" assumes uc: "uniformly_continuous_on X f" obtains g where "uniformly_continuous_on (closure X) g" "\x. x \ X \ f x = g x" "\Y h x. X \ Y \ Y \ closure X \ continuous_on Y h \ (\x. x \ X \ f x = h x) \ x \ Y \ h x = g x" proof - from uc have cont_f: "continuous_on X f" by (simp add: uniformly_continuous_imp_continuous) obtain y where y: "(f \ y x) (at x within X)" if "x \ closure X" for x apply atomize_elim apply (rule choice) using uniformly_continuous_on_extension_at_closure[OF assms] by metis let ?g = "\x. if x \ X then f x else y x" have "uniformly_continuous_on (closure X) ?g" unfolding uniformly_continuous_on_def proof safe fix e::real assume "e > 0" define e' where "e' \ e / 3" have "e' > 0" using \e > 0\ by (simp add: e'_def) from uc[unfolded uniformly_continuous_on_def, rule_format, OF \0 < e'\] obtain d where "d > 0" and d: "\x x'. x \ X \ x' \ X \ dist x' x < d \ dist (f x') (f x) < e'" by auto define d' where "d' = d / 3" have "d' > 0" using \d > 0\ by (simp add: d'_def) show "\d>0. \x\closure X. \x'\closure X. dist x' x < d \ dist (?g x') (?g x) < e" proof (safe intro!: exI[where x=d'] \d' > 0\) fix x x' assume x: "x \ closure X" and x': "x' \ closure X" and dist: "dist x' x < d'" then obtain xs xs' where xs: "xs \ x" "\n. xs n \ X" and xs': "xs' \ x'" "\n. xs' n \ X" by (auto simp: closure_sequential) have "\\<^sub>F n in sequentially. dist (xs' n) x' < d'" and "\\<^sub>F n in sequentially. dist (xs n) x < d'" by (auto intro!: \0 < d'\ order_tendstoD tendsto_eq_intros xs xs') moreover have "(\x. f (xs x)) \ y x" if "x \ closure X" "x \ X" "xs \ x" "\n. xs n \ X" for xs x using that not_eventuallyD by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at) then have "(\x. f (xs' x)) \ ?g x'" "(\x. f (xs x)) \ ?g x" using x x' by (auto intro!: continuous_on_tendsto_compose[OF cont_f] simp: xs' xs) then have "\\<^sub>F n in sequentially. dist (f (xs' n)) (?g x') < e'" "\\<^sub>F n in sequentially. dist (f (xs n)) (?g x) < e'" by (auto intro!: \0 < e'\ order_tendstoD tendsto_eq_intros) ultimately have "\\<^sub>F n in sequentially. dist (?g x') (?g x) < e" proof eventually_elim case (elim n) have "dist (?g x') (?g x) \ dist (f (xs' n)) (?g x') + dist (f (xs' n)) (f (xs n)) + dist (f (xs n)) (?g x)" by (metis add.commute add_le_cancel_left dist_commute dist_triangle dist_triangle_le) also from \dist (xs' n) x' < d'\ \dist x' x < d'\ \dist (xs n) x < d'\ have "dist (xs' n) (xs n) < d" unfolding d'_def by metric with \xs _ \ X\ \xs' _ \ X\ have "dist (f (xs' n)) (f (xs n)) < e'" by (rule d) also note \dist (f (xs' n)) (?g x') < e'\ also note \dist (f (xs n)) (?g x) < e'\ finally show ?case by (simp add: e'_def) qed then show "dist (?g x') (?g x) < e" by simp qed qed moreover have "f x = ?g x" if "x \ X" for x using that by simp moreover { fix Y h x assume Y: "x \ Y" "X \ Y" "Y \ closure X" and cont_h: "continuous_on Y h" and extension: "(\x. x \ X \ f x = h x)" { assume "x \ X" have "x \ closure X" using Y by auto then obtain xs where xs: "xs \ x" "\n. xs n \ X" by (auto simp: closure_sequential) from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y have hx: "(\x. f (xs x)) \ h x" by (auto simp: subsetD extension) then have "(\x. f (xs x)) \ y x" using \x \ X\ not_eventuallyD xs(2) by (force intro!: filterlim_compose[OF y[OF \x \ closure X\]] simp: filterlim_at xs) with hx have "h x = y x" by (rule LIMSEQ_unique) } then have "h x = ?g x" using extension by auto } ultimately show ?thesis .. qed lemma bounded_uniformly_continuous_image: fixes f :: "'a :: heine_borel \ 'b :: heine_borel" assumes "uniformly_continuous_on S f" "bounded S" shows "bounded(f ` S)" by (metis (no_types, lifting) assms bounded_closure_image compact_closure compact_continuous_image compact_eq_bounded_closed image_cong uniformly_continuous_imp_continuous uniformly_continuous_on_extension_on_closure) subsection \With Abstract Topology (TODO: move and remove dependency?)\ lemma openin_contains_ball: "openin (top_of_set T) S \ S \ T \ (\x \ S. \e. 0 < e \ ball x e \ T \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: openin_open) apply (metis Int_commute Int_mono inf.cobounded2 open_contains_ball order_refl subsetCE) done next assume ?rhs then show ?lhs apply (simp add: openin_euclidean_subtopology_iff) by (metis (no_types) Int_iff dist_commute inf.absorb_iff2 mem_ball) qed lemma openin_contains_cball: "openin (top_of_set T) S \ S \ T \ (\x \ S. \e. 0 < e \ cball x e \ T \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp add: openin_contains_ball intro: exI [where x="_/2"]) next assume ?rhs then show ?lhs by (force simp add: openin_contains_ball) qed subsection \Closed Nest\ text \Bounded closed nest property (proof does not use Heine-Borel)\ lemma bounded_closed_nest: fixes S :: "nat \ ('a::heine_borel) set" assumes "\n. closed (S n)" and "\n. S n \ {}" and "\m n. m \ n \ S n \ S m" and "bounded (S 0)" obtains a where "\n. a \ S n" proof - from assms(2) obtain x where x: "\n. x n \ S n" using choice[of "\n x. x \ S n"] by auto from assms(4,1) have "seq_compact (S 0)" by (simp add: bounded_closed_imp_seq_compact) then obtain l r where lr: "l \ S 0" "strict_mono r" "(x \ r) \ l" using x and assms(3) unfolding seq_compact_def by blast have "\n. l \ S n" proof fix n :: nat have "closed (S n)" using assms(1) by simp moreover have "\i. (x \ r) i \ S i" using x and assms(3) and lr(2) [THEN seq_suble] by auto then have "\i. (x \ r) (i + n) \ S n" using assms(3) by (fast intro!: le_add2) moreover have "(\i. (x \ r) (i + n)) \ l" using lr(3) by (rule LIMSEQ_ignore_initial_segment) ultimately show "l \ S n" by (rule closed_sequentially) qed then show ?thesis using that by blast qed text \Decreasing case does not even need compactness, just completeness.\ lemma decreasing_closed_nest: fixes S :: "nat \ ('a::complete_space) set" assumes "\n. closed (S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x\S n. \y\S n. dist x y < e" obtains a where "\n. a \ S n" proof - have "\n. \x. x \ S n" using assms(2) by auto then have "\t. \n. t n \ S n" using choice[of "\n x. x \ S n"] by auto then obtain t where t: "\n. t n \ S n" by auto { fix e :: real assume "e > 0" then obtain N where N: "\x\S N. \y\S N. dist x y < e" using assms(4) by blast { fix m n :: nat assume "N \ m \ N \ n" then have "t m \ S N" "t n \ S N" using assms(3) t unfolding subset_eq t by blast+ then have "dist (t m) (t n) < e" using N by auto } then have "\N. \m n. N \ m \ N \ n \ dist (t m) (t n) < e" by auto } then have "Cauchy t" unfolding cauchy_def by auto then obtain l where l:"(t \ l) sequentially" using complete_UNIV unfolding complete_def by auto { fix n :: nat { fix e :: real assume "e > 0" then obtain N :: nat where N: "\n\N. dist (t n) l < e" using l[unfolded lim_sequentially] by auto have "t (max n N) \ S n" by (meson assms(3) contra_subsetD max.cobounded1 t) then have "\y\S n. dist y l < e" using N max.cobounded2 by blast } then have "l \ S n" using closed_approachable[of "S n" l] assms(1) by auto } then show ?thesis using that by blast qed text \Strengthen it to the intersection actually being a singleton.\ lemma decreasing_closed_nest_sing: fixes S :: "nat \ 'a::complete_space set" assumes "\n. closed(S n)" "\n. S n \ {}" "\m n. m \ n \ S n \ S m" "\e. e>0 \ \n. \x \ (S n). \ y\(S n). dist x y < e" shows "\a. \(range S) = {a}" proof - obtain a where a: "\n. a \ S n" using decreasing_closed_nest[of S] using assms by auto { fix b assume b: "b \ \(range S)" { fix e :: real assume "e > 0" then have "dist a b < e" using assms(4) and b and a by blast } then have "dist a b = 0" by (metis dist_eq_0_iff dist_nz less_le) } with a have "\(range S) = {a}" unfolding image_def by auto then show ?thesis .. qed subsection\<^marker>\tag unimportant\ \Making a continuous function avoid some value in a neighbourhood\ lemma continuous_within_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x within s) f" and "f x \ a" shows "\e>0. \y \ s. dist x y < e --> f y \ a" proof - obtain U where "open U" and "f x \ U" and "a \ U" using t1_space [OF \f x \ a\] by fast have "(f \ f x) (at x within s)" using assms(1) by (simp add: continuous_within) then have "eventually (\y. f y \ U) (at x within s)" using \open U\ and \f x \ U\ unfolding tendsto_def by fast then have "eventually (\y. f y \ a) (at x within s)" using \a \ U\ by (fast elim: eventually_mono) then show ?thesis using \f x \ a\ by (auto simp: dist_commute eventually_at) qed lemma continuous_at_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous (at x) f" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms continuous_within_avoid[of x UNIV f a] by simp lemma continuous_on_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "x \ s" and "f x \ a" shows "\e>0. \y \ s. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)] continuous_within_avoid[of x s f a] using assms(3) by auto lemma continuous_on_open_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" assumes "continuous_on s f" and "open s" and "x \ s" and "f x \ a" shows "\e>0. \y. dist x y < e \ f y \ a" using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] using continuous_at_avoid[of x f a] assms(4) by auto subsection \Consequences for Real Numbers\ lemma closed_contains_Inf: fixes S :: "real set" shows "S \ {} \ bdd_below S \ closed S \ Inf S \ S" by (metis closure_contains_Inf closure_closed) lemma closed_subset_contains_Inf: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_below A \ Inf A \ C" by (metis closure_contains_Inf closure_minimal subset_eq) lemma closed_contains_Sup: fixes S :: "real set" shows "S \ {} \ bdd_above S \ closed S \ Sup S \ S" by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup) lemma closed_subset_contains_Sup: fixes A C :: "real set" shows "closed C \ A \ C \ A \ {} \ bdd_above A \ Sup A \ C" by (metis closure_contains_Sup closure_minimal subset_eq) lemma atLeastAtMost_subset_contains_Inf: fixes A :: "real set" and a b :: real shows "A \ {} \ a \ b \ A \ {a..b} \ Inf A \ {a..b}" by (rule closed_subset_contains_Inf) (auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a]) lemma bounded_real: "bounded (S::real set) \ (\a. \x\S. \x\ \ a)" by (simp add: bounded_iff) lemma bounded_imp_bdd_above: "bounded S \ bdd_above (S :: real set)" by (auto simp: bounded_def bdd_above_def dist_real_def) (metis abs_le_D1 abs_minus_commute diff_le_eq) lemma bounded_imp_bdd_below: "bounded S \ bdd_below (S :: real set)" by (auto simp: bounded_def bdd_below_def dist_real_def) (metis abs_le_D1 add.commute diff_le_eq) lemma bounded_has_Sup: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Sup S" and "\b. (\x\S. x \ b) \ Sup S \ b" proof show "\b. (\x\S. x \ b) \ Sup S \ b" using assms by (metis cSup_least) qed (metis cSup_upper assms(1) bounded_imp_bdd_above) lemma Sup_insert: fixes S :: "real set" shows "bounded S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) lemma bounded_has_Inf: fixes S :: "real set" assumes "bounded S" and "S \ {}" shows "\x\S. x \ Inf S" and "\b. (\x\S. x \ b) \ Inf S \ b" proof show "\b. (\x\S. x \ b) \ Inf S \ b" using assms by (metis cInf_greatest) qed (metis cInf_lower assms(1) bounded_imp_bdd_below) lemma Inf_insert: fixes S :: "real set" shows "bounded S \ Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) lemma open_real: fixes s :: "real set" shows "open s \ (\x \ s. \e>0. \x'. \x' - x\ < e --> x' \ s)" unfolding open_dist dist_norm by simp lemma islimpt_approachable_real: fixes s :: "real set" shows "x islimpt s \ (\e>0. \x'\ s. x' \ x \ \x' - x\ < e)" unfolding islimpt_approachable dist_norm by simp lemma closed_real: fixes s :: "real set" shows "closed s \ (\x. (\e>0. \x' \ s. x' \ x \ \x' - x\ < e) \ x \ s)" unfolding closed_limpt islimpt_approachable dist_norm by simp lemma continuous_at_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous (at x) f \ (\e>0. \d>0. \x'. norm(x' - x) < d --> \f x' - f x\ < e)" unfolding continuous_at unfolding Lim_at unfolding dist_norm apply auto apply (erule_tac x=e in allE, auto) apply (rule_tac x=d in exI, auto) apply (erule_tac x=x' in allE, auto) apply (erule_tac x=e in allE, auto) done lemma continuous_on_real_range: fixes f :: "'a::real_normed_vector \ real" shows "continuous_on s f \ (\x \ s. \e>0. \d>0. (\x' \ s. norm(x' - x) < d \ \f x' - f x\ < e))" unfolding continuous_on_iff dist_norm by simp lemma continuous_on_closed_Collect_le: fixes f g :: "'a::topological_space \ real" assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s" shows "closed {x \ s. f x \ g x}" proof - have "closed ((\x. g x - f x) -` {0..} \ s)" using closed_real_atLeast continuous_on_diff [OF g f] by (simp add: continuous_on_closed_vimage [OF s]) also have "((\x. g x - f x) -` {0..} \ s) = {x\s. f x \ g x}" by auto finally show ?thesis . qed lemma continuous_le_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. x \ a}" ] assms continuous_on_closed_Collect_le[of "UNIV" "\x. x" "\x. a"] by auto lemma continuous_ge_on_closure: fixes a::real assumes f: "continuous_on (closure s) f" and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" using image_closure_subset [OF f, where T=" {x. a \ x}"] assms continuous_on_closed_Collect_le[of "UNIV" "\x. a" "\x. x"] by auto subsection\The infimum of the distance between two sets\ definition\<^marker>\tag important\ setdist :: "'a::metric_space set \ 'a set \ real" where "setdist s t \ (if s = {} \ t = {} then 0 else Inf {dist x y| x y. x \ s \ y \ t})" lemma setdist_empty1 [simp]: "setdist {} t = 0" by (simp add: setdist_def) lemma setdist_empty2 [simp]: "setdist t {} = 0" by (simp add: setdist_def) lemma setdist_pos_le [simp]: "0 \ setdist s t" by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) lemma le_setdistI: assumes "s \ {}" "t \ {}" "\x y. \x \ s; y \ t\ \ d \ dist x y" shows "d \ setdist s t" using assms by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) lemma setdist_le_dist: "\x \ s; y \ t\ \ setdist s t \ dist x y" unfolding setdist_def by (auto intro!: bdd_belowI [where m=0] cInf_lower) lemma le_setdist_iff: "d \ setdist S T \ (\x \ S. \y \ T. d \ dist x y) \ (S = {} \ T = {} \ d \ 0)" apply (cases "S = {} \ T = {}") apply (force simp add: setdist_def) apply (intro iffI conjI) using setdist_le_dist apply fastforce apply (auto simp: intro: le_setdistI) done lemma setdist_ltE: assumes "setdist S T < b" "S \ {}" "T \ {}" obtains x y where "x \ S" "y \ T" "dist x y < b" using assms by (auto simp: not_le [symmetric] le_setdist_iff) lemma setdist_refl: "setdist S S = 0" apply (cases "S = {}") apply (force simp add: setdist_def) apply (rule antisym [OF _ setdist_pos_le]) apply (metis all_not_in_conv dist_self setdist_le_dist) done lemma setdist_sym: "setdist S T = setdist T S" by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf]) lemma setdist_triangle: "setdist S T \ setdist S {a} + setdist {a} T" proof (cases "S = {} \ T = {}") case True then show ?thesis using setdist_pos_le by fastforce next case False then have "\x. x \ S \ setdist S T - dist x a \ setdist {a} T" apply (intro le_setdistI) apply (simp_all add: algebra_simps) apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist]) done then have "setdist S T - setdist {a} T \ setdist S {a}" using False by (fastforce intro: le_setdistI) then show ?thesis by (simp add: algebra_simps) qed lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y" by (simp add: setdist_def) lemma setdist_Lipschitz: "\setdist {x} S - setdist {y} S\ \ dist x y" apply (subst setdist_singletons [symmetric]) by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym) lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\y. (setdist {y} S))" by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma continuous_on_setdist [continuous_intros]: "continuous_on T (\y. (setdist {y} S))" by (metis continuous_at_setdist continuous_at_imp_continuous_on) lemma uniformly_continuous_on_setdist: "uniformly_continuous_on T (\y. (setdist {y} S))" by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) lemma setdist_subset_right: "\T \ {}; T \ u\ \ setdist S u \ setdist S T" apply (cases "S = {} \ u = {}", force) apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono) done lemma setdist_subset_left: "\S \ {}; S \ T\ \ setdist T u \ setdist S u" by (metis setdist_subset_right setdist_sym) lemma setdist_closure_1 [simp]: "setdist (closure S) T = setdist S T" proof (cases "S = {} \ T = {}") case True then show ?thesis by force next case False { fix y assume "y \ T" have "continuous_on (closure S) (\a. dist a y)" by (auto simp: continuous_intros dist_norm) then have *: "\x. x \ closure S \ setdist S T \ dist x y" by (fast intro: setdist_le_dist \y \ T\ continuous_ge_on_closure) } note * = this show ?thesis apply (rule antisym) using False closure_subset apply (blast intro: setdist_subset_left) using False * apply (force intro!: le_setdistI) done qed lemma setdist_closure_2 [simp]: "setdist T (closure S) = setdist T S" by (metis setdist_closure_1 setdist_sym) lemma setdist_eq_0I: "\x \ S; x \ T\ \ setdist S T = 0" by (metis antisym dist_self setdist_le_dist setdist_pos_le) lemma setdist_unique: "\a \ S; b \ T; \x y. x \ S \ y \ T ==> dist a b \ dist x y\ \ setdist S T = dist a b" by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) lemma setdist_le_sing: "x \ S ==> setdist S T \ setdist {x} T" using setdist_subset_left by auto lemma infdist_eq_setdist: "infdist x A = setdist {x} A" by (simp add: infdist_def setdist_def Setcompr_eq_image) lemma setdist_eq_infdist: "setdist A B = (if A = {} then 0 else INF a\A. infdist a B)" proof - have "Inf {dist x y |x y. x \ A \ y \ B} = (INF x\A. Inf (dist x ` B))" if "b \ B" "a \ A" for a b proof (rule order_antisym) have "Inf {dist x y |x y. x \ A \ y \ B} \ Inf (dist x ` B)" if "b \ B" "a \ A" "x \ A" for x proof - have *: "\b'. b' \ B \ Inf {dist x y |x y. x \ A \ y \ B} \ dist x b'" by (metis (mono_tags, lifting) ex_in_conv setdist_def setdist_le_dist that(3)) show ?thesis using that by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: *)+ qed then show "Inf {dist x y |x y. x \ A \ y \ B} \ (INF x\A. Inf (dist x ` B))" using that by (subst conditionally_complete_lattice_class.le_cInf_iff) (auto simp: bdd_below_def) next have *: "\x y. \b \ B; a \ A; x \ A; y \ B\ \ \a\A. Inf (dist a ` B) \ dist x y" by (meson bdd_below_image_dist cINF_lower) show "(INF x\A. Inf (dist x ` B)) \ Inf {dist x y |x y. x \ A \ y \ B}" proof (rule conditionally_complete_lattice_class.cInf_mono) show "bdd_below ((\x. Inf (dist x ` B)) ` A)" by (metis (no_types, lifting) bdd_belowI2 ex_in_conv infdist_def infdist_nonneg that(1)) qed (use that in \auto simp: *\) qed then show ?thesis by (auto simp: setdist_def infdist_def) qed lemma infdist_mono: assumes "A \ B" "A \ {}" shows "infdist x B \ infdist x A" by (simp add: assms infdist_eq_setdist setdist_subset_right) lemma infdist_singleton [simp]: "infdist x {y} = dist x y" by (simp add: infdist_eq_setdist) proposition setdist_attains_inf: assumes "compact B" "B \ {}" obtains y where "y \ B" "setdist A B = infdist y A" proof (cases "A = {}") case True then show thesis by (metis assms diameter_compact_attained infdist_def setdist_def that) next case False obtain y where "y \ B" and min: "\y'. y' \ B \ infdist y A \ infdist y' A" by (metis continuous_attains_inf [OF assms continuous_on_infdist] continuous_on_id) show thesis proof have "setdist A B = (INF y\B. infdist y A)" by (metis \B \ {}\ setdist_eq_infdist setdist_sym) also have "\ = infdist y A" proof (rule order_antisym) show "(INF y\B. infdist y A) \ infdist y A" proof (rule cInf_lower) show "infdist y A \ (\y. infdist y A) ` B" using \y \ B\ by blast show "bdd_below ((\y. infdist y A) ` B)" by (meson bdd_belowI2 infdist_nonneg) qed next show "infdist y A \ (INF y\B. infdist y A)" by (simp add: \B \ {}\ cINF_greatest min) qed finally show "setdist A B = infdist y A" . qed (fact \y \ B\) qed end diff --git a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy @@ -1,5020 +1,5020 @@ (* Title: HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Huge cleanup by LCP *) theory Equivalence_Lebesgue_Henstock_Integration imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure Set_Integral Homeomorphism Cartesian_Euclidean_Space begin lemma LIMSEQ_if_less: "(\k. if i < k then a else b) \ a" by (rule_tac k="Suc i" in LIMSEQ_offset) auto text \Note that the rhs is an implication. This lemma plays a specific role in one proof.\ lemma le_left_mono: "x \ y \ y \ a \ x \ (a::'a::preorder)" by (auto intro: order_trans) lemma ball_trans: assumes "y \ ball z q" "r + q \ s" shows "ball y r \ ball z s" using assms by metric lemma has_integral_implies_lebesgue_measurable_cbox: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) (cbox x y)" shows "f \ lebesgue_on (cbox x y) \\<^sub>M borel" proof (rule cld_measure.borel_measurable_cld) let ?L = "lebesgue_on (cbox x y)" let ?\ = "emeasure ?L" let ?\' = "outer_measure_of ?L" interpret L: finite_measure ?L proof show "?\ (space ?L) \ \" by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq) qed show "cld_measure ?L" proof fix B A assume "B \ A" "A \ null_sets ?L" then show "B \ sets ?L" using null_sets_completion_subset[OF \B \ A\, of lborel] by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: ) next fix A assume "A \ space ?L" "\B. B \ sets ?L \ ?\ B < \ \ A \ B \ sets ?L" from this(1) this(2)[of "space ?L"] show "A \ sets ?L" by (auto simp: Int_absorb2 less_top[symmetric]) qed auto then interpret cld_measure ?L . have content_eq_L: "A \ sets borel \ A \ cbox x y \ content A = measure ?L A" for A by (subst measure_restrict_space) (auto simp: measure_def) fix E and a b :: real assume "E \ sets ?L" "a < b" "0 < ?\ E" "?\ E < \" then obtain M :: real where "?\ E = M" "0 < M" by (cases "?\ E") auto define e where "e = M / (4 + 2 / (b - a))" from \a < b\ \0 have "0 < e" by (auto intro!: divide_pos_pos simp: field_simps e_def) have "e < M / (3 + 2 / (b - a))" using \a < b\ \0 < M\ unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps) then have "2 * e < (b - a) * (M - e * 3)" using \0 \0 < e\ \a < b\ by (simp add: field_simps) have e_less_M: "e < M / 1" unfolding e_def using \a < b\ \0 by (intro divide_strict_left_mono) (auto simp: field_simps) obtain d where "gauge d" and integral_f: "\p. p tagged_division_of cbox x y \ d fine p \ norm ((\(x,k) \ p. content k *\<^sub>R f x) - I) < e" using \0 f unfolding has_integral by auto define C where "C X m = X \ {x. ball x (1/Suc m) \ d x}" for X m have "incseq (C X)" for X unfolding C_def [abs_def] by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto { fix X assume "X \ space ?L" and eq: "?\' X = ?\ E" have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\m. C X m)" using \X \ space ?L\ by (intro SUP_outer_measure_of_incseq \incseq (C X)\) (auto simp: C_def) also have "(\m. C X m) = X" proof - { fix x obtain e where "0 < e" "ball x e \ d x" using gaugeD[OF \gauge d\, of x] unfolding open_contains_ball by auto moreover obtain n where "1 / (1 + real n) < e" using reals_Archimedean[OF \0] by (auto simp: inverse_eq_divide) then have "ball x (1 / (1 + real n)) \ ball x e" by (intro subset_ball) auto ultimately have "\n. ball x (1 / (1 + real n)) \ d x" by blast } then show ?thesis by (auto simp: C_def) qed finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\ E" using eq by auto also have "\ > M - e" using \0 < M\ \?\ E = M\ \0 by (auto intro!: ennreal_lessI) finally have "\m. M - e < outer_measure_of ?L (C X m)" unfolding less_SUP_iff by auto } note C = this let ?E = "{x\E. f x \ a}" and ?F = "{x\E. b \ f x}" have "\ (?\' ?E = ?\ E \ ?\' ?F = ?\ E)" proof assume eq: "?\' ?E = ?\ E \ ?\' ?F = ?\ E" with C[of ?E] C[of ?F] \E \ sets ?L\[THEN sets.sets_into_space] obtain ma mb where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)" by auto moreover define m where "m = max ma mb" ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)" using incseqD[OF \incseq (C ?E)\, of ma m, THEN outer_measure_of_mono] incseqD[OF \incseq (C ?F)\, of mb m, THEN outer_measure_of_mono] by (auto intro: less_le_trans) define d' where "d' x = d x \ ball x (1 / (3 * Suc m))" for x have "gauge d'" unfolding d'_def by (intro gauge_Int \gauge d\ gauge_ball) auto then obtain p where p: "p tagged_division_of cbox x y" "d' fine p" by (rule fine_division_exists) then have "d fine p" unfolding d'_def[abs_def] fine_def by auto define s where "s = {(x::'a, k). k \ (C ?E m) \ {} \ k \ (C ?F m) \ {}}" define T where "T E k = (SOME x. x \ k \ C E m)" for E k let ?A = "(\(x, k). (T ?E k, k)) ` (p \ s) \ (p - s)" let ?B = "(\(x, k). (T ?F k, k)) ` (p \ s) \ (p - s)" { fix X assume X_eq: "X = ?E \ X = ?F" let ?T = "(\(x, k). (T X k, k))" let ?p = "?T ` (p \ s) \ (p - s)" have in_s: "(x, k) \ s \ T X k \ k \ C X m" for x k using someI_ex[of "\x. x \ k \ C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def) { fix x k assume "(x, k) \ p" "(x, k) \ s" have k: "k \ ball x (1 / (3 * Suc m))" using \d' fine p\[THEN fineD, OF \(x, k) \ p\] by (auto simp: d'_def) then have "x \ ball (T X k) (1 / (3 * Suc m))" using in_s[OF \(x, k) \ s\] by (auto simp: C_def subset_eq dist_commute) then have "ball x (1 / (3 * Suc m)) \ ball (T X k) (1 / Suc m)" by (rule ball_trans) (auto simp: field_split_simps) with k in_s[OF \(x, k) \ s\] have "k \ d (T X k)" by (auto simp: C_def) } then have "d fine ?p" using \d fine p\ by (auto intro!: fineI) moreover have "?p tagged_division_of cbox x y" proof (rule tagged_division_ofI) show "finite ?p" using p(1) by auto next fix z k assume *: "(z, k) \ ?p" then consider "(z, k) \ p" "(z, k) \ s" | x' where "(x', k) \ p" "(x', k) \ s" "z = T X k" by (auto simp: T_def) then have "z \ k \ k \ cbox x y \ (\a b. k = cbox a b)" using p(1) by cases (auto dest: in_s) then show "z \ k" "k \ cbox x y" "\a b. k = cbox a b" by auto next fix z k z' k' assume "(z, k) \ ?p" "(z', k') \ ?p" "(z, k) \ (z', k')" with tagged_division_ofD(5)[OF p(1), of _ k _ k'] show "interior k \ interior k' = {}" by (auto simp: T_def dest: in_s) next have "{k. \x. (x, k) \ ?p} = {k. \x. (x, k) \ p}" by (auto simp: T_def image_iff Bex_def) then show "\{k. \x. (x, k) \ ?p} = cbox x y" using p(1) by auto qed ultimately have I: "norm ((\(x,k) \ ?p. content k *\<^sub>R f x) - I) < e" using integral_f by auto have "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) + (\(x,k) \ p - s. content k *\<^sub>R f x)" using p(1)[THEN tagged_division_ofD(1)] by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def) also have "(\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k))" proof (subst sum.reindex_nontrivial, safe) fix x1 x2 k assume 1: "(x1, k) \ p" "(x1, k) \ s" and 2: "(x2, k) \ p" "(x2, k) \ s" and eq: "content k *\<^sub>R f (T X k) \ 0" with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k] show "x1 = x2" by (auto simp: content_eq_0_interior) qed (use p in \auto intro!: sum.cong\) finally have eq: "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k)) + (\(x,k) \ p - s. content k *\<^sub>R f x)" . have in_T: "(x, k) \ s \ T X k \ X" for x k using in_s[of x k] by (auto simp: C_def) note I eq in_T } note parts = this have p_in_L: "(x, k) \ p \ k \ sets ?L" for x k using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space) have [simp]: "finite p" using tagged_division_ofD(1)[OF p(1)] . have "(M - 3*e) * (b - a) \ (\(x,k) \ p \ s. content k) * (b - a)" proof (intro mult_right_mono) have fin: "?\ (E \ \{k\snd`p. k \ C X m = {}}) < \" for X using \?\ E < \\ by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \E \ sets ?L\) have sets: "(E \ \{k\snd`p. k \ C X m = {}}) \ sets ?L" for X using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \E \ sets ?L\ sets.finite_Union sets.Int) (auto intro: p_in_L) { fix X assume "X \ E" "M - e < ?\' (C X m)" have "M - e \ ?\' (C X m)" by (rule less_imp_le) fact also have "\ \ ?\' (E - (E \ \{k\snd`p. k \ C X m = {}}))" proof (intro outer_measure_of_mono subsetI) fix v assume "v \ C X m" then have "v \ cbox x y" "v \ E" using \E \ space ?L\ \X \ E\ by (auto simp: space_restrict_space C_def) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto then show "v \ E - E \ (\{k\snd`p. k \ C X m = {}})" using \v \ C X m\ \v \ E\ by auto qed also have "\ = ?\ E - ?\ (E \ \{k\snd`p. k \ C X m = {}})" using \E \ sets ?L\ fin[of X] sets[of X] by (auto intro!: emeasure_Diff) finally have "?\ (E \ \{k\snd`p. k \ C X m = {}}) \ e" using \0 < e\ e_less_M by (cases "?\ (E \ \{k\snd`p. k \ C X m = {}})") (auto simp add: \?\ E = M\ ennreal_minus ennreal_le_iff2) note this } note upper_bound = this have "?\ (E \ \(snd`(p - s))) = ?\ ((E \ \{k\snd`p. k \ C ?E m = {}}) \ (E \ \{k\snd`p. k \ C ?F m = {}}))" by (intro arg_cong[where f="?\"]) (auto simp: s_def image_def Bex_def) also have "\ \ ?\ (E \ \{k\snd`p. k \ C ?E m = {}}) + ?\ (E \ \{k\snd`p. k \ C ?F m = {}})" using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto also have "\ \ e + ennreal e" using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto finally have "?\ E - 2*e \ ?\ (E - (E \ \(snd`(p - s))))" using \0 < e\ \E \ sets ?L\ tagged_division_ofD(1)[OF p(1)] by (subst emeasure_Diff) (auto simp: top_unique simp flip: ennreal_plus intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ \ ?\ (\x\p \ s. snd x)" proof (safe intro!: emeasure_mono subsetI) fix v assume "v \ E" and not: "v \ (\x\p \ s. snd x)" then have "v \ cbox x y" using \E \ space ?L\ by (auto simp: space_restrict_space) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto with not show "v \ \(snd ` (p - s))" by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"]) qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ = measure ?L (\x\p \ s. snd x)" by (auto intro!: emeasure_eq_ennreal_measure) finally have "M - 2 * e \ measure ?L (\x\p \ s. snd x)" unfolding \?\ E = M\ using \0 < e\ by (simp add: ennreal_minus) also have "measure ?L (\x\p \ s. snd x) = content (\x\p \ s. snd x)" using tagged_division_ofD(1,3,4) [OF p(1)] by (intro content_eq_L[symmetric]) (fastforce intro!: sets.finite_UN UN_least del: subsetI)+ also have "content (\x\p \ s. snd x) \ (\k\p \ s. content (snd k))" using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite dest!: p(1)[THEN tagged_division_ofD(4)]) finally show "M - 3 * e \ (\(x, y)\p \ s. content y)" using \0 < e\ by (simp add: split_beta) qed (use \a < b\ in auto) also have "\ = (\(x,k) \ p \ s. content k * (b - a))" by (simp add: sum_distrib_right split_beta') also have "\ \ (\(x,k) \ p \ s. content k * (f (T ?F k) - f (T ?E k)))" using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono) also have "\ = (\(x,k) \ p \ s. content k * f (T ?F k)) - (\(x,k) \ p \ s. content k * f (T ?E k))" by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric]) also have "\ = (\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x)" by (subst (1 2) parts) auto also have "\ \ norm ((\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x))" by auto also have "\ \ e + e" using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto finally show False using \2 * e < (b - a) * (M - e * 3)\ by (auto simp: field_simps) qed moreover have "?\' ?E \ ?\ E" "?\' ?F \ ?\ E" unfolding outer_measure_of_eq[OF \E \ sets ?L\, symmetric] by (auto intro!: outer_measure_of_mono) ultimately show "min (?\' ?E) (?\' ?F) < ?\ E" unfolding min_less_iff_disj by (auto simp: less_le) qed lemma has_integral_implies_lebesgue_measurable_real: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) \" shows "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof - define B :: "nat \ 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n show "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof (rule measurable_piecewise_restrict) have "(\n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \ \(B ` UNIV)" unfolding B_def by (intro UN_mono box_subset_cbox order_refl) then show "countable (range B)" "space lebesgue \ \(B ` UNIV)" by (auto simp: B_def UN_box_eq_UNIV) next fix \' assume "\' \ range B" then obtain n where \': "\' = B n" by auto then show "\' \ space lebesgue \ sets lebesgue" by (auto simp: B_def) have "f integrable_on \" using f by auto then have "(\x. f x * indicator \ x) integrable_on \" by (auto simp: integrable_on_def cong: has_integral_cong) then have "(\x. f x * indicator \ x) integrable_on (\ \ B n)" by (rule integrable_on_superset) auto then have "(\x. f x * indicator \ x) integrable_on B n" unfolding B_def by (rule integrable_on_subcbox) auto then show "(\x. f x * indicator \ x) \ lebesgue_on \' \\<^sub>M borel" unfolding B_def \' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def) qed qed lemma has_integral_implies_lebesgue_measurable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "(f has_integral I) \" shows "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI) fix i :: "'b" assume "i \ Basis" have "(\x. (f x \ i) * indicator \ x) \ borel_measurable (completion lborel)" using has_integral_linear[OF f bounded_linear_inner_left, of i] by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def) then show "(\x. indicator \ x *\<^sub>R f x \ i) \ borel_measurable (completion lborel)" by (simp add: ac_simps) qed subsection \Equivalence Lebesgue integral on \<^const>\lborel\ and HK-integral\ lemma has_integral_measure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and finite: "emeasure lborel A < \" shows "((\x. 1) has_integral measure lborel A) A" proof - { fix l u :: 'a have "((\x. 1) has_integral measure lborel (box l u)) (box l u)" proof cases assume "\b\Basis. l \ b \ u \ b" then show ?thesis using has_integral_const[of "1::real" l u] by (simp flip: has_integral_restrict[OF box_subset_cbox] add: has_integral_spike_interior) next assume "\ (\b\Basis. l \ b \ u \ b)" then have "box l u = {}" unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le) then show ?thesis by simp qed } note has_integral_box = this { fix a b :: 'a let ?M = "\A. measure lborel (A \ box a b)" have "Int_stable (range (\(a, b). box a b))" by (auto simp: Int_stable_def box_Int_box) moreover have "(range (\(a, b). box a b)) \ Pow UNIV" by auto moreover have "A \ sigma_sets UNIV (range (\(a, b). box a b))" using A unfolding borel_eq_box by simp ultimately have "((\x. 1) has_integral ?M A) (A \ box a b)" proof (induction rule: sigma_sets_induct_disjoint) case (basic A) then show ?case by (auto simp: box_Int_box has_integral_box) next case empty then show ?case by simp next case (compl A) then have [measurable]: "A \ sets borel" by (simp add: borel_eq_box) have "((\x. 1) has_integral ?M (box a b)) (box a b)" by (simp add: has_integral_box) moreover have "((\x. if x \ A \ box a b then 1 else 0) has_integral ?M A) (box a b)" by (subst has_integral_restrict) (auto intro: compl) ultimately have "((\x. 1 - (if x \ A \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_diff) then have "((\x. (if x \ (UNIV - A) \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_cong[THEN iffD1, rotated 1]) auto then have "((\x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \ box a b)" by (subst (asm) has_integral_restrict) auto also have "?M (box a b) - ?M A = ?M (UNIV - A)" by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2) finally show ?case . next case (union F) then have [measurable]: "\i. F i \ sets borel" by (simp add: borel_eq_box subset_eq) have "((\x. if x \ \(F ` UNIV) \ box a b then 1 else 0) has_integral ?M (\i. F i)) (box a b)" proof (rule has_integral_monotone_convergence_increasing) let ?f = "\k x. \i F i \ box a b then 1 else 0 :: real" show "\k. (?f k has_integral (\ik x. ?f k x \ ?f (Suc k) x" by (intro sum_mono2) auto from union(1) have *: "\x i j. x \ F i \ x \ F j \ j = i" by (auto simp add: disjoint_family_on_def) show "(\k. ?f k x) \ (if x \ \(F ` UNIV) \ box a b then 1 else 0)" for x by (auto simp: * sum.If_cases Iio_Int_singleton if_distrib LIMSEQ_if_less cong: if_cong) have *: "emeasure lborel ((\x. F x) \ box a b) \ emeasure lborel (box a b)" by (intro emeasure_mono) auto with union(1) show "(\k. \i ?M (\i. F i)" unfolding sums_def[symmetric] UN_extend_simps by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique) qed then show ?case by (subst (asm) has_integral_restrict) auto qed } note * = this show ?thesis proof (rule has_integral_monotone_convergence_increasing) let ?B = "\n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set" let ?f = "\n::nat. \x. if x \ A \ ?B n then 1 else 0 :: real" let ?M = "\n. measure lborel (A \ ?B n)" show "\n::nat. (?f n has_integral ?M n) A" using * by (subst has_integral_restrict) simp_all show "\k x. ?f k x \ ?f (Suc k) x" by (auto simp: box_def) { fix x assume "x \ A" moreover have "(\k. indicator (A \ ?B k) x :: real) \ indicator (\k::nat. A \ ?B k) x" by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def) ultimately show "(\k. if x \ A \ ?B k then 1 else 0::real) \ 1" by (simp add: indicator_def UN_box_eq_UNIV) } have "(\n. emeasure lborel (A \ ?B n)) \ emeasure lborel (\n::nat. A \ ?B n)" by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def) also have "(\n. emeasure lborel (A \ ?B n)) = (\n. measure lborel (A \ ?B n))" proof (intro ext emeasure_eq_ennreal_measure) fix n have "emeasure lborel (A \ ?B n) \ emeasure lborel (?B n)" by (intro emeasure_mono) auto then show "emeasure lborel (A \ ?B n) \ top" by (auto simp: top_unique) qed finally show "(\n. measure lborel (A \ ?B n)) \ measure lborel A" using emeasure_eq_ennreal_measure[of lborel A] finite by (simp add: UN_box_eq_UNIV less_top) qed qed lemma nn_integral_has_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" shows "(f has_integral r) UNIV" using f proof (induct f arbitrary: r rule: borel_measurable_induct_real) case (set A) then have "((\x. 1) has_integral measure lborel A) A" by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator) with set show ?case by (simp add: ennreal_indicator measure_def) (simp add: indicator_def) next case (mult g c) then have "ennreal c * (\\<^sup>+ x. g x \lborel) = ennreal r" by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult) with \0 \ r\ \0 \ c\ obtain r' where "(c = 0 \ r = 0) \ (0 \ r' \ (\\<^sup>+ x. ennreal (g x) \lborel) = ennreal r' \ r = c * r')" by (cases "\\<^sup>+ x. ennreal (g x) \lborel" rule: ennreal_cases) (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric]) with mult show ?case by (auto intro!: has_integral_cmult_real) next case (add g h) then have "(\\<^sup>+ x. h x + g x \lborel) = (\\<^sup>+ x. h x \lborel) + (\\<^sup>+ x. g x \lborel)" by (simp add: nn_integral_add) with add obtain a b where "0 \ a" "0 \ b" "(\\<^sup>+ x. h x \lborel) = ennreal a" "(\\<^sup>+ x. g x \lborel) = ennreal b" "r = a + b" by (cases "\\<^sup>+ x. h x \lborel" "\\<^sup>+ x. g x \lborel" rule: ennreal2_cases) (auto simp: add_top nn_integral_add top_add simp flip: ennreal_plus) with add show ?case by (auto intro!: has_integral_add) next case (seq U) note seq(1)[measurable] and f[measurable] have U_le_f: "U i x \ f x" for i x by (metis (no_types) LIMSEQ_le_const UNIV_I incseq_def le_fun_def seq.hyps(4) seq.hyps(5) space_borel) { fix i have "(\\<^sup>+x. U i x \lborel) \ (\\<^sup>+x. f x \lborel)" using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp then obtain p where "(\\<^sup>+x. U i x \lborel) = ennreal p" "p \ r" "0 \ p" using seq(6) \0\r\ by (cases "\\<^sup>+x. U i x \lborel" rule: ennreal_cases) (auto simp: top_unique) moreover note seq ultimately have "\p. (\\<^sup>+x. U i x \lborel) = ennreal p \ 0 \ p \ p \ r \ (U i has_integral p) UNIV" by auto } then obtain p where p: "\i. (\\<^sup>+x. ennreal (U i x) \lborel) = ennreal (p i)" and bnd: "\i. p i \ r" "\i. 0 \ p i" and U_int: "\i.(U i has_integral (p i)) UNIV" by metis have int_eq: "\i. integral UNIV (U i) = p i" using U_int by (rule integral_unique) have *: "f integrable_on UNIV \ (\k. integral UNIV (U k)) \ integral UNIV f" proof (rule monotone_convergence_increasing) show "\k. U k integrable_on UNIV" using U_int by auto show "\k x. x\UNIV \ U k x \ U (Suc k) x" using \incseq U\ by (auto simp: incseq_def le_fun_def) then show "bounded (range (\k. integral UNIV (U k)))" using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) show "\x. x\UNIV \ (\k. U k x) \ f x" using seq by auto qed moreover have "(\i. (\\<^sup>+x. U i x \lborel)) \ (\\<^sup>+x. f x \lborel)" using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto ultimately have "integral UNIV f = r" by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique) with * show ?case by (simp add: has_integral_integral) qed lemma nn_integral_lborel_eq_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "(\\<^sup>+x. f x \lborel) = integral UNIV f" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique) qed lemma nn_integral_integrable_on: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "f integrable_on UNIV" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f) qed lemma nn_integral_has_integral_lborel: fixes f :: "'a::euclidean_space \ real" assumes f_borel: "f \ borel_measurable borel" and nonneg: "\x. 0 \ f x" assumes I: "(f has_integral I) UNIV" shows "integral\<^sup>N lborel f = I" proof - from f_borel have "(\x. ennreal (f x)) \ borel_measurable lborel" by auto from borel_measurable_implies_simple_function_sequence'[OF this] obtain F where F: "\i. simple_function lborel (F i)" "incseq F" "\i x. F i x < top" "\x. (SUP i. F i x) = ennreal (f x)" by blast then have [measurable]: "\i. F i \ borel_measurable lborel" by (metis borel_measurable_simple_function) let ?B = "\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set" have "0 \ I" using I by (rule has_integral_nonneg) (simp add: nonneg) have F_le_f: "enn2real (F i x) \ f x" for i x using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\i. F i x"] by (cases "F i x" rule: ennreal_cases) auto let ?F = "\i x. F i x * indicator (?B i) x" have "(\\<^sup>+ x. ennreal (f x) \lborel) = (SUP i. integral\<^sup>N lborel (\x. ?F i x))" proof (subst nn_integral_monotone_convergence_SUP[symmetric]) { fix x obtain j where j: "x \ ?B j" using UN_box_eq_UNIV by auto have "ennreal (f x) = (SUP i. F i x)" using F(4)[of x] nonneg[of x] by (simp add: max_def) also have "\ = (SUP i. ?F i x)" proof (rule SUP_eq) fix i show "\j\UNIV. F i x \ ?F j x" using j F(2) by (intro bexI[of _ "max i j"]) (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def) qed (auto intro!: F split: split_indicator) finally have "ennreal (f x) = (SUP i. ?F i x)" . } then show "(\\<^sup>+ x. ennreal (f x) \lborel) = (\\<^sup>+ x. (SUP i. ?F i x) \lborel)" by simp qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator) also have "\ \ ennreal I" proof (rule SUP_least) fix i :: nat have finite_F: "(\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel) < \" proof (rule nn_integral_bound_simple_function) have "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} \ emeasure lborel (?B i)" by (intro emeasure_mono) (auto split: split_indicator) then show "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} < \" by (auto simp: less_top[symmetric] top_unique) qed (auto split: split_indicator intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal) have int_F: "(\x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV" using F(4) finite_F by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg) have "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) = (\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel)" using F(3,4) by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator) also have "\ = ennreal (integral UNIV (\x. enn2real (F i x) * indicator (?B i) x))" using F by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F]) (auto split: split_indicator intro: enn2real_nonneg) also have "\ \ ennreal I" by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f simp: \0 \ I\ split: split_indicator ) finally show "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) \ ennreal I" . qed finally have "(\\<^sup>+ x. ennreal (f x) \lborel) < \" by (auto simp: less_top[symmetric] top_unique) from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis by (simp add: integral_unique) qed lemma has_integral_iff_emeasure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and [simp]: "0 \ r" shows "((\x. 1) has_integral r) A \ emeasure lborel A = ennreal r" proof (cases "emeasure lborel A = \") case emeasure_A: True have "\ (\x. 1::real) integrable_on A" proof assume int: "(\x. 1::real) integrable_on A" then have "(indicator A::'a \ real) integrable_on UNIV" unfolding indicator_def[abs_def] integrable_restrict_UNIV . then obtain r where "((indicator A::'a\real) has_integral r) UNIV" by auto from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False by (simp add: ennreal_indicator) qed with emeasure_A show ?thesis by auto next case False then have "((\x. 1) has_integral measure lborel A) A" by (simp add: has_integral_measure_lborel less_top) with False show ?thesis by (auto simp: emeasure_eq_ennreal_measure has_integral_unique) qed lemma ennreal_max_0: "ennreal (max 0 x) = ennreal x" by (auto simp: max_def ennreal_neg) lemma has_integral_integral_real: fixes f::"'a::euclidean_space \ real" assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - from integrableE[OF f] obtain r q where "0 \ r" "0 \ q" and r: "(\\<^sup>+ x. ennreal (max 0 (f x)) \lborel) = ennreal r" and q: "(\\<^sup>+ x. ennreal (max 0 (- f x)) \lborel) = ennreal q" and f: "f \ borel_measurable lborel" and eq: "integral\<^sup>L lborel f = r - q" unfolding ennreal_max_0 by auto then have "((\x. max 0 (f x)) has_integral r) UNIV" "((\x. max 0 (- f x)) has_integral q) UNIV" using nn_integral_has_integral[OF _ _ r] nn_integral_has_integral[OF _ _ q] by auto note has_integral_diff[OF this] moreover have "(\x. max 0 (f x) - max 0 (- f x)) = f" by auto ultimately show ?thesis by (simp add: eq) qed lemma has_integral_AE: assumes ae: "AE x in lborel. x \ \ \ f x = g x" shows "(f has_integral x) \ = (g has_integral x) \" proof - from ae obtain N where N: "N \ sets borel" "emeasure lborel N = 0" "{x. \ (x \ \ \ f x = g x)} \ N" by (auto elim!: AE_E) then have not_N: "AE x in lborel. x \ N" by (simp add: AE_iff_measurable) show ?thesis proof (rule has_integral_spike_eq[symmetric]) show "\x. x\\ - N \ f x = g x" using N(3) by auto show "negligible N" unfolding negligible_def proof (intro allI) fix a b :: "'a" let ?F = "\x::'a. if x \ cbox a b then indicator N x else 0 :: real" have "integrable lborel ?F = integrable lborel (\x::'a. 0::real)" using not_N N(1) by (intro integrable_cong_AE) auto moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)" using not_N N(1) by (intro integral_cong_AE) auto ultimately have "(?F has_integral 0) UNIV" using has_integral_integral_real[of ?F] by simp then show "(indicator N has_integral (0::real)) (cbox a b)" unfolding has_integral_restrict_UNIV . qed qed qed lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto finally show ?thesis using nonneg by auto qed lemma has_integral_iff_nn_integral_lebesgue: assumes f: "\x. 0 \ f x" shows "(f has_integral r) UNIV \ (f \ lebesgue \\<^sub>M borel \ integral\<^sup>N lebesgue f = r \ 0 \ r)" (is "?I = ?N") proof assume ?I have "0 \ r" using has_integral_nonneg[OF \?I\] f by auto then show ?N using nn_integral_has_integral_lebesgue[OF f \?I\] has_integral_implies_lebesgue_measurable[OF \?I\] by (auto simp: nn_integral_completion) next assume ?N then obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" by (auto dest: completion_ex_borel_measurable_real) moreover have "(\\<^sup>+ x. ennreal \f' x\ \lborel) = (\\<^sup>+ x. ennreal \f x\ \lborel)" using f' by (intro nn_integral_cong_AE) auto moreover have "((\x. \f' x\) has_integral r) UNIV \ ((\x. \f x\) has_integral r) UNIV" using f' by (intro has_integral_AE) auto moreover note nn_integral_has_integral[of "\x. \f' x\" r] \?N\ ultimately show ?I using f by (auto simp: nn_integral_completion) qed context fixes f::"'a::euclidean_space \ 'b::euclidean_space" begin lemma has_integral_integral_lborel: assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lborel f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma integrable_on_lborel: "integrable lborel f \ f integrable_on UNIV" using has_integral_integral_lborel by auto lemma integral_lborel: "integrable lborel f \ integral UNIV f = (\x. f x \lborel)" using has_integral_integral_lborel by auto end context begin private lemma has_integral_integral_lebesgue_real: fixes f :: "'a::euclidean_space \ real" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" using completion_ex_borel_measurable_real[OF borel_measurable_integrable[OF f]] by auto moreover have "(\\<^sup>+ x. ennreal (norm (f x)) \lborel) = (\\<^sup>+ x. ennreal (norm (f' x)) \lborel)" using f' by (intro nn_integral_cong_AE) auto ultimately have "integrable lborel f'" using f by (auto simp: integrable_iff_bounded nn_integral_completion cong: nn_integral_cong_AE) note has_integral_integral_real[OF this] moreover have "integral\<^sup>L lebesgue f = integral\<^sup>L lebesgue f'" using f' f by (intro integral_cong_AE) (auto intro: AE_completion measurable_completion) moreover have "integral\<^sup>L lebesgue f' = integral\<^sup>L lborel f'" using f' by (simp add: integral_completion) moreover have "(f' has_integral integral\<^sup>L lborel f') UNIV \ (f has_integral integral\<^sup>L lborel f') UNIV" using f' by (intro has_integral_AE) auto ultimately show ?thesis by auto qed lemma has_integral_integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lebesgue f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma has_integral_integral_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "(f has_integral (integral\<^sup>L (lebesgue_on S) f)) S" proof - let ?f = "\x. if x \ S then f x else 0" have "integrable lebesgue (\x. indicat_real S x *\<^sub>R f x)" using indicator_scaleR_eq_if [of S _ f] assms by (metis (full_types) integrable_restrict_space sets.Int_space_eq2) then have "integrable lebesgue ?f" using indicator_scaleR_eq_if [of S _ f] assms by auto then have "(?f has_integral (integral\<^sup>L lebesgue ?f)) UNIV" by (rule has_integral_integral_lebesgue) then have "(f has_integral (integral\<^sup>L lebesgue ?f)) S" using has_integral_restrict_UNIV by blast moreover have "S \ space lebesgue \ sets lebesgue" by (simp add: assms) then have "(integral\<^sup>L lebesgue ?f) = (integral\<^sup>L (lebesgue_on S) f)" by (simp add: integral_restrict_space indicator_scaleR_eq_if) ultimately show ?thesis by auto qed lemma lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on S) f = integral S f" by (metis has_integral_integral_lebesgue_on assms integral_unique) lemma integrable_on_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ f integrable_on UNIV" using has_integral_integral_lebesgue by auto lemma integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ integral UNIV f = (\x. f x \lebesgue)" using has_integral_integral_lebesgue by auto end subsection \Absolute integrability (this is the same as Lebesgue integrability)\ translations "LBINT x. f" == "CONST lebesgue_integral CONST lborel (\x. f)" translations "LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\x. f)" lemma set_integral_reflect: fixes S and f :: "real \ 'a :: {banach, second_countable_topology}" shows "(LBINT x : S. f x) = (LBINT x : {x. - x \ S}. f (- x))" unfolding set_lebesgue_integral_def by (subst lborel_integral_real_affine[where c="-1" and t=0]) (auto intro!: Bochner_Integration.integral_cong split: split_indicator) lemma borel_integrable_atLeastAtMost': fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes f: "continuous_on {a..b} f" shows "set_integrable lborel {a..b} f" unfolding set_integrable_def by (intro borel_integrable_compact compact_Icc f) lemma integral_FTC_atLeastAtMost: fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" and F: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" and f: "continuous_on {a .. b} f" shows "integral\<^sup>L lborel (\x. indicator {a .. b} x *\<^sub>R f x) = F b - F a" proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have "(?f has_integral (\x. ?f x \lborel)) UNIV" using borel_integrable_atLeastAtMost'[OF f] unfolding set_integrable_def by (rule has_integral_integral_lborel) moreover have "(f has_integral F b - F a) {a .. b}" by (intro fundamental_theorem_of_calculus ballI assms) auto then have "(?f has_integral F b - F a) {a .. b}" by (subst has_integral_cong[where g=f]) auto then have "(?f has_integral F b - F a) UNIV" by (intro has_integral_on_superset[where T=UNIV and S="{a..b}"]) auto ultimately show "integral\<^sup>L lborel ?f = F b - F a" by (rule has_integral_unique) qed lemma set_borel_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "set_integrable lborel S f" shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f" proof - let ?f = "\x. indicator S x *\<^sub>R f x" have "(?f has_integral LINT x : S | lborel. f x) UNIV" using assms has_integral_integral_lborel unfolding set_integrable_def set_lebesgue_integral_def by blast hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S" by (simp add: indicator_scaleR_eq_if) thus "f integrable_on S" by (auto simp add: integrable_on_def) with 1 have "(f has_integral (integral S f)) S" by (intro integrable_integral, auto simp add: integrable_on_def) thus "LINT x : S | lborel. f x = integral S f" by (intro has_integral_unique [OF 1]) qed lemma has_integral_set_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "(f has_integral (LINT x:S|lebesgue. f x)) S" using has_integral_integral_lebesgue f by (fastforce simp add: set_integrable_def set_lebesgue_integral_def indicator_def if_distrib[of "\x. x *\<^sub>R f _"] cong: if_cong) lemma set_lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f" using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def) lemma lmeasurable_iff_has_integral: "S \ lmeasurable \ ((indicator S) has_integral measure lebesgue S) UNIV" by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI) abbreviation absolutely_integrable_on :: "('a::euclidean_space \ 'b::{banach, second_countable_topology}) \ 'a set \ bool" (infixr "absolutely'_integrable'_on" 46) where "f absolutely_integrable_on s \ set_integrable lebesgue s f" lemma absolutely_integrable_zero [simp]: "(\x. 0) absolutely_integrable_on S" by (simp add: set_integrable_def) lemma absolutely_integrable_on_def: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ f integrable_on S \ (\x. norm (f x)) integrable_on S" proof safe assume f: "f absolutely_integrable_on S" then have nf: "integrable lebesgue (\x. norm (indicator S x *\<^sub>R f x))" using integrable_norm set_integrable_def by blast show "f integrable_on S" by (rule set_lebesgue_integral_eq_integral[OF f]) have "(\x. norm (indicator S x *\<^sub>R f x)) = (\x. if x \ S then norm (f x) else 0)" by auto with integrable_on_lebesgue[OF nf] show "(\x. norm (f x)) integrable_on S" by (simp add: integrable_restrict_UNIV) next assume f: "f integrable_on S" and nf: "(\x. norm (f x)) integrable_on S" show "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule integrableI_bounded) show "(\x. indicator S x *\<^sub>R f x) \ borel_measurable lebesgue" using f has_integral_implies_lebesgue_measurable[of f _ S] by (auto simp: integrable_on_def) show "(\\<^sup>+ x. ennreal (norm (indicator S x *\<^sub>R f x)) \lebesgue) < \" using nf nn_integral_has_integral_lebesgue[of "\x. norm (f x)" _ S] by (auto simp: integrable_on_def nn_integral_completion) qed qed lemma integrable_on_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable (lebesgue_on S) f" and S: "S \ sets lebesgue" shows "f integrable_on S" proof - have "integrable lebesgue (\x. indicator S x *\<^sub>R f x)" using S f inf_top.comm_neutral integrable_restrict_space by blast then show ?thesis using absolutely_integrable_on_def set_integrable_def by blast qed lemma absolutely_integrable_imp_integrable: assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" by (meson assms integrable_restrict_space set_integrable_def sets.Int sets.top) lemma absolutely_integrable_on_null [intro]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "content (cbox a b) = 0 \ f absolutely_integrable_on (cbox a b)" by (auto simp: absolutely_integrable_on_def) lemma absolutely_integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f absolutely_integrable_on box a b \ f absolutely_integrable_on cbox a b" by (auto simp: integrable_on_open_interval absolutely_integrable_on_def) lemma absolutely_integrable_restrict_UNIV: "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ f absolutely_integrable_on S" unfolding set_integrable_def by (intro arg_cong2[where f=integrable]) auto lemma absolutely_integrable_onI: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f integrable_on S \ (\x. norm (f x)) integrable_on S \ f absolutely_integrable_on S" unfolding absolutely_integrable_on_def by auto lemma nonnegative_absolutely_integrable_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "\x. x \ A \ 0 \ f x" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_onI [OF f]) (use assms in \simp add: integrable_eq\) lemma absolutely_integrable_on_iff_nonneg: fixes f :: "'a :: euclidean_space \ real" assumes "\x. x \ S \ 0 \ f x" shows "f absolutely_integrable_on S \ f integrable_on S" proof - { assume "f integrable_on S" then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" using \f integrable_on S\ absolutely_integrable_restrict_UNIV assms nonnegative_absolutely_integrable_1 by blast then have "f absolutely_integrable_on S" using absolutely_integrable_restrict_UNIV by blast } then show ?thesis unfolding absolutely_integrable_on_def by auto qed lemma absolutely_integrable_on_scaleR_iff: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S \ c = 0 \ f absolutely_integrable_on S" proof (cases "c=0") case False then show ?thesis unfolding absolutely_integrable_on_def by (simp add: norm_mult) qed auto lemma absolutely_integrable_spike: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on T" and S: "negligible S" "\x. x \ T - S \ g x = f x" shows "g absolutely_integrable_on T" using assms integrable_spike unfolding absolutely_integrable_on_def by metis lemma absolutely_integrable_negligible: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" shows "f absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def integrable_negligible) lemma absolutely_integrable_spike_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" "\x. x \ T - S \ g x = f x" shows "(f absolutely_integrable_on T \ g absolutely_integrable_on T)" using assms by (blast intro: absolutely_integrable_spike sym) lemma absolutely_integrable_spike_set_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f absolutely_integrable_on S \ f absolutely_integrable_on T)" proof - have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ (\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" proof (rule absolutely_integrable_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto with absolutely_integrable_restrict_UNIV show ?thesis by blast qed lemma absolutely_integrable_spike_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f absolutely_integrable_on T" using absolutely_integrable_spike_set_eq f neg by blast lemma absolutely_integrable_reflect[simp]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on cbox (-b) (-a) \ f absolutely_integrable_on cbox a b" unfolding absolutely_integrable_on_def by (metis (mono_tags, lifting) integrable_eq integrable_reflect) lemma absolutely_integrable_reflect_real[simp]: fixes f :: "real \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on {-b .. -a} \ f absolutely_integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule absolutely_integrable_reflect) lemma absolutely_integrable_on_subcbox: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; cbox a b \ S\ \ f absolutely_integrable_on cbox a b" by (meson absolutely_integrable_on_def integrable_on_subcbox) lemma absolutely_integrable_on_subinterval: fixes f :: "real \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; {a..b} \ S\ \ f absolutely_integrable_on {a..b}" using absolutely_integrable_on_subcbox by fastforce lemma integrable_subinterval: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" and "{c..d} \ {a..b}" shows "integrable (lebesgue_on {c..d}) f" proof (rule absolutely_integrable_imp_integrable) show "f absolutely_integrable_on {c..d}" proof - have "f integrable_on {c..d}" using assms integrable_on_lebesgue_on integrable_on_subinterval by fastforce moreover have "(\x. norm (f x)) integrable_on {c..d}" proof (rule integrable_on_subinterval) show "(\x. norm (f x)) integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) qed (use assms in auto) ultimately show ?thesis by (auto simp: absolutely_integrable_on_def) qed qed auto lemma indefinite_integral_continuous_real: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" shows "continuous_on {a..b} (\x. integral\<^sup>L (lebesgue_on {a..x}) f)" proof - have "f integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) then have "continuous_on {a..b} (\x. integral {a..x} f)" using indefinite_integral_continuous_1 by blast moreover have "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" if "a \ x" "x \ b" for x proof - have "{a..x} \ {a..b}" using that by auto then have "integrable (lebesgue_on {a..x}) f" using integrable_subinterval assms by blast then show "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" by (simp add: lebesgue_integral_eq_integral) qed ultimately show ?thesis by (metis (no_types, lifting) atLeastAtMost_iff continuous_on_cong) qed lemma lmeasurable_iff_integrable_on: "S \ lmeasurable \ (\x. 1::real) integrable_on S" by (subst absolutely_integrable_on_iff_nonneg[symmetric]) (simp_all add: lmeasurable_iff_integrable set_integrable_def) lemma lmeasure_integral_UNIV: "S \ lmeasurable \ measure lebesgue S = integral UNIV (indicator S)" by (simp add: lmeasurable_iff_has_integral integral_unique) lemma lmeasure_integral: "S \ lmeasurable \ measure lebesgue S = integral S (\x. 1::real)" by (fastforce simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on) lemma integrable_on_const: "S \ lmeasurable \ (\x. c) integrable_on S" unfolding lmeasurable_iff_integrable by (metis (mono_tags, lifting) integrable_eq integrable_on_scaleR_left lmeasurable_iff_integrable lmeasurable_iff_integrable_on scaleR_one) lemma integral_indicator: assumes "(S \ T) \ lmeasurable" shows "integral T (indicator S) = measure lebesgue (S \ T)" proof - have "integral UNIV (indicator (S \ T)) = integral UNIV (\a. if a \ S \ T then 1::real else 0)" by (meson indicator_def) moreover have "(indicator (S \ T) has_integral measure lebesgue (S \ T)) UNIV" using assms by (simp add: lmeasurable_iff_has_integral) ultimately have "integral UNIV (\x. if x \ S \ T then 1 else 0) = measure lebesgue (S \ T)" by (metis (no_types) integral_unique) moreover have "integral T (\a. if a \ S then 1::real else 0) = integral (S \ T \ UNIV) (\a. 1)" by (simp add: Henstock_Kurzweil_Integration.integral_restrict_Int) moreover have "integral T (indicat_real S) = integral T (\a. if a \ S then 1 else 0)" by (meson indicator_def) ultimately show ?thesis by (simp add: assms lmeasure_integral) qed lemma measurable_integrable: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ (indicat_real S) integrable_on UNIV" by (auto simp: lmeasurable_iff_integrable absolutely_integrable_on_iff_nonneg [symmetric] set_integrable_def) lemma integrable_on_indicator: fixes S :: "'a::euclidean_space set" shows "indicat_real S integrable_on T \ (S \ T) \ lmeasurable" unfolding measurable_integrable unfolding integrable_restrict_UNIV [of T, symmetric] by (fastforce simp add: indicator_def elim: integrable_eq) lemma assumes \: "\ division_of S" shows lmeasurable_division: "S \ lmeasurable" (is ?l) and content_division: "(\k\\. measure lebesgue k) = measure lebesgue S" (is ?m) proof - { fix d1 d2 assume *: "d1 \ \" "d2 \ \" "d1 \ d2" then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d" using division_ofD(4)[OF \] by blast with division_ofD(5)[OF \ *] have "d1 \ sets lborel" "d2 \ sets lborel" "d1 \ d2 \ (cbox a b - box a b) \ (cbox c d - box c d)" by auto moreover have "(cbox a b - box a b) \ (cbox c d - box c d) \ null_sets lborel" by (intro null_sets.Un null_sets_cbox_Diff_box) ultimately have "d1 \ d2 \ null_sets lborel" by (blast intro: null_sets_subset) } then show ?l ?m unfolding division_ofD(6)[OF \, symmetric] using division_ofD(1,4)[OF \] by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def) qed lemma has_measure_limit: assumes "S \ lmeasurable" "e > 0" obtains B where "B > 0" "\a b. ball 0 B \ cbox a b \ \measure lebesgue (S \ cbox a b) - measure lebesgue S\ < e" using assms unfolding lmeasurable_iff_has_integral has_integral_alt' by (force simp: integral_indicator integrable_on_indicator) lemma lmeasurable_iff_indicator_has_integral: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ m = measure lebesgue S \ (indicat_real S has_integral m) UNIV" by (metis has_integral_iff lmeasurable_iff_has_integral measurable_integrable) lemma has_measure_limit_iff: fixes f :: "'n::euclidean_space \ 'a::banach" shows "S \ lmeasurable \ m = measure lebesgue S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (S \ cbox a b) \ lmeasurable \ \measure lebesgue (S \ cbox a b) - m\ < e)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson has_measure_limit fmeasurable.Int lmeasurable_cbox) next assume RHS [rule_format]: ?rhs then show ?lhs apply (simp add: lmeasurable_iff_indicator_has_integral has_integral' [where i=m]) by (metis (full_types) integral_indicator integrable_integral integrable_on_indicator) qed subsection\Applications to Negligibility\ lemma negligible_iff_null_sets: "negligible S \ S \ null_sets lebesgue" proof assume "negligible S" then have "(indicator S has_integral (0::real)) UNIV" by (auto simp: negligible) then show "S \ null_sets lebesgue" by (subst (asm) has_integral_iff_nn_integral_lebesgue) (auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff) next assume S: "S \ null_sets lebesgue" show "negligible S" unfolding negligible_def proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2] has_integral_restrict_UNIV[where s="cbox _ _", THEN iffD1]) fix a b show "(\x. if x \ cbox a b then indicator S x else 0) \ lebesgue \\<^sub>M borel" using S by (auto intro!: measurable_If) then show "(\\<^sup>+ x. ennreal (if x \ cbox a b then indicator S x else 0) \lebesgue) = ennreal 0" using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) qed auto qed corollary eventually_ae_filter_negligible: "eventually P (ae_filter lebesgue) \ (\N. negligible N \ {x. \ P x} \ N)" by (auto simp: completion.AE_iff_null_sets negligible_iff_null_sets null_sets_completion_subset) lemma starlike_negligible: assumes "closed S" and eq1: "\c x. (a + c *\<^sub>R x) \ S \ 0 \ c \ a + x \ S \ c = 1" shows "negligible S" proof - have "negligible ((+) (-a) ` S)" proof (subst negligible_on_intervals, intro allI) fix u v show "negligible ((+) (- a) ` S \ cbox u v)" using \closed S\ eq1 by (auto simp add: negligible_iff_null_sets algebra_simps intro!: closed_translation_subtract starlike_negligible_compact cong: image_cong_simp) (metis add_diff_eq diff_add_cancel scale_right_diff_distrib) qed then show ?thesis by (rule negligible_translation_rev) qed lemma starlike_negligible_strong: assumes "closed S" and star: "\c x. \0 \ c; c < 1; a+x \ S\ \ a + c *\<^sub>R x \ S" shows "negligible S" proof - show ?thesis proof (rule starlike_negligible [OF \closed S\, of a]) fix c x assume cx: "a + c *\<^sub>R x \ S" "0 \ c" "a + x \ S" with star have "\ (c < 1)" by auto moreover have "\ (c > 1)" using star [of "1/c" "c *\<^sub>R x"] cx by force ultimately show "c = 1" by arith qed qed lemma negligible_hyperplane: assumes "a \ 0 \ b \ 0" shows "negligible {x. a \ x = b}" proof - obtain x where x: "a \ x \ b" using assms by (metis euclidean_all_zero_iff inner_zero_right) moreover have "c = 1" if "a \ (x + c *\<^sub>R w) = b" "a \ (x + w) = b" for c w using that by (metis (no_types, hide_lams) add_diff_eq diff_0 diff_minus_eq_add inner_diff_right inner_scaleR_right mult_cancel_right2 right_minus_eq x) ultimately show ?thesis using starlike_negligible [OF closed_hyperplane, of x] by simp qed lemma negligible_lowdim: fixes S :: "'N :: euclidean_space set" assumes "dim S < DIM('N)" shows "negligible S" proof - obtain a where "a \ 0" and a: "span S \ {x. a \ x = 0}" using lowdim_subset_hyperplane [OF assms] by blast have "negligible (span S)" using \a \ 0\ a negligible_hyperplane by (blast intro: negligible_subset) then show ?thesis using span_base by (blast intro: negligible_subset) qed proposition negligible_convex_frontier: fixes S :: "'N :: euclidean_space set" assumes "convex S" shows "negligible(frontier S)" proof - have nf: "negligible(frontier S)" if "convex S" "0 \ S" for S :: "'N set" proof - obtain B where "B \ S" and indB: "independent B" and spanB: "S \ span B" and cardB: "card B = dim S" by (metis basis_exists) consider "dim S < DIM('N)" | "dim S = DIM('N)" using dim_subset_UNIV le_eq_less_or_eq by auto then show ?thesis proof cases case 1 show ?thesis by (rule negligible_subset [of "closure S"]) (simp_all add: frontier_def negligible_lowdim 1) next case 2 obtain a where "a \ interior (convex hull insert 0 B)" proof (rule interior_simplex_nonempty [OF indB]) show "finite B" by (simp add: indB independent_finite) show "card B = DIM('N)" by (simp add: cardB 2) qed then have a: "a \ interior S" by (metis \B \ S\ \0 \ S\ \convex S\ insert_absorb insert_subset interior_mono subset_hull) show ?thesis proof (rule starlike_negligible_strong [where a=a]) fix c::real and x have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)" by (simp add: algebra_simps) assume "0 \ c" "c < 1" "a + x \ frontier S" then show "a + c *\<^sub>R x \ frontier S" using eq mem_interior_closure_convex_shrink [OF \convex S\ a, of _ "1-c"] unfolding frontier_def by (metis Diff_iff add_diff_cancel_left' add_diff_eq diff_gt_0_iff_gt group_cancel.rule0 not_le) qed auto qed qed show ?thesis proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain a where "a \ S" by auto show ?thesis using nf [of "(\x. -a + x) ` S"] by (metis \a \ S\ add.left_inverse assms convex_translation_eq frontier_translation image_eqI negligible_translation_rev) qed qed corollary negligible_sphere: "negligible (sphere a e)" using frontier_cball negligible_convex_frontier convex_cball by (blast intro: negligible_subset) lemma non_negligible_UNIV [simp]: "\ negligible UNIV" unfolding negligible_iff_null_sets by (auto simp: null_sets_def) lemma negligible_interval: "negligible (cbox a b) \ box a b = {}" "negligible (box a b) \ box a b = {}" by (auto simp: negligible_iff_null_sets null_sets_def prod_nonneg inner_diff_left box_eq_empty not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq intro: eq_refl antisym less_imp_le) proposition open_not_negligible: assumes "open S" "S \ {}" shows "\ negligible S" proof assume neg: "negligible S" obtain a where "a \ S" using \S \ {}\ by blast then obtain e where "e > 0" "cball a e \ S" using \open S\ open_contains_cball_eq by blast let ?p = "a - (e / DIM('a)) *\<^sub>R One" let ?q = "a + (e / DIM('a)) *\<^sub>R One" have "cbox ?p ?q \ cball a e" proof (clarsimp simp: mem_box dist_norm) fix x assume "\i\Basis. ?p \ i \ x \ i \ x \ i \ ?q \ i" then have ax: "\(a - x) \ i\ \ e / real DIM('a)" if "i \ Basis" for i using that by (auto simp: algebra_simps) have "norm (a - x) \ (\i\Basis. \(a - x) \ i\)" by (rule norm_le_l1) also have "\ \ DIM('a) * (e / real DIM('a))" by (intro sum_bounded_above ax) also have "\ = e" by simp finally show "norm (a - x) \ e" . qed then have "negligible (cbox ?p ?q)" by (meson \cball a e \ S\ neg negligible_subset) with \e > 0\ show False by (simp add: negligible_interval box_eq_empty algebra_simps field_split_simps mult_le_0_iff) qed lemma negligible_convex_interior: "convex S \ (negligible S \ interior S = {})" by (metis Diff_empty closure_subset frontier_def interior_subset negligible_convex_frontier negligible_subset open_interior open_not_negligible) lemma measure_eq_0_null_sets: "S \ null_sets M \ measure M S = 0" by (auto simp: measure_def null_sets_def) lemma negligible_imp_measure0: "negligible S \ measure lebesgue S = 0" by (simp add: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_emeasure0: "S \ sets lebesgue \ (negligible S \ emeasure lebesgue S = 0)" by (auto simp: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_measure0: "S \ lmeasurable \ (negligible S \ measure lebesgue S = 0)" by (metis (full_types) completion.null_sets_outer negligible_iff_null_sets negligible_imp_measure0 order_refl) lemma negligible_imp_sets: "negligible S \ S \ sets lebesgue" by (simp add: negligible_iff_null_sets null_setsD2) lemma negligible_imp_measurable: "negligible S \ S \ lmeasurable" by (simp add: fmeasurableI_null_sets negligible_iff_null_sets) lemma negligible_iff_measure: "negligible S \ S \ lmeasurable \ measure lebesgue S = 0" by (fastforce simp: negligible_iff_measure0 negligible_imp_measurable dest: negligible_imp_measure0) lemma negligible_outer: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T < e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis negligible_iff_measure order_refl) next assume ?rhs then show "negligible S" by (meson completion.null_sets_outer negligible_iff_null_sets) qed lemma negligible_outer_le: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T \ e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis dual_order.strict_implies_order negligible_imp_measurable negligible_imp_measure0 order_refl) next assume ?rhs then show "negligible S" by (metis le_less_trans negligible_outer field_lbound_gt_zero) qed lemma negligible_UNIV: "negligible S \ (indicat_real S has_integral 0) UNIV" (is "_=?rhs") by (metis lmeasurable_iff_indicator_has_integral negligible_iff_measure) lemma sets_negligible_symdiff: "\S \ sets lebesgue; negligible((S - T) \ (T - S))\ \ T \ sets lebesgue" by (metis Diff_Diff_Int Int_Diff_Un inf_commute negligible_Un_eq negligible_imp_sets sets.Diff sets.Un) lemma lmeasurable_negligible_symdiff: "\S \ lmeasurable; negligible((S - T) \ (T - S))\ \ T \ lmeasurable" using integrable_spike_set_eq lmeasurable_iff_integrable_on by blast lemma measure_Un3_negligible: assumes meas: "S \ lmeasurable" "T \ lmeasurable" "U \ lmeasurable" and neg: "negligible(S \ T)" "negligible(S \ U)" "negligible(T \ U)" and V: "S \ T \ U = V" shows "measure lebesgue V = measure lebesgue S + measure lebesgue T + measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ T) = 0" using neg(1) negligible_imp_measure0 by blast have [simp]: "measure lebesgue (S \ U \ T \ U) = 0" using neg(2) neg(3) negligible_Un negligible_imp_measure0 by blast have "measure lebesgue V = measure lebesgue (S \ T \ U)" using V by simp also have "\ = measure lebesgue S + measure lebesgue T + measure lebesgue U" by (simp add: measure_Un3 meas fmeasurable.Un Int_Un_distrib2) finally show ?thesis . qed lemma measure_translate_add: assumes meas: "S \ lmeasurable" "T \ lmeasurable" and U: "S \ ((+)a ` T) = U" and neg: "negligible(S \ ((+)a ` T))" shows "measure lebesgue S + measure lebesgue T = measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ (+) a ` T) = 0" using neg negligible_imp_measure0 by blast have "measure lebesgue (S \ ((+)a ` T)) = measure lebesgue S + measure lebesgue T" by (simp add: measure_Un3 meas measurable_translation measure_translation fmeasurable.Un) then show ?thesis using U by auto qed lemma measure_negligible_symdiff: assumes S: "S \ lmeasurable" and neg: "negligible (S - T \ (T - S))" shows "measure lebesgue T = measure lebesgue S" proof - have "measure lebesgue (S - T) = 0" using neg negligible_Un_eq negligible_imp_measure0 by blast then show ?thesis by (metis S Un_commute add.right_neutral lmeasurable_negligible_symdiff measure_Un2 neg negligible_Un_eq negligible_imp_measure0) qed lemma measure_closure: assumes "bounded S" and neg: "negligible (frontier S)" shows "measure lebesgue (closure S) = measure lebesgue S" proof - have "measure lebesgue (frontier S) = 0" by (metis neg negligible_imp_measure0) then show ?thesis by (metis assms lmeasurable_iff_integrable_on eq_iff_diff_eq_0 has_integral_interior integrable_on_def integral_unique lmeasurable_interior lmeasure_integral measure_frontier) qed lemma measure_interior: "\bounded S; negligible(frontier S)\ \ measure lebesgue (interior S) = measure lebesgue S" using measure_closure measure_frontier negligible_imp_measure0 by fastforce lemma measurable_Jordan: assumes "bounded S" and neg: "negligible (frontier S)" shows "S \ lmeasurable" proof - have "closure S \ lmeasurable" by (metis lmeasurable_closure \bounded S\) moreover have "interior S \ lmeasurable" by (simp add: lmeasurable_interior \bounded S\) moreover have "interior S \ S" by (simp add: interior_subset) ultimately show ?thesis using assms by (metis (full_types) closure_subset completion.complete_sets_sandwich_fmeasurable measure_closure measure_interior) qed lemma measurable_convex: "\convex S; bounded S\ \ S \ lmeasurable" by (simp add: measurable_Jordan negligible_convex_frontier) lemma content_cball_conv_ball: "content (cball c r) = content (ball c r)" proof - have "ball c r - cball c r \ (cball c r - ball c r) = sphere c r" by auto hence "measure lebesgue (cball c r) = measure lebesgue (ball c r)" using negligible_sphere[of c r] by (intro measure_negligible_symdiff) simp_all thus ?thesis by simp qed subsection\Negligibility of image under non-injective linear map\ lemma negligible_Union_nat: assumes "\n::nat. negligible(S n)" shows "negligible(\n. S n)" proof - have "negligible (\m\k. S m)" for k using assms by blast then have 0: "integral UNIV (indicat_real (\m\k. S m)) = 0" and 1: "(indicat_real (\m\k. S m)) integrable_on UNIV" for k by (auto simp: negligible has_integral_iff) have 2: "\k x. indicat_real (\m\k. S m) x \ (indicat_real (\m\Suc k. S m) x)" by (simp add: indicator_def) have 3: "\x. (\k. indicat_real (\m\k. S m) x) \ (indicat_real (\n. S n) x)" by (force simp: indicator_def eventually_sequentially intro: tendsto_eventually) have 4: "bounded (range (\k. integral UNIV (indicat_real (\m\k. S m))))" by (simp add: 0) have *: "indicat_real (\n. S n) integrable_on UNIV \ (\k. integral UNIV (indicat_real (\m\k. S m))) \ (integral UNIV (indicat_real (\n. S n)))" by (intro monotone_convergence_increasing 1 2 3 4) then have "integral UNIV (indicat_real (\n. S n)) = (0::real)" using LIMSEQ_unique by (auto simp: 0) then show ?thesis using * by (simp add: negligible_UNIV has_integral_iff) qed lemma negligible_linear_singular_image: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\ inj f" shows "negligible (f ` S)" proof - obtain a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms linear_singular_image_hyperplane by blast then show "negligible (f ` S)" by (metis negligible_hyperplane negligible_subset) qed lemma measure_negligible_finite_Union: assumes "finite \" and meas: "\S. S \ \ \ S \ lmeasurable" and djointish: "pairwise (\S T. negligible (S \ T)) \" shows "measure lebesgue (\\) = (\S\\. measure lebesgue S)" using assms proof (induction) case empty then show ?case by auto next case (insert S \) then have "S \ lmeasurable" "\\ \ lmeasurable" "pairwise (\S T. negligible (S \ T)) \" by (simp_all add: fmeasurable.finite_Union insert.hyps(1) insert.prems(1) pairwise_insert subsetI) then show ?case proof (simp add: measure_Un3 insert) have *: "\T. T \ (\) S ` \ \ negligible T" using insert by (force simp: pairwise_def) have "negligible(S \ \\)" unfolding Int_Union by (rule negligible_Union) (simp_all add: * insert.hyps(1)) then show "measure lebesgue (S \ \\) = 0" using negligible_imp_measure0 by blast qed qed lemma measure_negligible_finite_Union_image: assumes "finite S" and meas: "\x. x \ S \ f x \ lmeasurable" and djointish: "pairwise (\x y. negligible (f x \ f y)) S" shows "measure lebesgue (\(f ` S)) = (\x\S. measure lebesgue (f x))" proof - have "measure lebesgue (\(f ` S)) = sum (measure lebesgue) (f ` S)" using assms by (auto simp: pairwise_mono pairwise_image intro: measure_negligible_finite_Union) also have "\ = sum (measure lebesgue \ f) S" using djointish [unfolded pairwise_def] by (metis inf.idem negligible_imp_measure0 sum.reindex_nontrivial [OF \finite S\]) also have "\ = (\x\S. measure lebesgue (f x))" by simp finally show ?thesis . qed subsection \Negligibility of a Lipschitz image of a negligible set\ text\The bound will be eliminated by a sort of onion argument\ lemma locally_Lipschitz_negl_bounded: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "0 < B" "bounded S" "negligible S" and lips: "\x. x \ S \ \T. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" unfolding negligible_iff_null_sets proof (clarsimp simp: completion.null_sets_outer) fix e::real assume "0 < e" have "S \ lmeasurable" using \negligible S\ by (simp add: negligible_iff_null_sets fmeasurableI_null_sets) then have "S \ sets lebesgue" by blast have e22: "0 < e/2 / (2 * B * real DIM('M)) ^ DIM('N)" using \0 < e\ \0 < B\ by (simp add: field_split_simps) obtain T where "open T" "S \ T" "(T - S) \ lmeasurable" "measure lebesgue (T - S) < e/2 / (2 * B * DIM('M)) ^ DIM('N)" using sets_lebesgue_outer_open [OF \S \ sets lebesgue\ e22] by (metis emeasure_eq_measure2 ennreal_leI linorder_not_le) then have T: "measure lebesgue T \ e/2 / (2 * B * DIM('M)) ^ DIM('N)" using \negligible S\ by (simp add: measure_Diff_null_set negligible_iff_null_sets) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. norm(y - x) < r \ y \ T \ (y \ S \ norm(f y - f x) \ B * norm(y - x))))" for x proof (cases "x \ S") case True obtain U where "open U" "x \ U" and U: "\y. y \ S \ U \ norm(f y - f x) \ B * norm(y - x)" using lips [OF \x \ S\] by auto have "x \ T \ U" using \S \ T\ \x \ U\ \x \ S\ by auto then obtain \ where "0 < \" "ball x \ \ T \ U" by (metis \open T\ \open U\ openE open_Int) then show ?thesis by (rule_tac x="min (1/2) \" in exI) (simp add: U dist_norm norm_minus_commute subset_iff) next case False then show ?thesis by (rule_tac x="1/4" in exI) auto qed then obtain R where R12: "\x. 0 < R x \ R x \ 1/2" and RT: "\x y. \x \ S; norm(y - x) < R x\ \ y \ T" and RB: "\x y. \x \ S; y \ S; norm(y - x) < R x\ \ norm(f y - f x) \ B * norm(y - x)" by metis+ then have gaugeR: "gauge (\x. ball x (R x))" by (simp add: gauge_def) obtain c where c: "S \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" "box (-c *\<^sub>R One:: 'M) (c *\<^sub>R One) \ {}" proof - obtain B where B: "\x. x \ S \ norm x \ B" using \bounded S\ bounded_iff by blast show ?thesis proof (rule_tac c = "abs B + 1" in that) show "S \ cbox (- (\B\ + 1) *\<^sub>R One) ((\B\ + 1) *\<^sub>R One)" using norm_bound_Basis_le Basis_le_norm by (fastforce simp: mem_box dest!: B intro: order_trans) show "box (- (\B\ + 1) *\<^sub>R One) ((\B\ + 1) *\<^sub>R One) \ {}" by (simp add: box_eq_empty(1)) qed qed obtain \ where "countable \" and Dsub: "\\ \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and pw: "pairwise (\A B. interior A \ interior B = {}) \" and Ksub: "\K. K \ \ \ \x \ S \ K. K \ (\x. ball x (R x)) x" and exN: "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (2*c) / 2^n" and "S \ \\" using covering_lemma [OF c gaugeR] by force have "\u v z. K = cbox u v \ box u v \ {} \ z \ S \ z \ cbox u v \ cbox u v \ ball z (R z)" if "K \ \" for K proof - obtain u v where "K = cbox u v" using \K \ \\ cbox by blast with that show ?thesis by (metis Int_iff interior_cbox cbox Ksub) qed then obtain uf vf zf where uvz: "\K. K \ \ \ K = cbox (uf K) (vf K) \ box (uf K) (vf K) \ {} \ zf K \ S \ zf K \ cbox (uf K) (vf K) \ cbox (uf K) (vf K) \ ball (zf K) (R (zf K))" by metis define prj1 where "prj1 \ \x::'M. x \ (SOME i. i \ Basis)" define fbx where "fbx \ \D. cbox (f(zf D) - (B * DIM('M) * (prj1(vf D - uf D))) *\<^sub>R One::'N) (f(zf D) + (B * DIM('M) * prj1(vf D - uf D)) *\<^sub>R One)" have vu_pos: "0 < prj1 (vf X - uf X)" if "X \ \" for X using uvz [OF that] by (simp add: prj1_def box_ne_empty SOME_Basis inner_diff_left) have prj1_idem: "prj1 (vf X - uf X) = (vf X - uf X) \ i" if "X \ \" "i \ Basis" for X i proof - have "cbox (uf X) (vf X) \ \" using uvz \X \ \\ by auto with exN obtain n where "\i. i \ Basis \ vf X \ i - uf X \ i = (2*c) / 2^n" by blast then show ?thesis by (simp add: \i \ Basis\ SOME_Basis inner_diff prj1_def) qed have countbl: "countable (fbx ` \)" using \countable \\ by blast have "(\k\fbx`\'. measure lebesgue k) \ e/2" if "\' \ \" "finite \'" for \' proof - have BM_ge0: "0 \ B * (DIM('M) * prj1 (vf X - uf X))" if "X \ \'" for X using \0 < B\ \\' \ \\ that vu_pos by fastforce have "{} \ \'" using cbox \\' \ \\ interior_empty by blast have "(\k\fbx`\'. measure lebesgue k) \ sum (measure lebesgue o fbx) \'" by (rule sum_image_le [OF \finite \'\]) (force simp: fbx_def) also have "\ \ (\X\\'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)" proof (rule sum_mono) fix X assume "X \ \'" then have "X \ \" using \\' \ \\ by blast then have ufvf: "cbox (uf X) (vf X) = X" using uvz by blast have "prj1 (vf X - uf X) ^ DIM('M) = (\i::'M \ Basis. prj1 (vf X - uf X))" by (rule prod_constant [symmetric]) also have "\ = (\i\Basis. vf X \ i - uf X \ i)" by (simp add: \X \ \\ inner_diff_left prj1_idem cong: prod.cong) finally have prj1_eq: "prj1 (vf X - uf X) ^ DIM('M) = (\i\Basis. vf X \ i - uf X \ i)" . have "uf X \ cbox (uf X) (vf X)" "vf X \ cbox (uf X) (vf X)" using uvz [OF \X \ \\] by (force simp: mem_box)+ moreover have "cbox (uf X) (vf X) \ ball (zf X) (1/2)" by (meson R12 order_trans subset_ball uvz [OF \X \ \\]) ultimately have "uf X \ ball (zf X) (1/2)" "vf X \ ball (zf X) (1/2)" by auto then have "dist (vf X) (uf X) \ 1" unfolding mem_ball by (metis dist_commute dist_triangle_half_l dual_order.order_iff_strict) then have 1: "prj1 (vf X - uf X) \ 1" unfolding prj1_def dist_norm using Basis_le_norm SOME_Basis order_trans by fastforce have 0: "0 \ prj1 (vf X - uf X)" using \X \ \\ prj1_def vu_pos by fastforce have "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))" apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \X \ \'\ \0 < B\ flip: prj1_eq) using MleN 0 1 uvz \X \ \\ by (fastforce simp add: box_ne_empty power_decreasing) also have "\ = (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" by (subst (3) ufvf[symmetric]) simp finally show "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" . qed also have "\ = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \'" by (simp add: sum_distrib_left) also have "\ \ e/2" proof - have "\K. K \ \' \ \a b. K = cbox a b" using cbox that by blast then have div: "\' division_of \\'" using pairwise_subset [OF pw \\' \ \\] unfolding pairwise_def by (force simp: \finite \'\ \{} \ \'\ division_of_def) have le_meaT: "measure lebesgue (\\') \ measure lebesgue T" proof (rule measure_mono_fmeasurable) show "(\\') \ sets lebesgue" using div lmeasurable_division by auto have "\\' \ \\" using \\' \ \\ by blast also have "... \ T" proof (clarify) fix x D assume "x \ D" "D \ \" show "x \ T" using Ksub [OF \D \ \\] by (metis \x \ D\ Int_iff dist_norm mem_ball norm_minus_commute subsetD RT) qed finally show "\\' \ T" . show "T \ lmeasurable" using \S \ lmeasurable\ \S \ T\ \T - S \ lmeasurable\ fmeasurable_Diff_D by blast qed have "sum (measure lebesgue) \' = sum content \'" using \\' \ \\ cbox by (force intro: sum.cong) then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \' = (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\\')" using content_division [OF div] by auto also have "\ \ (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T" using \0 < B\ by (intro mult_left_mono [OF le_meaT]) (force simp add: algebra_simps) also have "\ \ e/2" using T \0 < B\ by (simp add: field_simps) finally show ?thesis . qed finally show ?thesis . qed then have e2: "sum (measure lebesgue) \ \ e/2" if "\ \ fbx ` \" "finite \" for \ by (metis finite_subset_image that) show "\W\lmeasurable. f ` S \ W \ measure lebesgue W < e" proof (intro bexI conjI) have "\X\\. f y \ fbx X" if "y \ S" for y proof - obtain X where "y \ X" "X \ \" using \S \ \\\ \y \ S\ by auto then have y: "y \ ball(zf X) (R(zf X))" using uvz by fastforce have conj_le_eq: "z - b \ y \ y \ z + b \ abs(y - z) \ b" for z y b::real by auto have yin: "y \ cbox (uf X) (vf X)" and zin: "(zf X) \ cbox (uf X) (vf X)" using uvz \X \ \\ \y \ X\ by auto have "norm (y - zf X) \ (\i\Basis. \(y - zf X) \ i\)" by (rule norm_le_l1) also have "\ \ real DIM('M) * prj1 (vf X - uf X)" proof (rule sum_bounded_above) fix j::'M assume j: "j \ Basis" show "\(y - zf X) \ j\ \ prj1 (vf X - uf X)" using yin zin j by (fastforce simp add: mem_box prj1_idem [OF \X \ \\ j] inner_diff_left) qed finally have nole: "norm (y - zf X) \ DIM('M) * prj1 (vf X - uf X)" by simp have fle: "\f y \ i - f(zf X) \ i\ \ B * DIM('M) * prj1 (vf X - uf X)" if "i \ Basis" for i proof - have "\f y \ i - f (zf X) \ i\ = \(f y - f (zf X)) \ i\" by (simp add: algebra_simps) also have "\ \ norm (f y - f (zf X))" by (simp add: Basis_le_norm that) also have "\ \ B * norm(y - zf X)" by (metis uvz RB \X \ \\ dist_commute dist_norm mem_ball \y \ S\ y) also have "\ \ B * real DIM('M) * prj1 (vf X - uf X)" using \0 < B\ by (simp add: nole) finally show ?thesis . qed show ?thesis by (rule_tac x=X in bexI) (auto simp: fbx_def prj1_idem mem_box conj_le_eq inner_add inner_diff fle \X \ \\) qed then show "f ` S \ (\D\\. fbx D)" by auto next have 1: "\D. D \ \ \ fbx D \ lmeasurable" by (auto simp: fbx_def) have 2: "I' \ \ \ finite I' \ measure lebesgue (\D\I'. fbx D) \ e/2" for I' by (rule order_trans[OF measure_Union_le e2]) (auto simp: fbx_def) show "(\D\\. fbx D) \ lmeasurable" by (intro fmeasurable_UN_bound[OF \countable \\ 1 2]) have "measure lebesgue (\D\\. fbx D) \ e/2" by (intro measure_UN_bound[OF \countable \\ 1 2]) then show "measure lebesgue (\D\\. fbx D) < e" using \0 < e\ by linarith qed qed proposition negligible_locally_Lipschitz_image: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and lips: "\x. x \ S \ \T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" proof - let ?S = "\n. ({x \ S. norm x \ n \ (\T. open T \ x \ T \ (\y\S \ T. norm (f y - f x) \ (real n + 1) * norm (y - x)))})" have negfn: "f ` ?S n \ null_sets lebesgue" for n::nat unfolding negligible_iff_null_sets[symmetric] apply (rule_tac B = "real n + 1" in locally_Lipschitz_negl_bounded) by (auto simp: MleN bounded_iff intro: negligible_subset [OF \negligible S\]) have "S = (\n. ?S n)" proof (intro set_eqI iffI) fix x assume "x \ S" with lips obtain T B where T: "open T" "x \ T" and B: "\y. y \ S \ T \ norm(f y - f x) \ B * norm(y - x)" by metis+ have no: "norm (f y - f x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" if "y \ S \ T" for y proof - have "B * norm(y - x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" by (meson max.cobounded1 mult_right_mono nat_ceiling_le_eq nat_le_iff_add norm_ge_zero order_trans) then show ?thesis using B order_trans that by blast qed have "norm x \ real (nat \max B (norm x)\)" by linarith then have "x \ ?S (nat (ceiling (max B (norm x))))" using T no by (force simp: \x \ S\ algebra_simps) then show "x \ (\n. ?S n)" by force qed auto then show ?thesis by (rule ssubst) (auto simp: image_Union negligible_iff_null_sets intro: negfn) qed corollary negligible_differentiable_image_negligible: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "\T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" if "x \ S" for x proof - obtain f' where "linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x)" using diff_f \x \ S\ by (auto simp: linear_linear differentiable_on_def differentiable_def has_derivative_within_alt) obtain B where "B > 0" and B: "\x. norm (f' x) \ B * norm x" using linear_bounded_pos \linear f'\ by blast obtain d where "d>0" and d: "\y. \y \ S; norm (y - x) < d\ \ norm (f y - f x - f' (y - x)) \ norm (y - x)" using f' [of 1] by (force simp:) show ?thesis proof (intro exI conjI ballI) show "norm (f y - f x) \ (B + 1) * norm (y - x)" if "y \ S \ ball x d" for y proof - have "norm (f y - f x) - B * norm (y - x) \ norm (f y - f x) - norm (f' (y - x))" by (simp add: B) also have "\ \ norm (f y - f x - f' (y - x))" by (rule norm_triangle_ineq2) also have "... \ norm (y - x)" by (metis IntE d dist_norm mem_ball norm_minus_commute that) finally show ?thesis by (simp add: algebra_simps) qed qed (use \d>0\ in auto) qed with negligible_locally_Lipschitz_image assms show ?thesis by metis qed corollary negligible_differentiable_image_lowdim: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MlessN: "DIM('M) < DIM('N)" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "x \ DIM('M) \ x \ DIM('N)" for x using MlessN by linarith obtain lift :: "'M * real \ 'N" and drop :: "'N \ 'M * real" and j :: 'N where "linear lift" "linear drop" and dropl [simp]: "\z. drop (lift z) = z" and "j \ Basis" and j: "\x. lift(x,0) \ j = 0" using lowerdim_embeddings [OF MlessN] by metis have "negligible ((\x. lift (x, 0)) ` S)" proof - have "negligible {x. x\j = 0}" by (metis \j \ Basis\ negligible_standard_hyperplane) moreover have "(\m. lift (m, 0)) ` S \ {n. n \ j = 0}" using j by force ultimately show ?thesis using negligible_subset by auto qed moreover have "f \ fst \ drop differentiable_on (\x. lift (x, 0)) ` S" using diff_f apply (clarsimp simp add: differentiable_on_def) apply (intro differentiable_chain_within linear_imp_differentiable [OF \linear drop\] linear_imp_differentiable [OF linear_fst]) apply (force simp: image_comp o_def) done moreover have "f = f \ fst \ drop \ (\x. lift (x, 0))" by (simp add: o_def) ultimately show ?thesis by (metis (no_types) image_comp negligible_differentiable_image_negligible order_refl) qed subsection\Measurability of countable unions and intersections of various kinds.\ lemma assumes S: "\n. S n \ lmeasurable" and leB: "\n. measure lebesgue (S n) \ B" and nest: "\n. S n \ S(Suc n)" shows measurable_nested_Union: "(\n. S n) \ lmeasurable" and measure_nested_Union: "(\n. measure lebesgue (S n)) \ measure lebesgue (\n. S n)" (is ?Lim) proof - have "indicat_real (\ (range S)) integrable_on UNIV \ (\n. integral UNIV (indicat_real (S n))) \ integral UNIV (indicat_real (\ (range S)))" proof (rule monotone_convergence_increasing) show "\n. (indicat_real (S n)) integrable_on UNIV" using S measurable_integrable by blast show "\n x::'a. indicat_real (S n) x \ (indicat_real (S (Suc n)) x)" by (simp add: indicator_leI nest rev_subsetD) have "\x. (\n. x \ S n) \ (\\<^sub>F n in sequentially. x \ S n)" by (metis eventually_sequentiallyI lift_Suc_mono_le nest subsetCE) then show "\x. (\n. indicat_real (S n) x) \ (indicat_real (\(S ` UNIV)) x)" by (simp add: indicator_def tendsto_eventually) show "bounded (range (\n. integral UNIV (indicat_real (S n))))" using leB by (auto simp: lmeasure_integral_UNIV [symmetric] S bounded_iff) qed then have "(\n. S n) \ lmeasurable \ ?Lim" by (simp add: lmeasure_integral_UNIV S cong: conj_cong) (simp add: measurable_integrable) then show "(\n. S n) \ lmeasurable" "?Lim" by auto qed lemma assumes S: "\n. S n \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and leB: "\n. (\k\n. measure lebesgue (S k)) \ B" shows measurable_countable_negligible_Union: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - have 1: "\ (S ` {..n}) \ lmeasurable" for n using S by blast have 2: "measure lebesgue (\ (S ` {..n})) \ B" for n proof - have "measure lebesgue (\ (S ` {..n})) \ (\k\n. measure lebesgue (S k))" by (simp add: S fmeasurableD measure_UNION_le) with leB show ?thesis using order_trans by blast qed have 3: "\n. \ (S ` {..n}) \ \ (S ` {..Suc n})" by (simp add: SUP_subset_mono) have eqS: "(\n. S n) = (\n. \ (S ` {..n}))" using atLeastAtMost_iff by blast also have "(\n. \ (S ` {..n})) \ lmeasurable" by (intro measurable_nested_Union [OF 1 2] 3) finally show "(\n. S n) \ lmeasurable" . have eqm: "(\i\n. measure lebesgue (S i)) = measure lebesgue (\ (S ` {..n}))" for n using assms by (simp add: measure_negligible_finite_Union_image pairwise_mono) have "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. \ (S ` {..n}))" by (simp add: sums_def' eqm atLeast0AtMost) (intro measure_nested_Union [OF 1 2] 3) then show ?Sums by (simp add: eqS) qed lemma negligible_countable_Union [intro]: assumes "countable \" and meas: "\S. S \ \ \ negligible S" shows "negligible (\\)" proof (cases "\ = {}") case False then show ?thesis by (metis from_nat_into range_from_nat_into assms negligible_Union_nat) qed simp lemma assumes S: "\n. (S n) \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and bo: "bounded (\n. S n)" shows measurable_countable_negligible_Union_bounded: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union_bounded: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - obtain a b where ab: "(\n. S n) \ cbox a b" using bo bounded_subset_cbox_symmetric by metis then have B: "(\k\n. measure lebesgue (S k)) \ measure lebesgue (cbox a b)" for n proof - have "(\k\n. measure lebesgue (S k)) = measure lebesgue (\ (S ` {..n}))" using measure_negligible_finite_Union_image [OF _ _ pairwise_subset] djointish by (metis S finite_atMost subset_UNIV) also have "\ \ measure lebesgue (cbox a b)" proof (rule measure_mono_fmeasurable) show "\ (S ` {..n}) \ sets lebesgue" using S by blast qed (use ab in auto) finally show ?thesis . qed show "(\n. S n) \ lmeasurable" by (rule measurable_countable_negligible_Union [OF S djointish B]) show ?Sums by (rule measure_countable_negligible_Union [OF S djointish B]) qed lemma measure_countable_Union_approachable: assumes "countable \" "e > 0" and measD: "\d. d \ \ \ d \ lmeasurable" and B: "\D'. \D' \ \; finite D'\ \ measure lebesgue (\D') \ B" obtains D' where "D' \ \" "finite D'" "measure lebesgue (\\) - e < measure lebesgue (\D')" proof (cases "\ = {}") case True then show ?thesis by (simp add: \e > 0\ that) next case False let ?S = "\n. \k \ n. from_nat_into \ k" have "(\n. measure lebesgue (?S n)) \ measure lebesgue (\n. ?S n)" proof (rule measure_nested_Union) show "?S n \ lmeasurable" for n by (simp add: False fmeasurable.finite_UN from_nat_into measD) show "measure lebesgue (?S n) \ B" for n by (metis (mono_tags, lifting) B False finite_atMost finite_imageI from_nat_into image_iff subsetI) show "?S n \ ?S (Suc n)" for n by force qed then obtain N where N: "\n. n \ N \ dist (measure lebesgue (?S n)) (measure lebesgue (\n. ?S n)) < e" using metric_LIMSEQ_D \e > 0\ by blast show ?thesis proof show "from_nat_into \ ` {..N} \ \" by (auto simp: False from_nat_into) have eq: "(\n. \k\n. from_nat_into \ k) = (\\)" using \countable \\ False by (auto intro: from_nat_into dest: from_nat_into_surj [OF \countable \\]) show "measure lebesgue (\\) - e < measure lebesgue (\ (from_nat_into \ ` {..N}))" using N [OF order_refl] by (auto simp: eq algebra_simps dist_norm) qed auto qed subsection\Negligibility is a local property\ lemma locally_negligible_alt: "negligible S \ (\x \ S. \U. openin (top_of_set S) U \ x \ U \ negligible U)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs using openin_subtopology_self by blast next assume ?rhs then obtain U where ope: "\x. x \ S \ openin (top_of_set S) (U x)" and cov: "\x. x \ S \ x \ U x" and neg: "\x. x \ S \ negligible (U x)" by metis obtain \ where "\ \ U ` S" "countable \" and eq: "\\ = \(U ` S)" using ope by (force intro: Lindelof_openin [of "U ` S" S]) then have "negligible (\\)" by (metis imageE neg negligible_countable_Union subset_eq) with eq have "negligible (\(U ` S))" by metis moreover have "S \ \(U ` S)" using cov by blast ultimately show "negligible S" using negligible_subset by blast qed lemma locally_negligible: "locally negligible S \ negligible S" unfolding locally_def by (metis locally_negligible_alt negligible_subset openin_imp_subset openin_subtopology_self) subsection\Integral bounds\ lemma set_integral_norm_bound: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" shows "set_integrable M k f \ norm (LINT x:k|M. f x) \ LINT x:k|M. norm (f x)" using integral_norm_bound[of M "\x. indicator k x *\<^sub>R f x"] by (simp add: set_lebesgue_integral_def) lemma set_integral_finite_UN_AE: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" and ae: "\i j. i \ I \ j \ I \ AE x in M. (x \ A i \ x \ A j) \ i = j" and [measurable]: "\i. i \ I \ A i \ sets M" and f: "\i. i \ I \ set_integrable M (A i) f" shows "LINT x:(\i\I. A i)|M. f x = (\i\I. LINT x:A i|M. f x)" using \finite I\ order_refl[of I] proof (induction I rule: finite_subset_induct') case (insert i I') have "AE x in M. (\j\I'. x \ A i \ x \ A j)" proof (intro AE_ball_countable[THEN iffD2] ballI) fix j assume "j \ I'" with \I' \ I\ \i \ I'\ have "i \ j" "j \ I" by auto then show "AE x in M. x \ A i \ x \ A j" using ae[of i j] \i \ I\ by auto qed (use \finite I'\ in \rule countable_finite\) then have "AE x\A i in M. \xa\I'. x \ A xa " by auto with insert.hyps insert.IH[symmetric] show ?case by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN) qed (simp add: set_lebesgue_integral_def) lemma set_integrable_norm: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f: "set_integrable M k f" shows "set_integrable M k (\x. norm (f x))" using integrable_norm f by (force simp add: set_integrable_def) lemma absolutely_integrable_bounded_variation: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on UNIV" obtains B where "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" proof (rule that[of "integral UNIV (\x. norm (f x))"]; safe) fix d :: "'a set set" assume d: "d division_of \d" have *: "k \ d \ f absolutely_integrable_on k" for k using f[THEN set_integrable_subset, of k] division_ofD(2,4)[OF d, of k] by auto note d' = division_ofD[OF d] have "(\k\d. norm (integral k f)) = (\k\d. norm (LINT x:k|lebesgue. f x))" by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *) also have "\ \ (\k\d. LINT x:k|lebesgue. norm (f x))" by (intro sum_mono set_integral_norm_bound *) also have "\ = (\k\d. integral k (\x. norm (f x)))" by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *) also have "\ \ integral (\d) (\x. norm (f x))" using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def by (subst integral_combine_division_topdown[OF _ d]) auto also have "\ \ integral UNIV (\x. norm (f x))" using integrable_on_subdivision[OF d] assms unfolding absolutely_integrable_on_def by (intro integral_subset_le) auto finally show "(\k\d. norm (integral k f)) \ integral UNIV (\x. norm (f x))" . qed lemma absdiff_norm_less: assumes "sum (\x. norm (f x - g x)) S < e" shows "\sum (\x. norm(f x)) S - sum (\x. norm(g x)) S\ < e" (is "?lhs < e") proof - have "?lhs \ (\i\S. \norm (f i) - norm (g i)\)" by (metis (no_types) sum_abs sum_subtractf) also have "... \ (\x\S. norm (f x - g x))" by (simp add: norm_triangle_ineq3 sum_mono) also have "... < e" using assms(1) by blast finally show ?thesis . qed proposition bounded_variation_absolutely_integrable_interval: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes f: "f integrable_on cbox a b" and *: "\d. d division_of (cbox a b) \ sum (\K. norm(integral K f)) d \ B" shows "f absolutely_integrable_on cbox a b" proof - let ?f = "\d. \K\d. norm (integral K f)" and ?D = "{d. d division_of (cbox a b)}" have D_1: "?D \ {}" by (rule elementary_interval[of a b]) auto have D_2: "bdd_above (?f`?D)" by (metis * mem_Collect_eq bdd_aboveI2) note D = D_1 D_2 let ?S = "SUP x\?D. ?f x" have *: "\\. gauge \ \ (\p. p tagged_division_of cbox a b \ \ fine p \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e)" if e: "e > 0" for e proof - have "?S - e/2 < ?S" using \e > 0\ by simp then obtain d where d: "d division_of (cbox a b)" "?S - e/2 < (\k\d. norm (integral k f))" unfolding less_cSUP_iff[OF D] by auto note d' = division_ofD[OF this(1)] have "\e>0. \i\d. x \ i \ ball x e \ i = {}" for x proof - have "\d'>0. \x'\\{i \ d. x \ i}. d' \ dist x x'" proof (rule separate_point_closed) show "closed (\{i \ d. x \ i})" using d' by force show "x \ \{i \ d. x \ i}" by auto qed then show ?thesis by force qed then obtain k where k: "\x. 0 < k x" "\i x. \i \ d; x \ i\ \ ball x (k x) \ i = {}" by metis have "e/2 > 0" using e by auto with Henstock_lemma[OF f] obtain \ where g: "gauge \" "\p. \p tagged_partial_division_of cbox a b; \ fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (metis (no_types, lifting)) let ?g = "\x. \ x \ ball x (k x)" show ?thesis proof (intro exI conjI allI impI) show "gauge ?g" using g(1) k(1) by (auto simp: gauge_def) next fix p assume "p tagged_division_of (cbox a b) \ ?g fine p" then have p: "p tagged_division_of cbox a b" "\ fine p" "(\x. ball x (k x)) fine p" by (auto simp: fine_Int) note p' = tagged_division_ofD[OF p(1)] define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" have gp': "\ fine p'" using p(2) by (auto simp: p'_def fine_def) have p'': "p' tagged_division_of (cbox a b)" proof (rule tagged_division_ofI) show "finite p'" proof (rule finite_subset) show "p' \ (\(k, x, l). (x, k \ l)) ` (d \ p)" by (force simp: p'_def image_iff) show "finite ((\(k, x, l). (x, k \ l)) ` (d \ p))" by (simp add: d'(1) p'(1)) qed next fix x K assume "(x, K) \ p'" then have "\i l. x \ i \ i \ d \ (x, l) \ p \ K = i \ l" unfolding p'_def by auto then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" by blast show "x \ K" and "K \ cbox a b" using p'(2-3)[OF il(3)] il by auto show "\a b. K = cbox a b" unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] by (meson Int_interval) next fix x1 K1 assume "(x1, K1) \ p'" then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ K1 = i \ l" unfolding p'_def by auto then obtain i1 l1 where il1: "x1 \ i1" "i1 \ d" "(x1, l1) \ p" "K1 = i1 \ l1" by blast fix x2 K2 assume "(x2,K2) \ p'" then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ K2 = i \ l" unfolding p'_def by auto then obtain i2 l2 where il2: "x2 \ i2" "i2 \ d" "(x2, l2) \ p" "K2 = i2 \ l2" by blast assume "(x1, K1) \ (x2, K2)" then have "interior i1 \ interior i2 = {} \ interior l1 \ interior l2 = {}" using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] by (auto simp: il1 il2) then show "interior K1 \ interior K2 = {}" unfolding il1 il2 by auto next have *: "\(x, X) \ p'. X \ cbox a b" unfolding p'_def using d' by blast show "\{K. \x. (x, K) \ p'} = cbox a b" proof show "\{k. \x. (x, k) \ p'} \ cbox a b" using * by auto next show "cbox a b \ \{k. \x. (x, k) \ p'}" proof fix y assume y: "y \ cbox a b" obtain x L where xl: "(x, L) \ p" "y \ L" using y unfolding p'(6)[symmetric] by auto obtain I where i: "I \ d" "y \ I" using y unfolding d'(6)[symmetric] by auto have "x \ I" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show "y \ \{K. \x. (x, K) \ p'}" proof - obtain x l where xl: "(x, l) \ p" "y \ l" using y unfolding p'(6)[symmetric] by auto obtain i where i: "i \ d" "y \ i" using y unfolding d'(6)[symmetric] by auto have "x \ i" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show ?thesis unfolding p'_def by (rule_tac X="i \ l" in UnionI) (use i xl in auto) qed qed qed qed then have sum_less_e2: "(\(x,K) \ p'. norm (content K *\<^sub>R f x - integral K f)) < e/2" using g(2) gp' tagged_division_of_def by blast have in_p': "(x, I \ L) \ p'" if x: "(x, L) \ p" "I \ d" and y: "y \ I" "y \ L" for x I L y proof - have "x \ I" using fineD[OF p(3) that(1)] k(2)[OF \I \ d\] y by auto with x have "(\i l. x \ i \ i \ d \ (x, l) \ p \ I \ L = i \ l)" by blast then have "(x, I \ L) \ p'" by (simp add: p'_def) with y show ?thesis by auto qed moreover have Ex_p_p': "\y i l. (x, K) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" if xK: "(x,K) \ p'" for x K proof - obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" using xK unfolding p'_def by auto then show ?thesis using p'(2) by fastforce qed ultimately have p'alt: "p' = {(x, I \ L) | x I L. (x,L) \ p \ I \ d \ I \ L \ {}}" by auto have sum_p': "(\(x,K) \ p'. norm (integral K f)) = (\k\snd ` p'. norm (integral k f))" proof (rule sum.over_tagged_division_lemma[OF p'']) show "\u v. box u v = {} \ norm (integral (cbox u v) f) = 0" by (auto intro: integral_null simp: content_eq_0_interior) qed have snd_p_div: "snd ` p division_of cbox a b" by (rule division_of_tagged_division[OF p(1)]) note snd_p = division_ofD[OF snd_p_div] have fin_d_sndp: "finite (d \ snd ` p)" by (simp add: d'(1) snd_p(1)) have *: "\sni sni' sf sf'. \\sf' - sni'\ < e/2; ?S - e/2 < sni; sni' \ ?S; sni \ sni'; sf' = sf\ \ \sf - ?S\ < e" by arith show "norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e" unfolding real_norm_def proof (rule *) show "\(\(x,K)\p'. norm (content K *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e/2" using p'' sum_less_e2 unfolding split_def by (force intro!: absdiff_norm_less) show "(\(x,k) \ p'. norm (integral k f)) \?S" by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) show "(\k\d. norm (integral k f)) \ (\(x,k) \ p'. norm (integral k f))" proof - have *: "{k \ l | k l. k \ d \ l \ snd ` p} = (\(k,l). k \ l) ` (d \ snd ` p)" by auto have "(\K\d. norm (integral K f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" proof (rule sum_mono) fix K assume k: "K \ d" from d'(4)[OF this] obtain u v where uv: "K = cbox u v" by metis define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" have uvab: "cbox u v \ cbox a b" using d(1) k uv by blast have d'_div: "d' division_of cbox u v" unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) moreover have "norm (\i\d'. integral i f) \ (\k\d'. norm (integral k f))" by (simp add: sum_norm_le) moreover have "f integrable_on K" using f integrable_on_subcbox uv uvab by blast moreover have "d' division_of K" using d'_div uv by blast ultimately have "norm (integral K f) \ sum (\k. norm (integral k f)) d'" by (simp add: integral_combine_division_topdown) also have "\ = (\I\{K \ L |L. L \ snd ` p}. norm (integral I f))" proof (rule sum.mono_neutral_left) show "finite {K \ L |L. L \ snd ` p}" by (simp add: snd_p(1)) show "\i\{K \ L |L. L \ snd ` p} - d'. norm (integral i f) = 0" "d' \ {K \ L |L. L \ snd ` p}" using d'_def image_eqI uv by auto qed also have "\ = (\l\snd ` p. norm (integral (K \ l) f))" unfolding Setcompr_eq_image proof (rule sum.reindex_nontrivial [unfolded o_def]) show "finite (snd ` p)" using snd_p(1) by blast show "norm (integral (K \ l) f) = 0" if "l \ snd ` p" "y \ snd ` p" "l \ y" "K \ l = K \ y" for l y proof - have "interior (K \ l) \ interior (l \ y)" by (metis Int_lower2 interior_mono le_inf_iff that(4)) then have "interior (K \ l) = {}" by (simp add: snd_p(5) that) moreover from d'(4)[OF k] snd_p(4)[OF that(1)] obtain u1 v1 u2 v2 where uv: "K = cbox u1 u2" "l = cbox v1 v2" by metis ultimately show ?thesis using that integral_null unfolding uv Int_interval content_eq_0_interior by (metis (mono_tags, lifting) norm_eq_zero) qed qed finally show "norm (integral K f) \ (\l\snd ` p. norm (integral (K \ l) f))" . qed also have "\ = (\(i,l) \ d \ snd ` p. norm (integral (i\l) f))" by (simp add: sum.cartesian_product) also have "\ = (\x \ d \ snd ` p. norm (integral (case_prod (\) x) f))" by (force simp: split_def intro!: sum.cong) also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k f))" proof - have eq0: " (integral (l1 \ k1) f) = 0" if "l1 \ k1 = l2 \ k2" "(l1, k1) \ (l2, k2)" "l1 \ d" "(j1,k1) \ p" "l2 \ d" "(j2,k2) \ p" for l1 l2 k1 k2 j1 j2 proof - obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2" using \(j1, k1) \ p\ \l1 \ d\ d'(4) p'(4) by blast have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6)) moreover have "interior (l1 \ k1) = interior (l2 \ k2)" by (simp add: that(1)) ultimately have "interior(l1 \ k1) = {}" by auto then show ?thesis unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed show ?thesis unfolding * apply (rule sum.reindex_nontrivial [OF fin_d_sndp, symmetric, unfolded o_def]) apply clarsimp by (metis eq0 fst_conv snd_conv) qed also have "\ = (\(x,k) \ p'. norm (integral k f))" unfolding sum_p' proof (rule sum.mono_neutral_right) show "finite {i \ l |i l. i \ d \ l \ snd ` p}" by (metis * finite_imageI[OF fin_d_sndp]) show "snd ` p' \ {i \ l |i l. i \ d \ l \ snd ` p}" by (clarsimp simp: p'_def) (metis image_eqI snd_conv) show "\i\{i \ l |i l. i \ d \ l \ snd ` p} - snd ` p'. norm (integral i f) = 0" by clarsimp (metis Henstock_Kurzweil_Integration.integral_empty disjoint_iff image_eqI in_p' snd_conv) qed finally show ?thesis . qed show "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" by force have fin_pd: "finite (p \ d)" using finite_cartesian_product[OF p'(1) d'(1)] by metis have "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ ?S. \content k\ * norm (f x))" unfolding norm_scaleR proof (rule sum.mono_neutral_left) show "finite {(x, i \ l) |x i l. (x, l) \ p \ i \ d}" by (simp add: "*" fin_pd) qed (use p'alt in \force+\) also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" proof - have "\content (l1 \ k1)\ * norm (f x1) = 0" if "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" "x1 = x2" "l1 \ k1 = l2 \ k2" "x1 \ x2 \ l1 \ l2 \ k1 \ k2" for x1 l1 k1 x2 l2 k2 proof - obtain u1 v1 u2 v2 where uv: "k1 = cbox u1 u2" "l1 = cbox v1 v2" by (meson \(x1, l1) \ p\ \k1 \ d\ d(1) division_ofD(4) p'(4)) have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" using that p'(5) d'(5) by (metis snd_conv) moreover have "interior (l1 \ k1) = interior (l2 \ k2)" unfolding that .. ultimately have "interior (l1 \ k1) = {}" by auto then show "\content (l1 \ k1)\ * norm (f x1) = 0" unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed then show ?thesis unfolding * apply (subst sum.reindex_nontrivial [OF fin_pd]) unfolding split_paired_all o_def split_def prod.inject by force+ qed also have "\ = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - have sumeq: "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" if "(x, l) \ p" for x l proof - note xl = p'(2-4)[OF that] then obtain u v where uv: "l = cbox u v" by blast have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" by (simp add: Int_commute uv) also have "\ = sum content {k \ cbox u v| k. k \ d}" proof - have eq0: "content (k \ cbox u v) = 0" if "k \ d" "y \ d" "k \ y" and eq: "k \ cbox u v = y \ cbox u v" for k y proof - from d'(4)[OF that(1)] d'(4)[OF that(2)] obtain \ \ where \: "k \ cbox u v = cbox \ \" by (meson Int_interval) have "{} = interior ((k \ y) \ cbox u v)" by (simp add: d'(5) that) also have "\ = interior (y \ (k \ cbox u v))" by auto also have "\ = interior (k \ cbox u v)" unfolding eq by auto finally show ?thesis unfolding \ content_eq_0_interior .. qed then show ?thesis unfolding Setcompr_eq_image by (fastforce intro: sum.reindex_nontrivial [OF \finite d\, unfolded o_def, symmetric]) qed also have "\ = sum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" proof (rule sum.mono_neutral_right) show "finite {k \ cbox u v |k. k \ d}" by (simp add: d'(1)) qed (fastforce simp: inf.commute)+ finally have "(\i\d. \content (l \ i)\) = content (cbox u v)" using additive_content_division[OF division_inter_1[OF d(1)]] uv xl(2) by auto then show "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" unfolding sum_distrib_right[symmetric] using uv by auto qed show ?thesis by (subst sum_Sigma_product[symmetric]) (auto intro!: sumeq sum.cong p' d') qed finally show ?thesis . qed qed (rule d) qed qed then show ?thesis using absolutely_integrable_onI [OF f has_integral_integrable] has_integral[of _ ?S] by blast qed lemma bounded_variation_absolutely_integrable: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on UNIV" and "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" shows "f absolutely_integrable_on UNIV" proof (rule absolutely_integrable_onI, fact) let ?f = "\D. \k\D. norm (integral k f)" and ?D = "{d. d division_of (\d)}" define SDF where "SDF \ SUP d\?D. ?f d" have D_1: "?D \ {}" by (rule elementary_interval) auto have D_2: "bdd_above (?f`?D)" using assms(2) by auto have f_int: "\a b. f absolutely_integrable_on cbox a b" using assms integrable_on_subcbox by (blast intro!: bounded_variation_absolutely_integrable_interval) have "\B>0. \a b. ball 0 B \ cbox a b \ \integral (cbox a b) (\x. norm (f x)) - SDF\ < e" if "0 < e" for e proof - have "\y \ ?f ` ?D. \ y \ SDF - e" proof (rule ccontr) assume "\ ?thesis" then have "SDF \ SDF - e" unfolding SDF_def by (metis (mono_tags) D_1 cSUP_least image_eqI) then show False using that by auto qed then obtain d K where ddiv: "d division_of \d" and "K = ?f d" "SDF - e < K" by (auto simp add: image_iff not_le) then have d: "SDF - e < ?f d" by auto note d'=division_ofD[OF ddiv] have "bounded (\d)" using ddiv by blast then obtain K where K: "0 < K" "\x\\d. norm x \ K" using bounded_pos by blast show ?thesis proof (intro conjI impI allI exI) fix a b :: 'n assume ab: "ball 0 (K + 1) \ cbox a b" have *: "\s s1. \SDF - e < s1; s1 \ s; s < SDF + e\ \ \s - SDF\ < e" by arith show "\integral (cbox a b) (\x. norm (f x)) - SDF\ < e" unfolding real_norm_def proof (rule * [OF d]) have "?f d \ sum (\k. integral k (\x. norm (f x))) d" proof (intro sum_mono) fix k assume "k \ d" with d'(4) f_int show "norm (integral k f) \ integral k (\x. norm (f x))" by (force simp: absolutely_integrable_on_def integral_norm_bound_integral) qed also have "\ = integral (\d) (\x. norm (f x))" by (metis (full_types) absolutely_integrable_on_def d'(4) ddiv f_int integral_combine_division_bottomup) also have "\ \ integral (cbox a b) (\x. norm (f x))" proof - have "\d \ cbox a b" using K(2) ab by fastforce then show ?thesis using integrable_on_subdivision[OF ddiv] f_int[of a b] unfolding absolutely_integrable_on_def by (auto intro!: integral_subset_le) qed finally show "?f d \ integral (cbox a b) (\x. norm (f x))" . next have "e/2>0" using \e > 0\ by auto moreover have f: "f integrable_on cbox a b" "(\x. norm (f x)) integrable_on cbox a b" using f_int by (auto simp: absolutely_integrable_on_def) ultimately obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of (cbox a b); d1 fine p\ \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm (f x))) < e/2" unfolding has_integral_integral has_integral by meson obtain d2 where "gauge d2" and d2: "\p. \p tagged_partial_division_of (cbox a b); d2 fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (blast intro: Henstock_lemma [OF f(1) \e/2>0\]) obtain p where p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" by (rule fine_division_exists [OF gauge_Int [OF \gauge d1\ \gauge d2\], of a b]) (auto simp add: fine_Int) have *: "\sf sf' si di. \sf' = sf; si \ SDF; \sf - si\ < e/2; \sf' - di\ < e/2\ \ di < SDF + e" by arith have "integral (cbox a b) (\x. norm (f x)) < SDF + e" proof (rule *) show "\(\(x,k)\p. norm (content k *\<^sub>R f x)) - (\(x,k)\p. norm (integral k f))\ < e/2" unfolding split_def proof (rule absdiff_norm_less) show "(\p\p. norm (content (snd p) *\<^sub>R f (fst p) - integral (snd p) f)) < e/2" using d2[of p] p(1,3) by (auto simp: tagged_division_of_def split_def) qed show "\(\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm(f x))\ < e/2" using d1[OF p(1,2)] by (simp only: real_norm_def) show "(\(x,k) \ p. content k *\<^sub>R norm (f x)) = (\(x,k) \ p. norm (content k *\<^sub>R f x))" by (auto simp: split_paired_all sum.cong [OF refl]) have "(\(x,k) \ p. norm (integral k f)) = (\k\snd ` p. norm (integral k f))" apply (rule sum.over_tagged_division_lemma[OF p(1)]) by (metis Henstock_Kurzweil_Integration.integral_empty integral_open_interval norm_zero) also have "... \ SDF" using partial_division_of_tagged_division[of p "cbox a b"] p(1) by (auto simp: SDF_def tagged_partial_division_of_def intro!: cSUP_upper2 D_1 D_2) finally show "(\(x,k) \ p. norm (integral k f)) \ SDF" . qed then show "integral (cbox a b) (\x. norm (f x)) < SDF + e" by simp qed qed (use K in auto) qed moreover have "\a b. (\x. norm (f x)) integrable_on cbox a b" using absolutely_integrable_on_def f_int by auto ultimately have "((\x. norm (f x)) has_integral SDF) UNIV" by (auto simp: has_integral_alt') then show "(\x. norm (f x)) integrable_on UNIV" by blast qed subsection\Outer and inner approximation of measurable sets by well-behaved sets.\ proposition measurable_outer_intervals_bounded: assumes "S \ lmeasurable" "S \ cbox a b" "e > 0" obtains \ where "countable \" "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" "pairwise (\A B. interior A \ interior B = {}) \" "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i)/2^n" "\K. \K \ \; box a b \ {}\ \ interior K \ {}" "S \ \\" "\\ \ lmeasurable" "measure lebesgue (\\) \ measure lebesgue S + e" proof (cases "box a b = {}") case True show ?thesis proof (cases "cbox a b = {}") case True with assms have [simp]: "S = {}" by auto show ?thesis proof show "countable {}" by simp qed (use \e > 0\ in auto) next case False show ?thesis proof show "countable {cbox a b}" by simp show "\u v. cbox u v \ {cbox a b} \ \n. \i\Basis. v \ i - u \ i = (b \ i - a \ i)/2 ^ n" using False by (force simp: eq_cbox intro: exI [where x=0]) show "measure lebesgue (\{cbox a b}) \ measure lebesgue S + e" using assms by (simp add: sum_content.box_empty_imp [OF True]) qed (use assms \cbox a b \ {}\ in auto) qed next case False let ?\ = "measure lebesgue" have "S \ cbox a b \ lmeasurable" using \S \ lmeasurable\ by blast then have indS_int: "(indicator S has_integral (?\ S)) (cbox a b)" by (metis integral_indicator \S \ cbox a b\ has_integral_integrable_integral inf.orderE integrable_on_indicator) with \e > 0\ obtain \ where "gauge \" and \: "\\. \\ tagged_division_of (cbox a b); \ fine \\ \ norm ((\(x,K)\\. content(K) *\<^sub>R indicator S x) - ?\ S) < e" by (force simp: has_integral) have inteq: "integral (cbox a b) (indicat_real S) = integral UNIV (indicator S)" using assms by (metis has_integral_iff indS_int lmeasure_integral_UNIV) obtain \ where \: "countable \" "\\ \ cbox a b" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ \ x" and close: "\u v. cbox u v \ \ \ \n. \i \ Basis. v\i - u\i = (b\i - a\i)/2^n" and covers: "S \ \\" using covering_lemma [of S a b \] \gauge \\ \box a b \ {}\ assms by force show ?thesis proof show "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" by (meson Sup_le_iff \(2) cbox interior_empty) have negl_int: "negligible(K \ L)" if "K \ \" "L \ \" "K \ L" for K L proof - have "interior K \ interior L = {}" using djointish pairwiseD that by fastforce moreover obtain u v x y where "K = cbox u v" "L = cbox x y" using cbox \K \ \\ \L \ \\ by blast ultimately show ?thesis by (simp add: Int_interval box_Int_box negligible_interval(1)) qed have fincase: "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" if "finite \" "\ \ \" for \ proof - obtain t where t: "\K. K \ \ \ t K \ S \ K \ K \ \(t K)" using covered \\ \ \\ subsetD by metis have "\K \ \. \L \ \. K \ L \ interior K \ interior L = {}" using that djointish by (simp add: pairwise_def) (metis subsetD) with cbox that \ have \div: "\ division_of (\\)" by (fastforce simp: division_of_def dest: cbox) then have 1: "\\ \ lmeasurable" by blast have norme: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x,K)\p. content K * indicator S x) - integral (cbox a b) (indicator S)) < e" by (auto simp: lmeasure_integral_UNIV assms inteq dest: \) have "\x K y L. (x,K) \ (\K. (t K,K)) ` \ \ (y,L) \ (\K. (t K,K)) ` \ \ (x,K) \ (y,L) \ interior K \ interior L = {}" using that djointish by (clarsimp simp: pairwise_def) (metis subsetD) with that \ have tagged: "(\K. (t K, K)) ` \ tagged_partial_division_of cbox a b" by (auto simp: tagged_partial_division_of_def dest: t cbox) have fine: "\ fine (\K. (t K, K)) ` \" using t by (auto simp: fine_def) have *: "y \ ?\ S \ \x - y\ \ e \ x \ ?\ S + e" for x y by arith have "?\ (\\) \ ?\ S + e" proof (rule *) have "(\K\\. ?\ (K \ S)) = ?\ (\C\\. C \ S)" proof (rule measure_negligible_finite_Union_image [OF \finite \\, symmetric]) show "\K. K \ \ \ K \ S \ lmeasurable" using \div \S \ lmeasurable\ by blast show "pairwise (\K y. negligible (K \ S \ (y \ S))) \" unfolding pairwise_def by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \\ \ \\) qed also have "\ = ?\ (\\ \ S)" by simp also have "\ \ ?\ S" by (simp add: "1" \S \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Int) finally show "(\K\\. ?\ (K \ S)) \ ?\ S" . next have "?\ (\\) = sum ?\ \" by (metis \div content_division) also have "\ = (\K\\. content K)" using \div by (force intro: sum.cong) also have "\ = (\x\\. content x * indicator S (t x))" using t by auto finally have eq1: "?\ (\\) = (\x\\. content x * indicator S (t x))" . have eq2: "(\K\\. ?\ (K \ S)) = (\K\\. integral K (indicator S))" apply (rule sum.cong [OF refl]) by (metis integral_indicator \div \S \ lmeasurable\ division_ofD(4) fmeasurable.Int inf.commute lmeasurable_cbox) have "\\(x,K)\(\K. (t K, K)) ` \. content K * indicator S x - integral K (indicator S)\ \ e" using Henstock_lemma_part1 [of "indicator S::'a\real", OF _ \e > 0\ \gauge \\ _ tagged fine] indS_int norme by auto then show "\?\ (\\) - (\K\\. ?\ (K \ S))\ \ e" by (simp add: eq1 eq2 comm_monoid_add_class.sum.reindex inj_on_def sum_subtractf) qed with 1 show ?thesis by blast qed have "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" proof (cases "finite \") case True with fincase show ?thesis by blast next case False let ?T = "from_nat_into \" have T: "bij_betw ?T UNIV \" by (simp add: False \(1) bij_betw_from_nat_into) have TM: "\n. ?T n \ lmeasurable" by (metis False cbox finite.emptyI from_nat_into lmeasurable_cbox) have TN: "\m n. m \ n \ negligible (?T m \ ?T n)" by (simp add: False \(1) from_nat_into infinite_imp_nonempty negl_int) have TB: "(\k\n. ?\ (?T k)) \ ?\ S + e" for n proof - have "(\k\n. ?\ (?T k)) = ?\ (\ (?T ` {..n}))" by (simp add: pairwise_def TM TN measure_negligible_finite_Union_image) also have "?\ (\ (?T ` {..n})) \ ?\ S + e" using fincase [of "?T ` {..n}"] T by (auto simp: bij_betw_def) finally show ?thesis . qed have "\\ \ lmeasurable" by (metis lmeasurable_compact T \(2) bij_betw_def cbox compact_cbox countable_Un_Int(1) fmeasurableD fmeasurableI2 rangeI) moreover have "?\ (\x. from_nat_into \ x) \ ?\ S + e" proof (rule measure_countable_Union_le [OF TM]) show "?\ (\x\n. from_nat_into \ x) \ ?\ S + e" for n by (metis (mono_tags, lifting) False fincase finite.emptyI finite_atMost finite_imageI from_nat_into imageE subsetI) qed ultimately show ?thesis by (metis T bij_betw_def) qed then show "\\ \ lmeasurable" "measure lebesgue (\\) \ ?\ S + e" by blast+ qed (use \ cbox djointish close covers in auto) qed subsection\Transformation of measure by linear maps\ lemma emeasure_lebesgue_ball_conv_unit_ball: fixes c :: "'a :: euclidean_space" assumes "r \ 0" shows "emeasure lebesgue (ball c r) = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" proof (cases "r = 0") case False with assms have r: "r > 0" by auto have "emeasure lebesgue ((\x. c + x) ` (\x. r *\<^sub>R x) ` ball (0 :: 'a) 1) = r ^ DIM('a) * emeasure lebesgue (ball (0 :: 'a) 1)" unfolding image_image using emeasure_lebesgue_affine[of r c "ball 0 1"] assms by (simp add: add_ac) also have "(\x. r *\<^sub>R x) ` ball 0 1 = ball (0 :: 'a) r" using r by (subst ball_scale) auto also have "(\x. c + x) ` \ = ball c r" by (subst image_add_ball) (simp_all add: algebra_simps) finally show ?thesis by simp qed auto lemma content_ball_conv_unit_ball: fixes c :: "'a :: euclidean_space" assumes "r \ 0" shows "content (ball c r) = r ^ DIM('a) * content (ball (0 :: 'a) 1)" proof - have "ennreal (content (ball c r)) = emeasure lebesgue (ball c r)" using emeasure_lborel_ball_finite[of c r] by (subst emeasure_eq_ennreal_measure) auto also have "\ = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" using assms by (intro emeasure_lebesgue_ball_conv_unit_ball) auto also have "\ = ennreal (r ^ DIM('a) * content (ball (0::'a) 1))" using emeasure_lborel_ball_finite[of "0::'a" 1] assms by (subst emeasure_eq_ennreal_measure) (auto simp: ennreal_mult') finally show ?thesis using assms by (subst (asm) ennreal_inj) auto qed lemma measurable_linear_image_interval: "linear f \ f ` (cbox a b) \ lmeasurable" by (metis bounded_linear_image linear_linear bounded_cbox closure_bounded_linear_image closure_cbox compact_closure lmeasurable_compact) proposition measure_linear_sufficient: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" and S: "S \ lmeasurable" and im: "\a b. measure lebesgue (f ` (cbox a b)) = m * measure lebesgue (cbox a b)" shows "f ` S \ lmeasurable \ m * measure lebesgue S = measure lebesgue (f ` S)" using le_less_linear [of 0 m] proof assume "m < 0" then show ?thesis using im [of 0 One] by auto next assume "m \ 0" let ?\ = "measure lebesgue" show ?thesis proof (cases "inj f") case False then have "?\ (f ` S) = 0" using \linear f\ negligible_imp_measure0 negligible_linear_singular_image by blast then have "m * ?\ (cbox 0 (One)) = 0" by (metis False \linear f\ cbox_borel content_unit im measure_completion negligible_imp_measure0 negligible_linear_singular_image sets_lborel) then show ?thesis using \linear f\ negligible_linear_singular_image negligible_imp_measure0 False by (auto simp: lmeasurable_iff_has_integral negligible_UNIV) next case True then obtain h where "linear h" and hf: "\x. h (f x) = x" and fh: "\x. f (h x) = x" using \linear f\ linear_injective_isomorphism by blast have fBS: "(f ` S) \ lmeasurable \ m * ?\ S = ?\ (f ` S)" if "bounded S" "S \ lmeasurable" for S proof - obtain a b where "S \ cbox a b" using \bounded S\ bounded_subset_cbox_symmetric by metis have fUD: "(f ` \\) \ lmeasurable \ ?\ (f ` \\) = (m * ?\ (\\))" if "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intint: "pairwise (\A B. interior A \ interior B = {}) \" for \ proof - have conv: "\K. K \ \ \ convex K" using cbox convex_box(1) by blast have neg: "negligible (g ` K \ g ` L)" if "linear g" "K \ \" "L \ \" "K \ L" for K L and g :: "'n\'n" proof (cases "inj g") case True have "negligible (frontier(g ` K \ g ` L) \ interior(g ` K \ g ` L))" proof (rule negligible_Un) show "negligible (frontier (g ` K \ g ` L))" by (simp add: negligible_convex_frontier convex_Int conv convex_linear_image that) next have "\p N. pairwise p N = (\Na. (Na::'n set) \ N \ (\Nb. Nb \ N \ Na \ Nb \ p Na Nb))" by (metis pairwise_def) then have "interior K \ interior L = {}" using intint that(2) that(3) that(4) by presburger then show "negligible (interior (g ` K \ g ` L))" by (metis True empty_imp_negligible image_Int image_empty interior_Int interior_injective_linear_image that(1)) qed moreover have "g ` K \ g ` L \ frontier (g ` K \ g ` L) \ interior (g ` K \ g ` L)" by (metis Diff_partition Int_commute calculation closure_Un_frontier frontier_def inf.absorb_iff2 inf_bot_right inf_sup_absorb negligible_Un_eq open_interior open_not_negligible sup_commute) ultimately show ?thesis by (rule negligible_subset) next case False then show ?thesis by (simp add: negligible_Int negligible_linear_singular_image \linear g\) qed have negf: "negligible ((f ` K) \ (f ` L))" and negid: "negligible (K \ L)" if "K \ \" "L \ \" "K \ L" for K L using neg [OF \linear f\] neg [OF linear_id] that by auto show ?thesis proof (cases "finite \") case True then have "?\ (\x\\. f ` x) = (\x\\. ?\ (f ` x))" using \linear f\ cbox measurable_linear_image_interval negf by (blast intro: measure_negligible_finite_Union_image [unfolded pairwise_def]) also have "\ = (\k\\. m * ?\ k)" by (metis (no_types, lifting) cbox im sum.cong) also have "\ = m * ?\ (\\)" unfolding sum_distrib_left [symmetric] by (metis True cbox lmeasurable_cbox measure_negligible_finite_Union [unfolded pairwise_def] negid) finally show ?thesis by (metis True \linear f\ cbox image_Union fmeasurable.finite_UN measurable_linear_image_interval) next case False with \countable \\ obtain X :: "nat \ 'n set" where S: "bij_betw X UNIV \" using bij_betw_from_nat_into by blast then have eq: "(\\) = (\n. X n)" "(f ` \\) = (\n. f ` X n)" by (auto simp: bij_betw_def) have meas: "\K. K \ \ \ K \ lmeasurable" using cbox by blast with S have 1: "\n. X n \ lmeasurable" by (auto simp: bij_betw_def) have 2: "pairwise (\m n. negligible (X m \ X n)) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negid range_eqI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have 3: "bounded (\n. X n)" using S unfolding bij_betw_def by blast have "(\n. X n) \ lmeasurable" by (rule measurable_countable_negligible_Union_bounded [OF 1 2 3]) with S have f1: "\n. f ` (X n) \ lmeasurable" unfolding bij_betw_def by (metis assms(1) cbox measurable_linear_image_interval rangeI) have f2: "pairwise (\m n. negligible (f ` (X m) \ f ` (X n))) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negf rangeI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have f3: "bounded (\n. f ` X n)" using S unfolding bij_betw_def by (metis bounded_linear_image linear_linear assms(1) image_Union range_composition) have "(\n. ?\ (X n)) sums ?\ (\n. X n)" by (rule measure_countable_negligible_Union_bounded [OF 1 2 3]) have meq: "?\ (\n. f ` X n) = m * ?\ (\(X ` UNIV))" proof (rule sums_unique2 [OF measure_countable_negligible_Union_bounded [OF f1 f2 f3]]) have m: "\n. ?\ (f ` X n) = (m * ?\ (X n))" using S unfolding bij_betw_def by (metis cbox im rangeI) show "(\n. ?\ (f ` X n)) sums (m * ?\ (\(X ` UNIV)))" unfolding m using measure_countable_negligible_Union_bounded [OF 1 2 3] sums_mult by blast qed show ?thesis using measurable_countable_negligible_Union_bounded [OF f1 f2 f3] meq by (auto simp: eq [symmetric]) qed qed show ?thesis unfolding completion.fmeasurable_measure_inner_outer_le proof (intro conjI allI impI) fix e :: real assume "e > 0" have 1: "cbox a b - S \ lmeasurable" by (simp add: fmeasurable.Diff that) have 2: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "cbox a b - S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ (cbox a b - S) + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of "cbox a b - S" a b "e/(1 + \m\)"]; use 1 2 pairwise_def in force) show "\T \ lmeasurable. T \ f ` S \ m * ?\ S - e \ ?\ T" proof (intro bexI conjI) show "f ` (cbox a b) - f ` (\\) \ f ` S" using \cbox a b - S \ \\\ by force have "m * ?\ S - e \ m * (?\ S - e / (1 + \m\))" using \m \ 0\ \e > 0\ by (simp add: field_simps) also have "\ \ ?\ (f ` cbox a b) - ?\ (f ` (\\))" proof - have "?\ (cbox a b - S) = ?\ (cbox a b) - ?\ S" by (simp add: measurable_measure_Diff \S \ cbox a b\ fmeasurableD that(2)) then have "(?\ S - e / (1 + m)) \ (content (cbox a b) - ?\ (\ \))" using \m \ 0\ le by auto then show ?thesis using \m \ 0\ \e > 0\ by (simp add: mult_left_mono im fUD [OF \countable \\ cbox intdisj] flip: right_diff_distrib) qed also have "\ = ?\ (f ` cbox a b - f ` \\)" proof (rule measurable_measure_Diff [symmetric]) show "f ` cbox a b \ lmeasurable" by (simp add: assms(1) measurable_linear_image_interval) show "f ` \ \ \ sets lebesgue" by (simp add: \countable \\ cbox fUD fmeasurableD intdisj) show "f ` \ \ \ f ` cbox a b" by (simp add: Sup_le_iff cbox image_mono) qed finally show "m * ?\ S - e \ ?\ (f ` cbox a b - f ` \\)" . show "f ` cbox a b - f ` \\ \ lmeasurable" by (simp add: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) qed next fix e :: real assume "e > 0" have em: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ S + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of S a b "e/(1 + \m\)"]; use \S \ lmeasurable\ \S \ cbox a b\ em in force) show "\U \ lmeasurable. f ` S \ U \ ?\ U \ m * ?\ S + e" proof (intro bexI conjI) show "f ` S \ f ` (\\)" by (simp add: DD(1) image_mono) have "?\ (f ` \\) \ m * (?\ S + e / (1 + \m\))" using \m \ 0\ le mult_left_mono by (auto simp: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) also have "\ \ m * ?\ S + e" using \m \ 0\ \e > 0\ by (simp add: fUD [OF \countable \\ cbox intdisj] field_simps) finally show "?\ (f ` \\) \ m * ?\ S + e" . show "f ` \\ \ lmeasurable" by (simp add: \countable \\ cbox fUD intdisj) qed qed qed show ?thesis unfolding has_measure_limit_iff proof (intro allI impI) fix e :: real assume "e > 0" obtain B where "B > 0" and B: "\a b. ball 0 B \ cbox a b \ \?\ (S \ cbox a b) - ?\ S\ < e / (1 + \m\)" using has_measure_limit [OF S] \e > 0\ by (metis abs_add_one_gt_zero zero_less_divide_iff) obtain c d::'n where cd: "ball 0 B \ cbox c d" by (metis bounded_subset_cbox_symmetric bounded_ball) with B have less: "\?\ (S \ cbox c d) - ?\ S\ < e / (1 + \m\)" . obtain D where "D > 0" and D: "cbox c d \ ball 0 D" by (metis bounded_cbox bounded_subset_ballD) obtain C where "C > 0" and C: "\x. norm (f x) \ C * norm x" using linear_bounded_pos \linear f\ by blast have "f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" if "ball 0 (D*C) \ cbox a b" for a b proof - have "bounded (S \ h ` cbox a b)" by (simp add: bounded_linear_image linear_linear \linear h\ bounded_Int) moreover have Shab: "S \ h ` cbox a b \ lmeasurable" by (simp add: S \linear h\ fmeasurable.Int measurable_linear_image_interval) moreover have fim: "f ` (S \ h ` (cbox a b)) = (f ` S) \ cbox a b" by (auto simp: hf rev_image_eqI fh) ultimately have 1: "(f ` S) \ cbox a b \ lmeasurable" and 2: "?\ ((f ` S) \ cbox a b) = m * ?\ (S \ h ` cbox a b)" using fBS [of "S \ (h ` (cbox a b))"] by auto have *: "\\z - m\ < e; z \ w; w \ m\ \ \w - m\ \ e" for w z m and e::real by auto have meas_adiff: "\?\ (S \ h ` cbox a b) - ?\ S\ \ e / (1 + \m\)" proof (rule * [OF less]) show "?\ (S \ cbox c d) \ ?\ (S \ h ` cbox a b)" proof (rule measure_mono_fmeasurable [OF _ _ Shab]) have "f ` ball 0 D \ ball 0 (C * D)" using C \C > 0\ apply (clarsimp simp: algebra_simps) by (meson le_less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono) then have "f ` ball 0 D \ cbox a b" by (metis mult.commute order_trans that) have "ball 0 D \ h ` cbox a b" by (metis \f ` ball 0 D \ cbox a b\ hf image_subset_iff subsetI) then show "S \ cbox c d \ S \ h ` cbox a b" using D by blast next show "S \ cbox c d \ sets lebesgue" using S fmeasurable_cbox by blast qed next show "?\ (S \ h ` cbox a b) \ ?\ S" by (simp add: S Shab fmeasurableD measure_mono_fmeasurable) qed have "\?\ (f ` S \ cbox a b) - m * ?\ S\ \ \?\ S - ?\ (S \ h ` cbox a b)\ * m" by (metis "2" \m \ 0\ abs_minus_commute abs_mult_pos mult.commute order_refl right_diff_distrib') also have "\ \ e / (1 + m) * m" by (metis \m \ 0\ abs_minus_commute abs_of_nonneg meas_adiff mult.commute mult_left_mono) also have "\ < e" using \e > 0\ \m \ 0\ by (simp add: field_simps) finally have "\?\ (f ` S \ cbox a b) - m * ?\ S\ < e" . with 1 show ?thesis by auto qed then show "\B>0. \a b. ball 0 B \ cbox a b \ f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" - using \C>0\ \D>0\ by (metis mult_zero_left real_mult_less_iff1) + using \C>0\ \D>0\ by (metis mult_zero_left mult_less_iff1) qed qed qed subsection\Lemmas about absolute integrability\ lemma absolutely_integrable_linear: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and h :: "'n::euclidean_space \ 'p::euclidean_space" shows "f absolutely_integrable_on s \ bounded_linear h \ (h \ f) absolutely_integrable_on s" using integrable_bounded_linear[of h lebesgue "\x. indicator s x *\<^sub>R f x"] by (simp add: linear_simps[of h] set_integrable_def) lemma absolutely_integrable_sum: fixes f :: "'a \ 'n::euclidean_space \ 'm::euclidean_space" assumes "finite T" and "\a. a \ T \ (f a) absolutely_integrable_on S" shows "(\x. sum (\a. f a x) T) absolutely_integrable_on S" using assms by induction auto lemma absolutely_integrable_integrable_bound: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes le: "\x. x\S \ norm (f x) \ g x" and f: "f integrable_on S" and g: "g integrable_on S" shows "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) have "g absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (g x)) integrable_on S" using le norm_ge_zero[of "f _"] by (intro integrable_spike_finite[OF _ _ g, of "{}"]) (auto intro!: abs_of_nonneg intro: order_trans simp del: norm_ge_zero) qed fact then show "integrable lebesgue (\x. indicat_real S x *\<^sub>R g x)" by (simp add: set_integrable_def) show "(\x. indicat_real S x *\<^sub>R f x) \ borel_measurable lebesgue" using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) qed (use le in \force intro!: always_eventually split: split_indicator\) corollary absolutely_integrable_on_const [simp]: fixes c :: "'a::euclidean_space" assumes "S \ lmeasurable" shows "(\x. c) absolutely_integrable_on S" by (metis (full_types) assms absolutely_integrable_integrable_bound integrable_on_const order_refl) lemma absolutely_integrable_continuous: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "continuous_on (cbox a b) f \ f absolutely_integrable_on cbox a b" using absolutely_integrable_integrable_bound by (simp add: absolutely_integrable_on_def continuous_on_norm integrable_continuous) lemma absolutely_integrable_continuous_real: fixes f :: "real \ 'b::euclidean_space" shows "continuous_on {a..b} f \ f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_continuous box_real(2)) lemma continuous_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on (cbox a b) f" shows "integrable (lebesgue_on (cbox a b)) f" proof - have "f absolutely_integrable_on cbox a b" by (simp add: absolutely_integrable_continuous assms) then show ?thesis by (simp add: integrable_restrict_space set_integrable_def) qed lemma continuous_imp_integrable_real: fixes f :: "real \ 'b::euclidean_space" assumes "continuous_on {a..b} f" shows "integrable (lebesgue_on {a..b}) f" by (metis assms continuous_imp_integrable interval_cbox) subsection \Componentwise\ proposition absolutely_integrable_componentwise_iff: shows "f absolutely_integrable_on A \ (\b\Basis. (\x. f x \ b) absolutely_integrable_on A)" proof - have *: "(\x. norm (f x)) integrable_on A \ (\b\Basis. (\x. norm (f x \ b)) integrable_on A)" (is "?lhs = ?rhs") if "f integrable_on A" proof assume ?lhs then show ?rhs by (metis absolutely_integrable_on_def Topology_Euclidean_Space.norm_nth_le absolutely_integrable_integrable_bound integrable_component that) next assume R: ?rhs have "f absolutely_integrable_on A" proof (rule absolutely_integrable_integrable_bound) show "(\x. \i\Basis. norm (f x \ i)) integrable_on A" using R by (force intro: integrable_sum) qed (use that norm_le_l1 in auto) then show ?lhs using absolutely_integrable_on_def by auto qed show ?thesis unfolding absolutely_integrable_on_def by (simp add: integrable_componentwise_iff [symmetric] ball_conj_distrib * cong: conj_cong) qed lemma absolutely_integrable_componentwise: shows "(\b. b \ Basis \ (\x. f x \ b) absolutely_integrable_on A) \ f absolutely_integrable_on A" using absolutely_integrable_componentwise_iff by blast lemma absolutely_integrable_component: "f absolutely_integrable_on A \ (\x. f x \ (b :: 'b :: euclidean_space)) absolutely_integrable_on A" by (drule absolutely_integrable_linear[OF _ bounded_linear_inner_left[of b]]) (simp add: o_def) lemma absolutely_integrable_scaleR_left: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S" proof - have "(\x. c *\<^sub>R x) o f absolutely_integrable_on S" by (simp add: absolutely_integrable_linear assms bounded_linear_scaleR_right) then show ?thesis using assms by blast qed lemma absolutely_integrable_scaleR_right: assumes "f absolutely_integrable_on S" shows "(\x. f x *\<^sub>R c) absolutely_integrable_on S" using assms by blast lemma absolutely_integrable_norm: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(norm o f) absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def o_def) lemma absolutely_integrable_abs: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) absolutely_integrable_on S" (is "?g absolutely_integrable_on S") proof - have *: "(\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" if "i \ Basis" for i proof - have "bounded_linear (\y. \j\Basis. if j = i then y *\<^sub>R j else 0)" by (simp add: linear_linear algebra_simps linearI) moreover have "(\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" using assms \i \ Basis\ unfolding o_def by (intro absolutely_integrable_norm [unfolded o_def]) (auto simp: algebra_simps dest: absolutely_integrable_component) ultimately show ?thesis by (subst comp_assoc) (blast intro: absolutely_integrable_linear) qed have eq: "?g = (\x. \i\Basis. ((\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm(\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f) x)" by (simp) show ?thesis unfolding eq by (rule absolutely_integrable_sum) (force simp: intro!: *)+ qed lemma abs_absolutely_integrableI_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "(\x. \f x\) integrable_on A" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_integrable_bound [OF _ assms]) auto lemma abs_absolutely_integrableI: assumes f: "f integrable_on S" and fcomp: "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" shows "f absolutely_integrable_on S" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" if "i \ Basis" for i proof - have "(\x. \f x \ i\) integrable_on S" using assms integrable_component [OF fcomp, where y=i] that by simp then have "(\x. f x \ i) absolutely_integrable_on S" using abs_absolutely_integrableI_1 f integrable_component by blast then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_abs_iff: "f absolutely_integrable_on S \ f integrable_on S \ (\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using absolutely_integrable_abs absolutely_integrable_on_def by blast next assume ?rhs moreover have "(\x. if x \ S then \i\Basis. \f x \ i\ *\<^sub>R i else 0) = (\x. \i\Basis. \(if x \ S then f x else 0) \ i\ *\<^sub>R i)" by force ultimately show ?lhs by (simp only: absolutely_integrable_restrict_UNIV [of S, symmetric] integrable_restrict_UNIV [of S, symmetric] abs_absolutely_integrableI) qed lemma absolutely_integrable_max: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i + \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i + \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib algebra_simps euclidean_representation) finally show "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] by (intro set_integral_add absolutely_integrable_scaleR_left assms) (simp add: algebra_simps) ultimately show ?thesis by metis qed corollary absolutely_integrable_max_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. max (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_max [OF assms] by simp lemma absolutely_integrable_min: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i - \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i - \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib sum_subtractf algebra_simps euclidean_representation) finally show "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] by (intro set_integral_add set_integral_diff absolutely_integrable_scaleR_left assms) (simp add: algebra_simps) ultimately show ?thesis by metis qed corollary absolutely_integrable_min_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. min (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_min [OF assms] by simp lemma nonnegative_absolutely_integrable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ 0 \ f x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" if "i \ Basis" for i proof - have "(\x. f x \ i) integrable_on A" by (simp add: assms(1) integrable_component) then have "(\x. f x \ i) absolutely_integrable_on A" by (metis that comp nonnegative_absolutely_integrable_1) then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_component_ubound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f integrable_on A" and g: "g absolutely_integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. g x - (g x - f x)) absolutely_integrable_on A" proof (rule set_integral_diff [OF g nonnegative_absolutely_integrable]) show "(\x. g x - f x) integrable_on A" using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast qed (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma absolutely_integrable_component_lbound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f absolutely_integrable_on A" and g: "g integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "g absolutely_integrable_on A" proof - have "(\x. f x + (g x - f x)) absolutely_integrable_on A" proof (rule set_integral_add [OF f nonnegative_absolutely_integrable]) show "(\x. g x - f x) integrable_on A" using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g by blast qed (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f integrable_on S \ (\x. f x $ 1) integrable_on S" by (auto simp: integrable_componentwise_iff [of f] Basis_vec_def cart_eq_inner_axis) lemma integral_on_1_eq: fixes f :: "'a::euclidean_space \ real^1" shows "integral S f = vec (integral S (\x. f x $ 1))" by (cases "f integrable_on S") (simp_all add: integrable_on_1_iff vec_eq_iff not_integrable_integral) lemma absolutely_integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f absolutely_integrable_on S \ (\x. f x $ 1) absolutely_integrable_on S" unfolding absolutely_integrable_on_def by (auto simp: integrable_on_1_iff norm_real) lemma absolutely_integrable_absolutely_integrable_lbound: fixes f :: "'m::euclidean_space \ real" assumes f: "f integrable_on S" and g: "g absolutely_integrable_on S" and *: "\x. x \ S \ g x \ f x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_lbound [OF g f]) (simp add: *) lemma absolutely_integrable_absolutely_integrable_ubound: fixes f :: "'m::euclidean_space \ real" assumes fg: "f integrable_on S" "g absolutely_integrable_on S" and *: "\x. x \ S \ f x \ g x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_ubound [OF fg]) (simp add: *) lemma has_integral_vec1_I_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) (cbox a b)" shows "((f \ vec) has_integral y) {a$1..b$1}" proof - have "((\x. f(vec x)) has_integral (1 / 1) *\<^sub>R y) ((\x. x $ 1) ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real^1. vec ` cbox u v = cbox w z" "content (vec ` cbox u v :: (real^1) set) = 1 * content (cbox u v)" for u v unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z. (\x. x $ 1) ` cbox u v = cbox w z" for u v :: "real^1" using vec_nth_cbox_1_eq by blast qed (auto simp: continuous_vec assms) then show ?thesis by (simp add: o_def) qed lemma has_integral_vec1_I: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) S" shows "(f \ vec has_integral y) ((\x. x $ 1) ` S)" proof - have *: "\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e" if int: "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)" and B: "ball 0 B \ {a..b}" for e B a b proof - have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force have B': "ball (0::real^1) B \ cbox (vec a) (vec b)" using B by (simp add: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box norm_real subset_iff) show ?thesis using int [OF B'] by (auto simp: image_iff o_def cong: if_cong dest!: has_integral_vec1_I_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_I_cbox split: if_split_asm) subgoal by (metis vector_one_nth) subgoal apply (erule all_forward imp_forward ex_forward asm_rl)+ by (blast intro!: *)+ done qed lemma has_integral_vec1_nth_cbox: fixes f :: "real \ 'a::real_normed_vector" assumes "(f has_integral y) {a..b}" shows "((\x::real^1. f(x$1)) has_integral y) (cbox (vec a) (vec b))" proof - have "((\x::real^1. f(x$1)) has_integral (1 / 1) *\<^sub>R y) (vec ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real. (\x. x $ 1) ` cbox u v = cbox w z" "content ((\x. x $ 1) ` cbox u v) = 1 * content (cbox u v)" for u v::"real^1" unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z::real^1. vec ` cbox u v = cbox w z" for u v :: "real" using vec_cbox_1_eq by auto qed (auto simp: continuous_vec assms) then show ?thesis using vec_cbox_1_eq by auto qed lemma has_integral_vec1_D_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) {a$1..b$1}" shows "(f has_integral y) (cbox a b)" by (metis (mono_tags, lifting) assms comp_apply has_integral_eq has_integral_vec1_nth_cbox vector_one_nth) lemma has_integral_vec1_D: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) ((\x. x $ 1) ` S)" shows "(f has_integral y) S" proof - have *: "\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e" if int: "\a b. ball 0 B \ {a..b} \ (\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e)" and B: "ball 0 B \ cbox a b" for e B and a b::"real^1" proof - have B': "ball 0 B \ {a$1..b$1}" proof (clarsimp) fix t assume "\t\ < B" then show "a $ 1 \ t \ t \ b $ 1" using subsetD [OF B] by (metis (mono_tags, hide_lams) mem_ball_0 mem_box_cart(2) norm_real vec_component) qed have eq: "(\x. if vec x \ S then f (vec x) else 0) = (\x. if x \ S then f x else 0) \ vec" by force have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force show ?thesis using int [OF B'] by (auto simp: image_iff eq cong: if_cong dest!: has_integral_vec1_D_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_D_cbox eq_cbox split: if_split_asm, blast) apply (intro conjI impI) subgoal by (metis vector_one_nth) apply (erule thin_rl) apply (erule all_forward ex_forward conj_forward)+ by (blast intro!: *)+ qed lemma integral_vec1_eq: fixes f :: "real^1 \ 'a::real_normed_vector" shows "integral S f = integral ((\x. x $ 1) ` S) (f \ vec)" using has_integral_vec1_I [of f] has_integral_vec1_D [of f] by (metis has_integral_iff not_integrable_integral) lemma absolutely_integrable_drop: fixes f :: "real^1 \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ (f \ vec) absolutely_integrable_on (\x. x $ 1) ` S" unfolding absolutely_integrable_on_def integrable_on_def proof safe fix y r assume "(f has_integral y) S" "((\x. norm (f x)) has_integral r) S" then show "\y. (f \ vec has_integral y) ((\x. x $ 1) ` S)" "\y. ((\x. norm ((f \ vec) x)) has_integral y) ((\x. x $ 1) ` S)" by (force simp: o_def dest!: has_integral_vec1_I)+ next fix y :: "'b" and r :: "real" assume "(f \ vec has_integral y) ((\x. x $ 1) ` S)" "((\x. norm ((f \ vec) x)) has_integral r) ((\x. x $ 1) ` S)" then show "\y. (f has_integral y) S" "\y. ((\x. norm (f x)) has_integral y) S" by (force simp: o_def intro: has_integral_vec1_D)+ qed subsection \Dominated convergence\ lemma dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. (f k) integrable_on S" and h: "h integrable_on S" and le: "\k x. x \ S \ norm (f k x) \ h x" and conv: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" "(\k. integral S (f k)) \ integral S g" proof - have 3: "h absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (h x)) integrable_on S" proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI) fix x assume "x \ S - {}" then show "norm (h x) = h x" by (metis Diff_empty abs_of_nonneg bot_set_def le norm_ge_zero order_trans real_norm_def) qed auto qed fact have 2: "set_borel_measurable lebesgue S (f k)" for k unfolding set_borel_measurable_def using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) then have 1: "set_borel_measurable lebesgue S g" unfolding set_borel_measurable_def by (rule borel_measurable_LIMSEQ_metric) (use conv in \auto split: split_indicator\) have 4: "AE x in lebesgue. (\i. indicator S x *\<^sub>R f i x) \ indicator S x *\<^sub>R g x" "AE x in lebesgue. norm (indicator S x *\<^sub>R f k x) \ indicator S x *\<^sub>R h x" for k using conv le by (auto intro!: always_eventually split: split_indicator) have g: "g absolutely_integrable_on S" using 1 2 3 4 unfolding set_borel_measurable_def set_integrable_def by (rule integrable_dominated_convergence) then show "g integrable_on S" by (auto simp: absolutely_integrable_on_def) have "(\k. (LINT x:S|lebesgue. f k x)) \ (LINT x:S|lebesgue. g x)" unfolding set_borel_measurable_def set_lebesgue_integral_def using 1 2 3 4 unfolding set_borel_measurable_def set_lebesgue_integral_def set_integrable_def by (rule integral_dominated_convergence) then show "(\k. integral S (f k)) \ integral S g" using g absolutely_integrable_integrable_bound[OF le f h] by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto qed lemma has_integral_dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes "\k. (f k has_integral y k) S" "h integrable_on S" "\k. \x\S. norm (f k x) \ h x" "\x\S. (\k. f k x) \ g x" and x: "y \ x" shows "(g has_integral x) S" proof - have int_f: "\k. (f k) integrable_on S" using assms by (auto simp: integrable_on_def) have "(g has_integral (integral S g)) S" by (metis assms(2-4) dominated_convergence(1) has_integral_integral int_f) moreover have "integral S g = x" proof (rule LIMSEQ_unique) show "(\i. integral S (f i)) \ x" using integral_unique[OF assms(1)] x by simp show "(\i. integral S (f i)) \ integral S g" by (metis assms(2) assms(3) assms(4) dominated_convergence(2) int_f) qed ultimately show ?thesis by simp qed lemma dominated_convergence_integrable_1: fixes f :: "nat \ 'n::euclidean_space \ real" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" proof - have habs: "h absolutely_integrable_on S" using h normg nonnegative_absolutely_integrable_1 norm_ge_zero order_trans by blast let ?f = "\n x. (min (max (- h x) (f n x)) (h x))" have h0: "h x \ 0" if "x \ S" for x using normg that by force have leh: "norm (?f k x) \ h x" if "x \ S" for k x using h0 that by force have limf: "(\k. ?f k x) \ g x" if "x \ S" for x proof - have "\e y. \f y x - g x\ < e \ \min (max (- h x) (f y x)) (h x) - g x\ < e" using h0 [OF that] normg [OF that] by simp then show ?thesis using lim [OF that] by (auto simp add: tendsto_iff dist_norm elim!: eventually_mono) qed show ?thesis proof (rule dominated_convergence [of ?f S h g]) have "(\x. - h x) absolutely_integrable_on S" using habs unfolding set_integrable_def by auto then show "?f k integrable_on S" for k by (intro set_lebesgue_integral_eq_integral absolutely_integrable_min_1 absolutely_integrable_max_1 f habs) qed (use assms leh limf in auto) qed lemma dominated_convergence_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" using f unfolding integrable_componentwise_iff [of g] absolutely_integrable_componentwise_iff [where f = "f k" for k] proof clarify fix b :: "'m" assume fb [rule_format]: "\k. \b\Basis. (\x. f k x \ b) absolutely_integrable_on S" and b: "b \ Basis" show "(\x. g x \ b) integrable_on S" proof (rule dominated_convergence_integrable_1 [OF fb h]) fix x assume "x \ S" show "norm (g x \ b) \ h x" using norm_nth_le \x \ S\ b normg order.trans by blast show "(\k. f k x \ b) \ g x \ b" using \x \ S\ b lim tendsto_componentwise_iff by fastforce qed (use b in auto) qed lemma dominated_convergence_absolutely_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g absolutely_integrable_on S" proof - have "g integrable_on S" by (rule dominated_convergence_integrable [OF assms]) with assms show ?thesis by (blast intro: absolutely_integrable_integrable_bound [where g=h]) qed proposition integral_countable_UN: fixes f :: "real^'m \ real^'n" assumes f: "f absolutely_integrable_on (\(range s))" and s: "\m. s m \ sets lebesgue" shows "\n. f absolutely_integrable_on (\m\n. s m)" and "(\n. integral (\m\n. s m) f) \ integral (\(s ` UNIV)) f" (is "?F \ ?I") proof - show fU: "f absolutely_integrable_on (\m\n. s m)" for n using assms by (blast intro: set_integrable_subset [OF f]) have fint: "f integrable_on (\ (range s))" using absolutely_integrable_on_def f by blast let ?h = "\x. if x \ \(s ` UNIV) then norm(f x) else 0" have "(\n. integral UNIV (\x. if x \ (\m\n. s m) then f x else 0)) \ integral UNIV (\x. if x \ \(s ` UNIV) then f x else 0)" proof (rule dominated_convergence) show "(\x. if x \ (\m\n. s m) then f x else 0) integrable_on UNIV" for n unfolding integrable_restrict_UNIV using fU absolutely_integrable_on_def by blast show "(\x. if x \ \(s ` UNIV) then norm(f x) else 0) integrable_on UNIV" by (metis (no_types) absolutely_integrable_on_def f integrable_restrict_UNIV) show "\x. (\n. if x \ (\m\n. s m) then f x else 0) \ (if x \ \(s ` UNIV) then f x else 0)" by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto then show "?F \ ?I" by (simp only: integral_restrict_UNIV) qed subsection \Fundamental Theorem of Calculus for the Lebesgue integral\ text \ For the positive integral we replace continuity with Borel-measurability. \ lemma fixes f :: "real \ real" assumes [measurable]: "f \ borel_measurable borel" assumes f: "\x. x \ {a..b} \ DERIV F x :> f x" "\x. x \ {a..b} \ 0 \ f x" and "a \ b" shows nn_integral_FTC_Icc: "(\\<^sup>+x. ennreal (f x) * indicator {a .. b} x \lborel) = F b - F a" (is ?nn) and has_bochner_integral_FTC_Icc_nonneg: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_nonneg: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) and integrable_FTC_Icc_nonneg: "integrable lborel (\x. f x * indicator {a .. b} x)" (is ?int) proof - have *: "(\x. f x * indicator {a..b} x) \ borel_measurable borel" "\x. 0 \ f x * indicator {a..b} x" using f(2) by (auto split: split_indicator) have F_mono: "a \ x \ x \ y \ y \ b\ F x \ F y" for x y using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) have "(f has_integral F b - F a) {a..b}" by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] intro: has_field_derivative_subset[OF f(1)] \a \ b\) then have i: "((\x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV" unfolding indicator_def if_distrib[where f="\x. a * x" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) then have nn: "(\\<^sup>+x. f x * indicator {a .. b} x \lborel) = F b - F a" by (rule nn_integral_has_integral_lborel[OF *]) then show ?has by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \a \ b\) then show ?eq ?int unfolding has_bochner_integral_iff by auto show ?nn by (subst nn[symmetric]) (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator) qed lemma fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" assumes "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" assumes cont: "continuous_on {a .. b} f" shows has_bochner_integral_FTC_Icc: "has_bochner_integral lborel (\x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has) and integral_FTC_Icc: "(\x. indicator {a .. b} x *\<^sub>R f x \lborel) = F b - F a" (is ?eq) proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have int: "integrable lborel ?f" using borel_integrable_compact[OF _ cont] by auto have "(f has_integral F b - F a) {a..b}" using assms(1,2) by (intro fundamental_theorem_of_calculus) auto moreover have "(f has_integral integral\<^sup>L lborel ?f) {a..b}" using has_integral_integral_lborel[OF int] unfolding indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) ultimately show ?eq by (auto dest: has_integral_unique) then show ?has using int by (auto simp: has_bochner_integral_iff) qed lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ DERIV F x :> f x" assumes cont: "\x. a \ x \ x \ b \ isCont f x" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" unfolding has_field_derivative_iff_has_vector_derivative[symmetric] using deriv by (auto intro: DERIV_subset) have 2: "continuous_on {a .. b} f" using cont by (intro continuous_at_imp_continuous_on) auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 2] integral_FTC_Icc[OF \a \ b\ 1 2] by (auto simp: mult.commute) qed lemma nn_integral_FTC_atLeast: fixes f :: "real \ real" assumes f_borel: "f \ borel_measurable borel" assumes f: "\x. a \ x \ DERIV F x :> f x" assumes nonneg: "\x. a \ x \ 0 \ f x" assumes lim: "(F \ T) at_top" shows "(\\<^sup>+x. ennreal (f x) * indicator {a ..} x \lborel) = T - F a" proof - let ?f = "\(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x" let ?fR = "\x. ennreal (f x) * indicator {a ..} x" have F_mono: "a \ x \ x \ y \ F x \ F y" for x y using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) then have F_le_T: "a \ x \ F x \ T" for x by (intro tendsto_lowerbound[OF lim]) (auto simp: eventually_at_top_linorder) have "(SUP i. ?f i x) = ?fR x" for x proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) obtain n where "x - a < real n" using reals_Archimedean2[of "x - a"] .. then have "eventually (\n. ?f n x = ?fR x) sequentially" by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) then show "(\n. ?f n x) \ ?fR x" by (rule tendsto_eventually) qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator) then have "integral\<^sup>N lborel ?fR = (\\<^sup>+ x. (SUP i. ?f i x) \lborel)" by simp also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \lborel))" proof (rule nn_integral_monotone_convergence_SUP) show "incseq ?f" using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) show "\i. (?f i) \ borel_measurable lborel" using f_borel by auto qed also have "\ = (SUP i. ennreal (F (a + real i) - F a))" by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto also have "\ = T - F a" proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) have "(\x. F (a + real x)) \ T" by (auto intro: filterlim_compose[OF lim filterlim_tendsto_add_at_top] filterlim_real_sequentially) then show "(\n. ennreal (F (a + real n) - F a)) \ ennreal (T - F a)" by (simp add: F_mono F_le_T tendsto_diff) qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono) finally show ?thesis . qed lemma integral_power: "a \ b \ (\x. x^k * indicator {a..b} x \lborel) = (b^Suc k - a^Suc k) / Suc k" proof (subst integral_FTC_Icc_real) fix x show "DERIV (\x. x^Suc k / Suc k) x :> x^k" by (intro derivative_eq_intros) auto qed (auto simp: field_simps simp del: of_nat_Suc) subsection \Integration by parts\ lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont) lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = LBINT x. F x * g x * indicat_real {a..b} x + f x * G x * indicat_real {a..b} x" by (meson vector_space_over_itself.scale_left_distrib) also have "... = (\x. (F x * g x) * indicator {a .. b} x \lborel) + \x. (f x * G x) * indicator {a .. b} x \lborel" proof (intro Bochner_Integration.integral_add borel_integrable_atLeastAtMost cont_f cont_g continuous_intros) show "\x. \a \ x; x \ b\ \ isCont F x" "\x. \a \ x; x \ b\ \ isCont G x" using DERIV_isCont by blast+ qed finally have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = (\x. (F x * g x) * indicator {a .. b} x \lborel) + \x. (f x * G x) * indicator {a .. b} x \lborel" . moreover have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" proof (intro integral_FTC_Icc_real derivative_eq_intros cont_f cont_g continuous_intros) show "\x. \a \ x; x \ b\ \ isCont F x" "\x. \a \ x; x \ b\ \ isCont G x" using DERIV_isCont by blast+ qed auto ultimately show ?thesis by auto qed lemma integral_by_parts': fixes f g F G::"real \ real" assumes "a \ b" assumes "!!x. a \x \ x\b \ isCont f x" assumes "!!x. a \x \ x\b \ isCont g x" assumes "!!x. DERIV F x :> f x" assumes "!!x. DERIV G x :> g x" shows "(\x. indicator {a .. b} x *\<^sub>R (F x * g x) \lborel) = F b * G b - F a * G a - \x. indicator {a .. b} x *\<^sub>R (f x * G x) \lborel" using integral_by_parts[OF assms] by (simp add: ac_simps) lemma has_bochner_integral_even_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes even: "\x. f (- x) = f x" shows "has_bochner_integral lborel f (2 *\<^sub>R x)" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator even f) with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by (simp add: scaleR_2) qed lemma has_bochner_integral_odd_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes odd: "\x. f (- x) = - f x" shows "has_bochner_integral lborel f 0" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. - indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator odd f) from has_bochner_integral_minus[OF this] have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) (- x)" by simp with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + - x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by simp qed lemma has_integral_0_closure_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (closure S) f" and nonneg_interior: "\x. x \ S \ 0 \ f x" and pos: "0 < emeasure lborel S" and finite: "emeasure lborel S < \" and regular: "emeasure lborel (closure S) = emeasure lborel S" and opn: "open S" assumes int: "(f has_integral 0) (closure S)" assumes x: "x \ closure S" shows "f x = 0" proof - have zero: "emeasure lborel (frontier S) = 0" using finite closure_subset regular unfolding frontier_def by (subst emeasure_Diff) (auto simp: frontier_def interior_open \open S\ ) have nonneg: "0 \ f x" if "x \ closure S" for x using continuous_ge_on_closure[OF f that nonneg_interior] by simp have "0 = integral (closure S) f" by (blast intro: int sym) also note intl = has_integral_integrable[OF int] have af: "f absolutely_integrable_on (closure S)" using nonneg by (intro absolutely_integrable_onI intl integrable_eq[OF intl]) simp then have "integral (closure S) f = set_lebesgue_integral lebesgue (closure S) f" by (intro set_lebesgue_integral_eq_integral(2)[symmetric]) also have "\ = 0 \ (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)" unfolding set_lebesgue_integral_def proof (rule integral_nonneg_eq_0_iff_AE) show "integrable lebesgue (\x. indicat_real (closure S) x *\<^sub>R f x)" by (metis af set_integrable_def) qed (use nonneg in \auto simp: indicator_def\) also have "\ \ (AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by (auto simp: indicator_def) finally have "(AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by simp moreover have "(AE x in lebesgue. x \ - frontier S)" using zero by (auto simp: eventually_ae_filter null_sets_def intro!: exI[where x="frontier S"]) ultimately have ae: "AE x \ S in lebesgue. x \ {x \ closure S. f x = 0}" (is ?th) by eventually_elim (use closure_subset in \auto simp: \) have "closed {0::real}" by simp with continuous_on_closed_vimage[OF closed_closure, of S f] f have "closed (f -` {0} \ closure S)" by blast then have "closed {x \ closure S. f x = 0}" by (auto simp: vimage_def Int_def conj_commute) with \open S\ have "x \ {x \ closure S. f x = 0}" if "x \ S" for x using ae that by (rule mem_closed_if_AE_lebesgue_open) then have "f x = 0" if "x \ S" for x using that by auto from continuous_constant_on_closure[OF f this \x \ closure S\] show "f x = 0" . qed lemma has_integral_0_cbox_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (cbox a b) f" and nonneg_interior: "\x. x \ box a b \ 0 \ f x" assumes int: "(f has_integral 0) (cbox a b)" assumes ne: "box a b \ {}" assumes x: "x \ cbox a b" shows "f x = 0" proof - have "0 < emeasure lborel (box a b)" using ne x unfolding emeasure_lborel_box_eq by (force intro!: prod_pos simp: mem_box algebra_simps) then show ?thesis using assms by (intro has_integral_0_closure_imp_0[of "box a b" f x]) (auto simp: emeasure_lborel_box_eq emeasure_lborel_cbox_eq algebra_simps mem_box) qed subsection\Various common equivalent forms of function measurability\ lemma indicator_sum_eq: fixes m::real and f :: "'a \ real" assumes "\m\ \ 2 ^ (2*n)" "m/2^n \ f x" "f x < (m+1)/2^n" "m \ \" shows "(\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) = m/2^n" (is "sum ?f ?S = _") proof - have "sum ?f ?S = sum (\k. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {m}" proof (rule comm_monoid_add_class.sum.mono_neutral_right) show "finite ?S" by (rule finite_abs_int_segment) show "{m} \ {k \ \. \k\ \ 2 ^ (2*n)}" using assms by auto show "\i\{k \ \. \k\ \ 2 ^ (2*n)} - {m}. ?f i = 0" using assms by (auto simp: indicator_def Ints_def abs_le_iff field_split_simps) qed also have "\ = m/2^n" using assms by (auto simp: indicator_def not_less) finally show ?thesis . qed lemma measurable_on_sf_limit_lemma1: fixes f :: "'a::euclidean_space \ real" assumes "\a b. {x \ S. a \ f x \ f x < b} \ sets (lebesgue_on S)" obtains g where "\n. g n \ borel_measurable (lebesgue_on S)" "\n. finite(range (g n))" "\x. (\n. g n x) \ f x" proof show "(\x. sum (\k::real. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {k \ \. \k\ \ 2 ^ (2*n)}) \ borel_measurable (lebesgue_on S)" (is "?g \ _") for n proof - have "\k. \k \ \; \k\ \ 2 ^ (2*n)\ \ Measurable.pred (lebesgue_on S) (\x. k / (2^n) \ f x \ f x < (k+1) / (2^n))" using assms by (force simp: pred_def space_restrict_space) then show ?thesis by (simp add: field_class.field_divide_inverse) qed show "finite (range (?g n))" for n proof - have "range (?g n) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof clarify fix x show "?g n x \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof (cases "\k::real. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ (f x) \ (f x) < (k+1)/2^n") case True then show ?thesis apply clarify by (subst indicator_sum_eq) auto next case False then have "?g n x = 0" by auto then show ?thesis by force qed qed moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (simp add: finite_abs_int_segment) ultimately show ?thesis using finite_subset by blast qed show "(\n. ?g n x) \ f x" for x proof (rule LIMSEQ_I) fix e::real assume "e > 0" obtain N1 where N1: "\f x\ < 2 ^ N1" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by force have "norm (?g n x - f x) < e" if n: "n \ max N1 N2" for n proof - define m where "m \ floor(2^n * (f x))" have "1 \ \2^n\ * e" using n N2 \e > 0\ less_eq_real_def less_le_trans by (fastforce simp add: field_split_simps) then have *: "\x \ y; y < x + 1\ \ abs(x - y) < \2^n\ * e" for x y::real by linarith have "\2^n\ * \m/2^n - f x\ = \2^n * (m/2^n - f x)\" by (simp add: abs_mult) also have "\ = \real_of_int \2^n * f x\ - f x * 2^n\" by (simp add: algebra_simps m_def) also have "\ < \2^n\ * e" by (rule *; simp add: mult.commute) finally have "\2^n\ * \m/2^n - f x\ < \2^n\ * e" . then have me: "\m/2^n - f x\ < e" by simp have "\real_of_int m\ \ 2 ^ (2*n)" proof (cases "f x < 0") case True then have "-\f x\ \ \(2::real) ^ N1\" using N1 le_floor_iff minus_le_iff by fastforce with n True have "\real_of_int\f x\\ \ 2 ^ N1" by linarith also have "\ \ 2^n" using n by (simp add: m_def) finally have "\real_of_int \f x\\ * 2^n \ 2^n * 2^n" by simp moreover have "\real_of_int \2^n * f x\\ \ \real_of_int \f x\\ * 2^n" proof - have "\real_of_int \2^n * f x\\ = - (real_of_int \2^n * f x\)" using True by (simp add: abs_if mult_less_0_iff) also have "\ \ - (real_of_int (\(2::real) ^ n\ * \f x\))" using le_mult_floor_Ints [of "(2::real)^n"] by simp also have "\ \ \real_of_int \f x\\ * 2^n" using True by simp finally show ?thesis . qed ultimately show ?thesis by (metis (no_types, hide_lams) m_def order_trans power2_eq_square power_even_eq) next case False with n N1 have "f x \ 2^n" by (simp add: not_less) (meson less_eq_real_def one_le_numeral order_trans power_increasing) moreover have "0 \ m" using False m_def by force ultimately show ?thesis - by (metis abs_of_nonneg floor_mono le_floor_iff m_def of_int_0_le_iff power2_eq_square power_mult real_mult_le_cancel_iff1 zero_less_numeral mult.commute zero_less_power) + by (metis abs_of_nonneg floor_mono le_floor_iff m_def of_int_0_le_iff power2_eq_square power_mult mult_le_cancel_iff1 zero_less_numeral mult.commute zero_less_power) qed then have "?g n x = m/2^n" by (rule indicator_sum_eq) (auto simp add: m_def field_split_simps, linarith) then have "norm (?g n x - f x) = norm (m/2^n - f x)" by simp also have "\ < e" by (simp add: me) finally show ?thesis . qed then show "\no. \n\no. norm (?g n x - f x) < e" by blast qed qed lemma borel_measurable_simple_function_limit: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x))" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x)" if f: "\a i. i \ Basis \ {x \ S. f x \ i < a} \ sets (lebesgue_on S)" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite(image (g n) UNIV)) \ (\x. ((\n. g n x) \ f x \ i))" if "i \ Basis" for i proof (rule measurable_on_sf_limit_lemma1 [of S "\x. f x \ i"]) show "{x \ S. a \ f x \ i \ f x \ i < b} \ sets (lebesgue_on S)" for a b proof - have "{x \ S. a \ f x \ i \ f x \ i < b} = {x \ S. f x \ i < b} - {x \ S. a > f x \ i}" by auto also have "\ \ sets (lebesgue_on S)" using f that by blast finally show ?thesis . qed qed blast then obtain g where g: "\i n. i \ Basis \ g i n \ borel_measurable (lebesgue_on S)" "\i n. i \ Basis \ finite(range (g i n))" "\i x. i \ Basis \ ((\n. g i n x) \ f x \ i)" by metis show ?thesis proof (intro conjI allI exI) show "(\x. \i\Basis. g i n x *\<^sub>R i) \ borel_measurable (lebesgue_on S)" for n by (intro borel_measurable_sum borel_measurable_scaleR) (auto intro: g) show "finite (range (\x. \i\Basis. g i n x *\<^sub>R i))" for n proof - have "range (\x. \i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` PiE Basis (\i. range (g i n))" proof clarify fix x show "(\i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` (\\<^sub>E i\Basis. range (g i n))" by (rule_tac x="\i\Basis. g i n x" in image_eqI) auto qed moreover have "finite(PiE Basis (\i. range (g i n)))" by (simp add: g finite_PiE) ultimately show ?thesis by (metis (mono_tags, lifting) finite_surj) qed show "(\n. \i\Basis. g i n x *\<^sub>R i) \ f x" for x proof - have "(\n. \i\Basis. g i n x *\<^sub>R i) \ (\i\Basis. (f x \ i) *\<^sub>R i)" by (auto intro!: tendsto_sum tendsto_scaleR g) moreover have "(\i\Basis. (f x \ i) *\<^sub>R i) = f x" using euclidean_representation by blast ultimately show ?thesis by metis qed qed qed moreover have "f \ borel_measurable (lebesgue_on S)" if meas_g: "\n. g n \ borel_measurable (lebesgue_on S)" and fin: "\n. finite (range (g n))" and to_f: "\x. (\n. g n x) \ f x" for g by (rule borel_measurable_LIMSEQ_metric [OF meas_g to_f]) ultimately show ?thesis using borel_measurable_vimage_halfspace_component_lt by blast qed subsection \Lebesgue sets and continuous images\ proposition lebesgue_regular_inner: assumes "S \ sets lebesgue" obtains K C where "negligible K" "\n::nat. compact(C n)" "S = (\n. C n) \ K" proof - have "\T. closed T \ T \ S \ (S - T) \ lmeasurable \ emeasure lebesgue (S - T) < ennreal ((1/2)^n)" for n using sets_lebesgue_inner_closed assms by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power) then obtain C where clo: "\n. closed (C n)" and subS: "\n. C n \ S" and mea: "\n. (S - C n) \ lmeasurable" and less: "\n. emeasure lebesgue (S - C n) < ennreal ((1/2)^n)" by metis have "\F. (\n::nat. compact(F n)) \ (\n. F n) = C m" for m::nat by (metis clo closed_Union_compact_subsets) then obtain D :: "[nat,nat] \ 'a set" where D: "\m n. compact(D m n)" "\m. (\n. D m n) = C m" by metis let ?C = "from_nat_into (\m. range (D m))" have "countable (\m. range (D m))" by blast have "range (from_nat_into (\m. range (D m))) = (\m. range (D m))" using range_from_nat_into by simp then have CD: "\m n. ?C k = D m n" for k by (metis (mono_tags, lifting) UN_iff rangeE range_eqI) show thesis proof show "negligible (S - (\n. C n))" proof (clarsimp simp: negligible_outer_le) fix e :: "real" assume "e > 0" then obtain n where n: "(1/2)^n < e" using real_arch_pow_inv [of e "1/2"] by auto show "\T. S - (\n. C n) \ T \ T \ lmeasurable \ measure lebesgue T \ e" proof (intro exI conjI) show "S - (\n. C n) \ S - C n" by blast show "S - C n \ lmeasurable" by (simp add: mea) show "measure lebesgue (S - C n) \ e" using less [of n] n by (simp add: emeasure_eq_measure2 less_le mea) qed qed show "compact (?C n)" for n using CD D by metis show "S = (\n. ?C n) \ (S - (\n. C n))" (is "_ = ?rhs") proof show "S \ ?rhs" using D by fastforce show "?rhs \ S" using subS D CD by auto (metis Sup_upper range_eqI subsetCE) qed qed qed lemma sets_lebesgue_continuous_image: assumes T: "T \ sets lebesgue" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" and "T \ S" shows "f ` T \ sets lebesgue" proof - obtain K C where "negligible K" and com: "\n::nat. compact(C n)" and Teq: "T = (\n. C n) \ K" using lebesgue_regular_inner [OF T] by metis then have comf: "\n::nat. compact(f ` C n)" by (metis Un_subset_iff Union_upper \T \ S\ compact_continuous_image contf continuous_on_subset rangeI) have "((\n. f ` C n) \ f ` K) \ sets lebesgue" proof (rule sets.Un) have "K \ S" using Teq \T \ S\ by blast show "(\n. f ` C n) \ sets lebesgue" proof (rule sets.countable_Union) show "range (\n. f ` C n) \ sets lebesgue" using borel_compact comf by (auto simp: borel_compact) qed auto show "f ` K \ sets lebesgue" by (simp add: \K \ S\ \negligible K\ negim negligible_imp_sets) qed then show ?thesis by (simp add: Teq image_Un image_Union) qed lemma differentiable_image_in_sets_lebesgue: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f`S \ sets lebesgue" proof (rule sets_lebesgue_continuous_image [OF S]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed auto lemma sets_lebesgue_on_continuous_image: assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" shows "f ` X \ sets (lebesgue_on (f ` S))" proof - have "X \ S" by (metis S X sets.Int_space_eq2 sets_restrict_space_iff) moreover have "f ` S \ sets lebesgue" using S contf negim sets_lebesgue_continuous_image by blast moreover have "f ` X \ sets lebesgue" by (metis S X contf negim sets_lebesgue_continuous_image sets_restrict_space_iff space_restrict_space space_restrict_space2) ultimately show ?thesis by (auto simp: sets_restrict_space_iff) qed lemma differentiable_image_in_sets_lebesgue_on: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f ` X \ sets (lebesgue_on (f`S))" proof (rule sets_lebesgue_on_continuous_image [OF S X]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed subsection \Affine lemmas\ lemma borel_measurable_affine: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f \ borel_measurable lebesgue" and "c \ 0" shows "(\x. f(t + c *\<^sub>R x)) \ borel_measurable lebesgue" proof - { fix a b have "{x. f x \ cbox a b} \ sets lebesgue" using f cbox_borel lebesgue_measurable_vimage_borel by blast then have "(\x. (x - t) /\<^sub>R c) ` {x. f x \ cbox a b} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "(\x. (x - t) /\<^sub>R c) differentiable_on {x. f x \ cbox a b}" unfolding differentiable_on_def differentiable_def by (rule \c \ 0\ derivative_eq_intros strip exI | simp)+ qed auto moreover have "{x. f(t + c *\<^sub>R x) \ cbox a b} = (\x. (x-t) /\<^sub>R c) ` {x. f x \ cbox a b}" using \c \ 0\ by (auto simp: image_def) ultimately have "{x. f(t + c *\<^sub>R x) \ cbox a b} \ sets lebesgue" by (auto simp: borel_measurable_vimage_closed_interval) } then show ?thesis by (subst lebesgue_on_UNIV_eq [symmetric]; auto simp: borel_measurable_vimage_closed_interval) qed lemma lebesgue_integrable_real_affine: fixes f :: "real \ 'a :: euclidean_space" assumes f: "integrable lebesgue f" and "c \ 0" shows "integrable lebesgue (\x. f(t + c * x))" proof - have "(\x. norm (f x)) \ borel_measurable lebesgue" by (simp add: borel_measurable_integrable f) then show ?thesis using assms borel_measurable_affine [of f c] unfolding integrable_iff_bounded by (subst (asm) nn_integral_real_affine_lebesgue[where c=c and t=t]) (auto simp: ennreal_mult_less_top) qed lemma lebesgue_integrable_real_affine_iff: fixes f :: "real \ 'a :: euclidean_space" shows "c \ 0 \ integrable lebesgue (\x. f(t + c * x)) \ integrable lebesgue f" using lebesgue_integrable_real_affine[of f c t] lebesgue_integrable_real_affine[of "\x. f(t + c * x)" "1/c" "-t/c"] by (auto simp: field_simps) lemma\<^marker>\tag important\ lebesgue_integral_real_affine: fixes f :: "real \ 'a :: euclidean_space" and c :: real assumes c: "c \ 0" shows "(\x. f x \ lebesgue) = \c\ *\<^sub>R (\x. f(t + c * x) \lebesgue)" proof cases have "(\x. t + c * x) \ lebesgue \\<^sub>M lebesgue" using lebesgue_affine_measurable[where c= "\x::real. c"] \c \ 0\ by simp moreover assume "integrable lebesgue f" ultimately show ?thesis by (subst lebesgue_real_affine[OF c, of t]) (auto simp: integral_density integral_distr) next assume "\ integrable lebesgue f" with c show ?thesis by (simp add: lebesgue_integrable_real_affine_iff not_integrable_integral_eq) qed lemma has_bochner_integral_lebesgue_real_affine_iff: fixes i :: "'a :: euclidean_space" shows "c \ 0 \ has_bochner_integral lebesgue f i \ has_bochner_integral lebesgue (\x. f(t + c * x)) (i /\<^sub>R \c\)" unfolding has_bochner_integral_iff lebesgue_integrable_real_affine_iff by (simp_all add: lebesgue_integral_real_affine[symmetric] divideR_right cong: conj_cong) lemma has_bochner_integral_reflect_real_lemma[intro]: fixes f :: "real \ 'a::euclidean_space" assumes "has_bochner_integral (lebesgue_on {a..b}) f i" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i" proof - have eq: "indicat_real {a..b} (- x) *\<^sub>R f(- x) = indicat_real {- b..- a} x *\<^sub>R f(- x)" for x by (auto simp: indicator_def) have i: "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) i" using assms by (auto simp: has_bochner_integral_restrict_space) then have "has_bochner_integral lebesgue (\x. indicator {-b..-a} x *\<^sub>R f(-x)) i" using has_bochner_integral_lebesgue_real_affine_iff [of "-1" "(\x. indicator {a..b} x *\<^sub>R f x)" i 0] by (auto simp: eq) then show ?thesis by (auto simp: has_bochner_integral_restrict_space) qed lemma has_bochner_integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i \ has_bochner_integral (lebesgue_on {a..b}) f i" by (auto simp: dest: has_bochner_integral_reflect_real_lemma) lemma integrable_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integrable (lebesgue_on {-b..-a}) (\x. f(-x)) \ integrable (lebesgue_on {a..b}) f" by (metis has_bochner_integral_iff has_bochner_integral_reflect_real) lemma integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integral\<^sup>L (lebesgue_on {-b .. -a}) (\x. f(-x)) = integral\<^sup>L (lebesgue_on {a..b::real}) f" using has_bochner_integral_reflect_real [of b a f] by (metis has_bochner_integral_iff not_integrable_integral_eq) subsection\More results on integrability\ lemma integrable_on_all_intervals_UNIV: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. f integrable_on cbox a b" and normf: "\x. norm(f x) \ g x" and g: "g integrable_on UNIV" shows "f integrable_on UNIV" proof - have intg: "(\a b. g integrable_on cbox a b)" and gle_e: "\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ \integral (cbox a b) g - integral (cbox c d) g\ < e" using g by (auto simp: integrable_alt_subset [of _ UNIV] intf) have le: "norm (integral (cbox a b) f - integral (cbox c d) f) \ \integral (cbox a b) g - integral (cbox c d) g\" if "cbox a b \ cbox c d" for a b c d proof - have "norm (integral (cbox a b) f - integral (cbox c d) f) = norm (integral (cbox c d - cbox a b) f)" using intf that by (simp add: norm_minus_commute integral_setdiff) also have "\ \ integral (cbox c d - cbox a b) g" proof (rule integral_norm_bound_integral [OF _ _ normf]) show "f integrable_on cbox c d - cbox a b" "g integrable_on cbox c d - cbox a b" by (meson integrable_integral integrable_setdiff intf intg negligible_setdiff that)+ qed also have "\ = integral (cbox c d) g - integral (cbox a b) g" using intg that by (simp add: integral_setdiff) also have "\ \ \integral (cbox a b) g - integral (cbox c d) g\" by simp finally show ?thesis . qed show ?thesis using gle_e apply (simp add: integrable_alt_subset [of _ UNIV] intf) apply (erule imp_forward all_forward ex_forward asm_rl)+ by (meson not_less order_trans le) qed lemma integrable_on_all_intervals_integrable_bound: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b" and normf: "\x. x \ S \ norm(f x) \ g x" and g: "g integrable_on S" shows "f integrable_on S" using integrable_on_all_intervals_UNIV [OF intf, of "(\x. if x \ S then g x else 0)"] by (simp add: g integrable_restrict_UNIV normf) lemma measurable_bounded_lemma: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable lebesgue" and g: "g integrable_on cbox a b" and normf: "\x. x \ cbox a b \ norm(f x) \ g x" shows "f integrable_on cbox a b" proof - have "g absolutely_integrable_on cbox a b" by (metis (full_types) add_increasing g le_add_same_cancel1 nonnegative_absolutely_integrable_1 norm_ge_zero normf) then have "integrable (lebesgue_on (cbox a b)) g" by (simp add: integrable_restrict_space set_integrable_def) then have "integrable (lebesgue_on (cbox a b)) f" proof (rule Bochner_Integration.integrable_bound) show "AE x in lebesgue_on (cbox a b). norm (f x) \ norm (g x)" by (rule AE_I2) (auto intro: normf order_trans) qed (simp add: f measurable_restrict_space1) then show ?thesis by (simp add: integrable_on_lebesgue_on) qed proposition measurable_bounded_by_integrable_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "g integrable_on S" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "f integrable_on S" proof (rule integrable_on_all_intervals_integrable_bound [OF _ normf g]) show "(\x. if x \ S then f x else 0) integrable_on cbox a b" for a b proof (rule measurable_bounded_lemma) show "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" by (simp add: S borel_measurable_if f) show "(\x. if x \ S then g x else 0) integrable_on cbox a b" by (simp add: g integrable_altD(1)) show "norm (if x \ S then f x else 0) \ (if x \ S then g x else 0)" for x using normf by simp qed qed lemma measurable_bounded_by_integrable_imp_lebesgue_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "integrable (lebesgue_on S) g" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" proof - have "f absolutely_integrable_on S" by (metis (no_types) S absolutely_integrable_integrable_bound f g integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_integrable normf) then show ?thesis by (simp add: S integrable_restrict_space set_integrable_def) qed lemma measurable_bounded_by_integrable_imp_integrable_real: fixes f :: "'a::euclidean_space \ real" assumes "f \ borel_measurable (lebesgue_on S)" "g integrable_on S" "\x. x \ S \ abs(f x) \ g x" "S \ sets lebesgue" shows "f integrable_on S" using measurable_bounded_by_integrable_imp_integrable [of f S g] assms by simp subsection\ Relation between Borel measurability and integrability.\ lemma integrable_imp_measurable_weak: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" by (metis (mono_tags, lifting) assms has_integral_implies_lebesgue_measurable borel_measurable_restrict_space_iff integrable_on_def sets.Int_space_eq2) lemma integrable_imp_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" proof - have "(UNIV::'a set) \ sets lborel" by simp then show ?thesis by (metis (mono_tags, lifting) assms borel_measurable_if_D integrable_imp_measurable_weak integrable_restrict_UNIV lebesgue_on_UNIV_eq sets_lebesgue_on_refl) qed lemma integrable_iff_integrable_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "(\\<^sup>+ x. ennreal (norm (f x)) \lebesgue_on S) < \" shows "integrable (lebesgue_on S) f \ f integrable_on S" using assms integrable_iff_bounded integrable_imp_measurable integrable_on_lebesgue_on by blast lemma absolutely_integrable_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (norm \ f)" (is "?lhs = ?rhs") proof assume L: ?lhs then have "f \ borel_measurable (lebesgue_on S)" by (simp add: absolutely_integrable_on_def integrable_imp_measurable) then show ?rhs using assms set_integrable_norm [of lebesgue S f] L by (simp add: integrable_restrict_space set_integrable_def) next assume ?rhs then show ?lhs using assms integrable_on_lebesgue_on by (metis absolutely_integrable_integrable_bound comp_def eq_iff measurable_bounded_by_integrable_imp_integrable) qed lemma absolutely_integrable_measurable_real: fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (\x. \f x\)" by (simp add: absolutely_integrable_measurable assms o_def) lemma absolutely_integrable_measurable_real': fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ (\x. \f x\) integrable_on S" by (metis abs_absolutely_integrableI_1 absolutely_integrable_measurable_real assms measurable_bounded_by_integrable_imp_integrable order_refl real_norm_def set_integrable_abs set_lebesgue_integral_eq_integral(1)) lemma absolutely_integrable_imp_borel_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S)" using absolutely_integrable_measurable assms by blast lemma measurable_bounded_by_integrable_imp_absolutely_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" and "g integrable_on S" and "\x. x \ S \ norm(f x) \ (g x)" shows "f absolutely_integrable_on S" using assms absolutely_integrable_integrable_bound measurable_bounded_by_integrable_imp_integrable by blast proposition negligible_differentiable_vimage: fixes f :: "'a \ 'a::euclidean_space" assumes "negligible T" and f': "\x. x \ S \ inj(f' x)" and derf: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "negligible {x \ S. f x \ T}" proof - define U where "U \ \n::nat. {x \ S. \y. y \ S \ norm(y - x) < 1/n \ norm(y - x) \ n * norm(f y - f x)}" have "negligible {x \ U n. f x \ T}" if "n > 0" for n proof (subst locally_negligible_alt, clarify) fix a assume a: "a \ U n" and fa: "f a \ T" define V where "V \ {x. x \ U n \ f x \ T} \ ball a (1 / n / 2)" show "\V. openin (top_of_set {x \ U n. f x \ T}) V \ a \ V \ negligible V" proof (intro exI conjI) have noxy: "norm(x - y) \ n * norm(f x - f y)" if "x \ V" "y \ V" for x y using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm by (meson norm_triangle_half_r) then have "inj_on f V" by (force simp: inj_on_def) then obtain g where g: "\x. x \ V \ g(f x) = x" by (metis inv_into_f_f) have "\T' B. open T' \ f x \ T' \ (\y\f ` V \ T \ T'. norm (g y - g (f x)) \ B * norm (y - f x))" if "f x \ T" "x \ V" for x using that noxy by (rule_tac x = "ball (f x) 1" in exI) (force simp: g) then have "negligible (g ` (f ` V \ T))" by (force simp: \negligible T\ negligible_Int intro!: negligible_locally_Lipschitz_image) moreover have "V \ g ` (f ` V \ T)" by (force simp: g image_iff V_def) ultimately show "negligible V" by (rule negligible_subset) qed (use a fa V_def that in auto) qed with negligible_countable_Union have "negligible (\n \ {0<..}. {x. x \ U n \ f x \ T})" by auto moreover have "{x \ S. f x \ T} \ (\n \ {0<..}. {x. x \ U n \ f x \ T})" proof clarsimp fix x assume "x \ S" and "f x \ T" then obtain inj: "inj(f' x)" and der: "(f has_derivative f' x) (at x within S)" using assms by metis moreover have "linear(f' x)" and eps: "\\. \ > 0 \ \\>0. \y\S. norm (y - x) < \ \ norm (f y - f x - f' x (y - x)) \ \ * norm (y - x)" using der by (auto simp: has_derivative_within_alt linear_linear) ultimately obtain g where "linear g" and g: "g \ f' x = id" using linear_injective_left_inverse by metis then obtain B where "B > 0" and B: "\z. B * norm z \ norm(f' x z)" using linear_invertible_bounded_below_pos \linear (f' x)\ by blast then obtain i where "i \ 0" and i: "1 / real i < B" by (metis (full_types) inverse_eq_divide real_arch_invD) then obtain \ where "\ > 0" and \: "\y. \y\S; norm (y - x) < \\ \ norm (f y - f x - f' x (y - x)) \ (B - 1 / real i) * norm (y - x)" using eps [of "B - 1/i"] by auto then obtain j where "j \ 0" and j: "inverse (real j) < \" using real_arch_inverse by blast have "norm (y - x)/(max i j) \ norm (f y - f x)" if "y \ S" and less: "norm (y - x) < 1 / (max i j)" for y proof - have "1 / real (max i j) < \" using j \j \ 0\ \0 < \\ by (auto simp: field_split_simps max_mult_distrib_left of_nat_max) then have "norm (y - x) < \" using less by linarith with \ \y \ S\ have le: "norm (f y - f x - f' x (y - x)) \ B * norm (y - x) - norm (y - x)/i" by (auto simp: algebra_simps) have "norm (y - x) / real (max i j) \ norm (y - x) / real i" using \i \ 0\ \j \ 0\ by (simp add: field_split_simps max_mult_distrib_left of_nat_max less_max_iff_disj) also have "... \ norm (f y - f x)" using B [of "y-x"] le norm_triangle_ineq3 [of "f y - f x" "f' x (y - x)"] by linarith finally show ?thesis . qed with \x \ S\ \i \ 0\ \j \ 0\ show "\n\{0<..}. x \ U n" by (rule_tac x="max i j" in bexI) (auto simp: field_simps U_def less_max_iff_disj) qed ultimately show ?thesis by (rule negligible_subset) qed lemma absolutely_integrable_Un: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "f absolutely_integrable_on S" and T: "f absolutely_integrable_on T" shows "f absolutely_integrable_on (S \ T)" proof - have [simp]: "{x. (if x \ A then f x else 0) \ 0} = {x \ A. f x \ 0}" for A by auto let ?ST = "{x \ S. f x \ 0} \ {x \ T. f x \ 0}" have "?ST \ sets lebesgue" proof (rule Sigma_Algebra.sets.Int) have "f integrable_on S" using S absolutely_integrable_on_def by blast then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ S then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ S. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto next have "f integrable_on T" using T absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ T. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto qed then have "f absolutely_integrable_on ?ST" by (rule set_integrable_subset [OF S]) auto then have Int: "(\x. if x \ ?ST then f x else 0) absolutely_integrable_on UNIV" using absolutely_integrable_restrict_UNIV by blast have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" "(\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" using S T absolutely_integrable_restrict_UNIV by blast+ then have "(\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)) absolutely_integrable_on UNIV" by (rule set_integral_add) then have "(\x. ((if x \ S then f x else 0) + (if x \ T then f x else 0)) - (if x \ ?ST then f x else 0)) absolutely_integrable_on UNIV" using Int by (rule set_integral_diff) then have "(\x. if x \ S \ T then f x else 0) absolutely_integrable_on UNIV" by (rule absolutely_integrable_spike) (auto intro: empty_imp_negligible) then show ?thesis unfolding absolutely_integrable_restrict_UNIV . qed lemma absolutely_integrable_on_combine: fixes f :: "real \ 'a::euclidean_space" assumes "f absolutely_integrable_on {a..c}" and "f absolutely_integrable_on {c..b}" and "a \ c" and "c \ b" shows "f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_Un assms ivl_disj_un_two_touch(4)) lemma uniform_limit_set_lebesgue_integral_at_top: fixes f :: "'a \ real \ 'b::{banach, second_countable_topology}" and g :: "real \ real" assumes bound: "\x y. x \ A \ y \ a \ norm (f x y) \ g y" assumes integrable: "set_integrable M {a..} g" assumes measurable: "\x. x \ A \ set_borel_measurable M {a..} (f x)" assumes "sets borel \ sets M" shows "uniform_limit A (\b x. LINT y:{a..b}|M. f x y) (\x. LINT y:{a..}|M. f x y) at_top" proof (cases "A = {}") case False then obtain x where x: "x \ A" by auto have g_nonneg: "g y \ 0" if "y \ a" for y proof - have "0 \ norm (f x y)" by simp also have "\ \ g y" using bound[OF x that] by simp finally show ?thesis . qed have integrable': "set_integrable M {a..} (\y. f x y)" if "x \ A" for x unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) show "integrable M (\x. indicator {a..} x * g x)" using integrable by (simp add: set_integrable_def) show "(\y. indicat_real {a..} y *\<^sub>R f x y) \ borel_measurable M" using measurable[OF that] by (simp add: set_borel_measurable_def) show "AE y in M. norm (indicat_real {a..} y *\<^sub>R f x y) \ norm (indicat_real {a..} y * g y)" using bound[OF that] by (intro AE_I2) (auto simp: indicator_def g_nonneg) qed show ?thesis proof (rule uniform_limitI) fix e :: real assume e: "e > 0" have sets [intro]: "A \ sets M" if "A \ sets borel" for A using that assms by blast have "((\b. LINT y:{a..b}|M. g y) \ (LINT y:{a..}|M. g y)) at_top" by (intro tendsto_set_lebesgue_integral_at_top assms sets) auto with e obtain b0 :: real where b0: "\b\b0. \(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: tendsto_iff eventually_at_top_linorder dist_real_def abs_minus_commute) define b where "b = max a b0" have "a \ b" by (simp add: b_def) from b0 have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: b_def) also have "{a..} = {a..b} \ {b<..}" by (auto simp: b_def) also have "\(LINT y:\|M. g y) - (LINT y:{a..b}|M. g y)\ = \(LINT y:{b<..}|M. g y)\" using \a \ b\ by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "\(LINT y:{b<..}|M. g y)\ = (LINT y:{b<..}|M. g y)" by simp finally have less: "(LINT y:{b<..}|M. g y) < e" . have "eventually (\b. b \ b0) at_top" by (rule eventually_ge_at_top) moreover have "eventually (\b. b \ a) at_top" by (rule eventually_ge_at_top) ultimately show "eventually (\b. \x\A. dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e) at_top" proof eventually_elim case (elim b) show ?case proof fix x assume x: "x \ A" have "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) = norm ((LINT y:{a..}|M. f x y) - (LINT y:{a..b}|M. f x y))" by (simp add: dist_norm norm_minus_commute) also have "{a..} = {a..b} \ {b<..}" using elim by auto also have "(LINT y:\|M. f x y) - (LINT y:{a..b}|M. f x y) = (LINT y:{b<..}|M. f x y)" using elim x by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable']) also have "norm \ \ (LINT y:{b<..}|M. norm (f x y))" using elim x by (intro set_integral_norm_bound set_integrable_subset[OF integrable']) auto also have "\ \ (LINT y:{b<..}|M. g y)" using elim x bound g_nonneg by (intro set_integral_mono set_integrable_norm set_integrable_subset[OF integrable'] set_integrable_subset[OF integrable]) auto also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "(LINT y:{b<..}|M. g y) = \(LINT y:{b<..}|M. g y)\" by simp also have "\ = \(LINT y:{a..b} \ {b<..}|M. g y) - (LINT y:{a..b}|M. g y)\" using elim by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "{a..b} \ {b<..} = {a..}" using elim by auto also have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" using b0 elim by blast finally show "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e" . qed qed qed qed auto subsubsection\Differentiability of inverse function (most basic form)\ proposition has_derivative_inverse_within: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes der_f: "(f has_derivative f') (at a within S)" and cont_g: "continuous (at (f a) within f ` S) g" and "a \ S" "linear g'" and id: "g' \ f' = id" and gf: "\x. x \ S \ g(f x) = x" shows "(g has_derivative g') (at (f a) within f ` S)" proof - have [simp]: "g' (f' x) = x" for x by (simp add: local.id pointfree_idE) have "bounded_linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - a) < d \ norm (f y - f a - f' (y - a)) \ e * norm (y - a)" using der_f by (auto simp: has_derivative_within_alt) obtain C where "C > 0" and C: "\x. norm (g' x) \ C * norm x" using linear_bounded_pos [OF \linear g'\] by metis obtain B k where "B > 0" "k > 0" and Bk: "\x. \x \ S; norm(f x - f a) < k\ \ norm(x - a) \ B * norm(f x - f a)" proof - obtain B where "B > 0" and B: "\x. B * norm x \ norm (f' x)" using linear_inj_bounded_below_pos [of f'] \linear g'\ id der_f has_derivative_linear linear_invertible_bounded_below_pos by blast then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ B / 2 * norm (y - a)" using f' [of "B/2"] by auto then obtain e where "e > 0" and e: "\x. \x \ S; norm (f x - f a) < e\ \ norm (g (f x) - g (f a)) < d" using cont_g by (auto simp: continuous_within_eps_delta dist_norm) show thesis proof show "2/B > 0" using \B > 0\ by simp show "norm (x - a) \ 2 / B * norm (f x - f a)" if "x \ S" "norm (f x - f a) < e" for x proof - have xa: "norm (x - a) < d" using e [OF that] gf by (simp add: \a \ S\ that) have *: "\norm(y - f') \ B / 2 * norm x; B * norm x \ norm f'\ \ norm y \ B / 2 * norm x" for y f'::'b and x::'a using norm_triangle_ineq3 [of y f'] by linarith show ?thesis using * [OF d [OF \x \ S\ xa] B] \B > 0\ by (simp add: field_simps) qed qed (use \e > 0\ in auto) qed show ?thesis unfolding has_derivative_within_alt proof (intro conjI impI allI) show "bounded_linear g'" using \linear g'\ by (simp add: linear_linear) next fix e :: "real" assume "e > 0" then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ e / (B * C) * norm (y - a)" using f' [of "e / (B * C)"] \B > 0\ \C > 0\ by auto have "norm (x - a - g' (f x - f a)) \ e * norm (f x - f a)" if "x \ S" and lt_k: "norm (f x - f a) < k" and lt_dB: "norm (f x - f a) < d/B" for x proof - have "norm (x - a) \ B * norm(f x - f a)" using Bk lt_k \x \ S\ by blast also have "\ < d" by (metis \0 < B\ lt_dB mult.commute pos_less_divide_eq) finally have lt_d: "norm (x - a) < d" . have "norm (x - a - g' (f x - f a)) \ norm(g'(f x - f a - (f' (x - a))))" by (simp add: linear_diff [OF \linear g'\] norm_minus_commute) also have "\ \ C * norm (f x - f a - f' (x - a))" using C by blast also have "\ \ e * norm (f x - f a)" proof - have "norm (f x - f a - f' (x - a)) \ e / (B * C) * norm (x - a)" using d [OF \x \ S\ lt_d] . also have "\ \ (norm (f x - f a) * e) / C" using \B > 0\ \C > 0\ \e > 0\ by (simp add: field_simps Bk lt_k \x \ S\) finally show ?thesis using \C > 0\ by (simp add: field_simps) qed finally show ?thesis . qed with \k > 0\ \B > 0\ \d > 0\ \a \ S\ show "\d>0. \y\f ` S. norm (y - f a) < d \ norm (g y - g (f a) - g' (y - f a)) \ e * norm (y - f a)" by (rule_tac x="min k (d / B)" in exI) (auto simp: gf) qed qed end diff --git a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy @@ -1,7592 +1,7592 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) Huge cleanup by LCP *) section \Henstock-Kurzweil Gauge Integration in Many Dimensions\ theory Henstock_Kurzweil_Integration imports Lebesgue_Measure Tagged_Division begin lemma norm_diff2: "\y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \ e1; norm(y2 - x2) \ e2\ \ norm(y-x) \ e" using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"] by (simp add: add_diff_add) lemma setcomp_dot1: "{z. P (z \ (i,0))} = {(x,y). P(x \ i)}" by auto lemma setcomp_dot2: "{z. P (z \ (0,i))} = {(x,y). P(y \ i)}" by auto lemma Sigma_Int_Paircomp1: "(Sigma A B) \ {(x, y). P x} = Sigma (A \ {x. P x}) B" by blast lemma Sigma_Int_Paircomp2: "(Sigma A B) \ {(x, y). P y} = Sigma A (\z. B z \ {y. P y})" by blast (* END MOVE *) subsection \Content (length, area, volume...) of an interval\ abbreviation content :: "'a::euclidean_space set \ real" where "content s \ measure lborel s" lemma content_cbox_cases: "content (cbox a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_cbox_eq inner_diff) lemma content_cbox: "\i\Basis. a\i \ b\i \ content (cbox a b) = (\i\Basis. b\i - a\i)" unfolding content_cbox_cases by simp lemma content_cbox': "cbox a b \ {} \ content (cbox a b) = (\i\Basis. b\i - a\i)" by (simp add: box_ne_empty inner_diff) lemma content_cbox_if: "content (cbox a b) = (if cbox a b = {} then 0 else \i\Basis. b\i - a\i)" by (simp add: content_cbox') lemma content_cbox_cart: "cbox a b \ {} \ content(cbox a b) = prod (\i. b$i - a$i) UNIV" by (simp add: content_cbox_if Basis_vec_def cart_eq_inner_axis axis_eq_axis prod.UNION_disjoint) lemma content_cbox_if_cart: "content(cbox a b) = (if cbox a b = {} then 0 else prod (\i. b$i - a$i) UNIV)" by (simp add: content_cbox_cart) lemma content_division_of: assumes "K \ \" "\ division_of S" shows "content K = (\i \ Basis. interval_upperbound K \ i - interval_lowerbound K \ i)" proof - obtain a b where "K = cbox a b" using cbox_division_memE assms by metis then show ?thesis using assms by (force simp: division_of_def content_cbox') qed lemma content_real: "a \ b \ content {a..b} = b - a" by simp lemma abs_eq_content: "\y - x\ = (if x\y then content {x..y} else content {y..x})" by (auto simp: content_real) lemma content_singleton: "content {a} = 0" by simp lemma content_unit[iff]: "content (cbox 0 (One::'a::euclidean_space)) = 1" by simp lemma content_pos_le [iff]: "0 \ content X" by simp corollary\<^marker>\tag unimportant\ content_nonneg [simp]: "\ content (cbox a b) < 0" using not_le by blast lemma content_pos_lt: "\i\Basis. a\i < b\i \ 0 < content (cbox a b)" by (auto simp: less_imp_le inner_diff box_eq_empty intro!: prod_pos) lemma content_eq_0: "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_cbox_cases not_le intro: less_imp_le antisym eq_refl) lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" unfolding content_eq_0 interior_cbox box_eq_empty by auto lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_empty [simp]: "content {} = 0" by simp lemma content_real_if [simp]: "content {a..b} = (if a \ b then b - a else 0)" by (simp add: content_real) lemma content_subset: "cbox a b \ cbox c d \ content (cbox a b) \ content (cbox c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq) lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" unfolding measure_lborel_cbox_eq Basis_prod_def apply (subst prod.union_disjoint) apply (auto simp: bex_Un ball_Un) apply (subst (1 2) prod.reindex_nontrivial) apply auto done lemma content_cbox_pair_eq0_D: "content (cbox (a,c) (b,d)) = 0 \ content (cbox a b) = 0 \ content (cbox c d) = 0" by (simp add: content_Pair) lemma content_cbox_plus: fixes x :: "'a::euclidean_space" shows "content(cbox x (x + h *\<^sub>R One)) = (if h \ 0 then h ^ DIM('a) else 0)" by (simp add: algebra_simps content_cbox_if box_eq_empty) lemma content_0_subset: "content(cbox a b) = 0 \ s \ cbox a b \ content s = 0" using emeasure_mono[of s "cbox a b" lborel] by (auto simp: measure_def enn2real_eq_0_iff emeasure_lborel_cbox_eq) lemma content_ball_pos: assumes "r > 0" shows "content (ball c r) > 0" proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto from ab have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" using ab by (intro emeasure_mono) auto also have "emeasure lborel (box a b) = ennreal (content (box a b))" using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto also have "emeasure lborel (ball c r) = ennreal (content (ball c r))" using emeasure_lborel_ball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto finally show ?thesis using \content (box a b) > 0\ by simp qed lemma content_cball_pos: assumes "r > 0" shows "content (cball c r) > 0" proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto from ab have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" using ab by (intro emeasure_mono) auto also have "\ \ emeasure lborel (cball c r)" by (intro emeasure_mono) auto also have "emeasure lborel (box a b) = ennreal (content (box a b))" using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto also have "emeasure lborel (cball c r) = ennreal (content (cball c r))" using emeasure_lborel_cball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto finally show ?thesis using \content (box a b) > 0\ by simp qed lemma content_split: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" \ \Prove using measure theory\ proof (cases "\i\Basis. a \ i \ b \ i") case True have 1: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" by (simp add: if_distrib prod.delta_remove assms) note simps = interval_split[OF assms] content_cbox_cases have 2: "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" by (metis (no_types, lifting) assms finite_Basis mult.commute prod.remove) have "\x. min (b \ k) c = max (a \ k) c \ x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" by (auto simp add: field_simps) moreover have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" by (auto intro!: prod.cong) have "\ a \ k \ c \ \ c \ b \ k \ False" unfolding not_le using True assms by auto ultimately show ?thesis using assms unfolding simps ** 1[of "\i x. b\i - x"] 1[of "\i x. x - a\i"] 2 by auto next case False then have "cbox a b = {}" unfolding box_eq_empty by (auto simp: not_le) then show ?thesis by (auto simp: not_le) qed lemma division_of_content_0: assumes "content (cbox a b) = 0" "d division_of (cbox a b)" "K \ d" shows "content K = 0" unfolding forall_in_division[OF assms(2)] by (meson assms content_0_subset division_of_def) lemma sum_content_null: assumes "content (cbox a b) = 0" and "p tagged_division_of (cbox a b)" shows "(\(x,K)\p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)" proof (rule sum.neutral, rule) fix y assume y: "y \ p" obtain x K where xk: "y = (x, K)" using surj_pair[of y] by blast then obtain c d where k: "K = cbox c d" "K \ cbox a b" by (metis assms(2) tagged_division_ofD(3) tagged_division_ofD(4) y) have "(\(x',K'). content K' *\<^sub>R f x') y = content K *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using assms(1) content_0_subset k(2) by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed global_interpretation sum_content: operative plus 0 content rewrites "comm_monoid_set.F plus 0 = sum" proof - interpret operative plus 0 content by standard (auto simp add: content_split [symmetric] content_eq_0_interior) show "operative plus 0 content" by standard show "comm_monoid_set.F plus 0 = sum" by (simp add: sum_def) qed lemma additive_content_division: "d division_of (cbox a b) \ sum content d = content (cbox a b)" by (fact sum_content.division) lemma additive_content_tagged_division: "d tagged_division_of (cbox a b) \ sum (\(x,l). content l) d = content (cbox a b)" by (fact sum_content.tagged_division) lemma subadditive_content_division: assumes "\ division_of S" "S \ cbox a b" shows "sum content \ \ content(cbox a b)" proof - have "\ division_of \\" "\\ \ cbox a b" using assms by auto then obtain \' where "\ \ \'" "\' division_of cbox a b" using partial_division_extend_interval by metis then have "sum content \ \ sum content \'" using sum_mono2 by blast also have "... \ content(cbox a b)" by (simp add: \\' division_of cbox a b\ additive_content_division less_eq_real_def) finally show ?thesis . qed lemma content_real_eq_0: "content {a..b::real} = 0 \ a \ b" by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" using content_empty unfolding empty_as_interval by auto lemma interval_bounds_nz_content [simp]: assumes "content (cbox a b) \ 0" shows "interval_upperbound (cbox a b) = b" and "interval_lowerbound (cbox a b) = a" by (metis assms content_empty interval_bounds')+ subsection \Gauge integral\ text \Case distinction to define it first on compact intervals first, then use a limit. This is only much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\ definition has_integral :: "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" (infixr "has'_integral" 46) where "(f has_integral I) s \ (if \a b. s = cbox a b then ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter s) else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\p. \(x,k)\p. content k *\<^sub>R (if x \ s then f x else 0)) \ z) (division_filter (cbox a b)) \ norm (z - I) < e)))" lemma has_integral_cbox: "(f has_integral I) (cbox a b) \ ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter (cbox a b))" by (auto simp add: has_integral_def) lemma has_integral: "(f has_integral y) (cbox a b) \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff) lemma has_integral_real: "(f has_integral y) {a..b::real} \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of {a..b} \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" unfolding box_real[symmetric] by (rule has_integral) lemma has_integralD[dest]: assumes "(f has_integral y) (cbox a b)" and "e > 0" obtains \ where "gauge \" and "\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,k)\\. content k *\<^sub>R f x) - y) < e" using assms unfolding has_integral by auto lemma has_integral_alt: "(f has_integral y) i \ (if \a b. i = cbox a b then (f has_integral y) i else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" by (subst has_integral_def) (auto simp add: has_integral_cbox) lemma has_integral_altD: assumes "(f has_integral y) i" and "\ (\a b. i = cbox a b)" and "e>0" obtains B where "B > 0" and "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" using assms has_integral_alt[of f y i] by auto definition integrable_on (infixr "integrable'_on" 46) where "f integrable_on i \ (\y. (f has_integral y) i)" definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" lemma integrable_integral[intro]: "f integrable_on i \ (f has_integral (integral i f)) i" unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) lemma not_integrable_integral: "\ f integrable_on i \ integral i f = 0" unfolding integrable_on_def integral_def by blast lemma has_integral_integrable[dest]: "(f has_integral i) s \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_integral: "f integrable_on s \ (f has_integral (integral s f)) s" by auto subsection \Basic theorems about integrals\ lemma has_integral_eq_rhs: "(f has_integral j) S \ i = j \ (f has_integral i) S" by (rule forw_subst) lemma has_integral_unique_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" shows "(f has_integral k1) (cbox a b) \ (f has_integral k2) (cbox a b) \ k1 = k2" by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty]) lemma has_integral_unique: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2" proof (rule ccontr) let ?e = "norm (k1 - k2)/2" let ?F = "(\x. if x \ i then f x else 0)" assume "k1 \ k2" then have e: "?e > 0" by auto have nonbox: "\ (\a b. i = cbox a b)" using \k1 \ k2\ assms has_integral_unique_cbox by blast obtain B1 where B1: "0 < B1" "\a b. ball 0 B1 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k1) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(1) nonbox,OF e]) blast obtain B2 where B2: "0 < B2" "\a b. ball 0 B2 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k2) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(2) nonbox,OF e]) blast obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (metis Un_subset_iff bounded_Un bounded_ball bounded_subset_cbox_symmetric) obtain w where w: "(?F has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2)/2" using B1(2)[OF ab(1)] by blast obtain z where z: "(?F has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2)/2" using B2(2)[OF ab(2)] by blast have "z = w" using has_integral_unique_cbox[OF w(1) z(1)] by auto then have "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] by (auto simp add: norm_minus_commute) also have "\ < norm (k1 - k2)/2 + norm (k1 - k2)/2" by (metis add_strict_mono z(2) w(2)) finally show False by auto qed lemma integral_unique [intro]: "(f has_integral y) k \ integral k f = y" unfolding integral_def by (rule some_equality) (auto intro: has_integral_unique) lemma has_integral_iff: "(f has_integral i) S \ (f integrable_on S \ integral S f = i)" by blast lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ \ f integrable_on k \ y=0" unfolding integral_def integrable_on_def apply (erule subst) apply (rule someI_ex) by blast lemma has_integral_const [intro]: fixes a b :: "'a::euclidean_space" shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" using eventually_division_filter_tagged_division[of "cbox a b"] additive_content_tagged_division[of _ a b] by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric] elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const]) lemma has_integral_const_real [intro]: fixes a b :: real shows "((\x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}" by (metis box_real(2) has_integral_const) lemma has_integral_integrable_integral: "(f has_integral i) s \ f integrable_on s \ integral s f = i" by blast lemma integral_const [simp]: fixes a b :: "'a::euclidean_space" shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) lemma integral_const_real [simp]: fixes a b :: real shows "integral {a..b} (\x. c) = content {a..b} *\<^sub>R c" by (metis box_real(2) integral_const) lemma has_integral_is_0_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ cbox a b \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] assms by (subst tendsto_cong[where g="\_. 0"]) (auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval) lemma has_integral_is_0: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ S \ f x = 0" shows "(f has_integral 0) S" proof (cases "(\a b. S = cbox a b)") case True with assms has_integral_is_0_cbox show ?thesis by blast next case False have *: "(\x. if x \ S then f x else 0) = (\x. 0)" by (auto simp add: assms) show ?thesis using has_integral_is_0_cbox False by (subst has_integral_alt) (force simp add: *) qed lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) S" by (rule has_integral_is_0) auto lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) S \ i = 0" using has_integral_unique[OF has_integral_0] by auto lemma has_integral_linear_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) (cbox a b)" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) (cbox a b)" proof - interpret bounded_linear h using h . show ?thesis unfolding has_integral_cbox using tendsto [OF f [unfolded has_integral_cbox]] by (simp add: sum scaleR split_beta') qed lemma has_integral_linear: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) S" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) S" proof (cases "(\a b. S = cbox a b)") case True with f h has_integral_linear_cbox show ?thesis by blast next case False interpret bounded_linear h using h . from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" by blast let ?S = "\f x. if x \ S then f x else 0" show ?thesis proof (subst has_integral_alt, clarsimp simp: False) fix e :: real assume e: "e > 0" have *: "0 < e/B" using e B(1) by simp obtain M where M: "M > 0" "\a b. ball 0 M \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - y) < e/B" using has_integral_altD[OF f False *] by blast show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e)" proof (rule exI, intro allI conjI impI) show "M > 0" using M by metis next fix a b::'n assume sb: "ball 0 M \ cbox a b" obtain z where z: "(?S f has_integral z) (cbox a b)" "norm (z - y) < e/B" using M(2)[OF sb] by blast have *: "?S(h \ f) = h \ ?S f" using zero by auto show "\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e" proof (intro exI conjI) show "(?S(h \ f) has_integral h z) (cbox a b)" by (simp add: * has_integral_linear_cbox[OF z(1) h]) show "norm (h z - h y) < e" by (metis B diff le_less_trans pos_less_divide_eq z(2)) qed qed qed qed lemma has_integral_scaleR_left: "(f has_integral y) S \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) S" using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) lemma integrable_on_scaleR_left: assumes "f integrable_on A" shows "(\x. f x *\<^sub>R y) integrable_on A" using assms has_integral_scaleR_left unfolding integrable_on_def by blast lemma has_integral_mult_left: fixes c :: "_ :: real_normed_algebra" shows "(f has_integral y) S \ ((\x. f x * c) has_integral (y * c)) S" using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) lemma has_integral_divide: fixes c :: "_ :: real_normed_div_algebra" shows "(f has_integral y) S \ ((\x. f x / c) has_integral (y / c)) S" unfolding divide_inverse by (simp add: has_integral_mult_left) text\The case analysis eliminates the condition \<^term>\f integrable_on S\ at the cost of the type class constraint \division_ring\\ corollary integral_mult_left [simp]: fixes c:: "'a::{real_normed_algebra,division_ring}" shows "integral S (\x. f x * c) = integral S f * c" proof (cases "f integrable_on S \ c = 0") case True then show ?thesis by (force intro: has_integral_mult_left) next case False then have "\ (\x. f x * c) integrable_on S" using has_integral_mult_left [of "(\x. f x * c)" _ S "inverse c"] by (auto simp add: mult.assoc) with False show ?thesis by (simp add: not_integrable_integral) qed corollary integral_mult_right [simp]: fixes c:: "'a::{real_normed_field}" shows "integral S (\x. c * f x) = c * integral S f" by (simp add: mult.commute [of c]) corollary integral_divide [simp]: fixes z :: "'a::real_normed_field" shows "integral S (\x. f x / z) = integral S (\x. f x) / z" using integral_mult_left [of S f "inverse z"] by (simp add: divide_inverse_commute) lemma has_integral_mult_right: fixes c :: "'a :: real_normed_algebra" shows "(f has_integral y) i \ ((\x. c * f x) has_integral (c * y)) i" using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) lemma has_integral_cmul: "(f has_integral k) S \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) S" unfolding o_def[symmetric] by (metis has_integral_linear bounded_linear_scaleR_right) lemma has_integral_cmult_real: fixes c :: real assumes "c \ 0 \ (f has_integral x) A" shows "((\x. c * f x) has_integral c * x) A" proof (cases "c = 0") case True then show ?thesis by simp next case False from has_integral_cmul[OF assms[OF this], of c] show ?thesis unfolding real_scaleR_def . qed lemma has_integral_neg: "(f has_integral k) S \ ((\x. -(f x)) has_integral -k) S" by (drule_tac c="-1" in has_integral_cmul) auto lemma has_integral_neg_iff: "((\x. - f x) has_integral k) S \ (f has_integral - k) S" using has_integral_neg[of f "- k"] has_integral_neg[of "\x. - f x" k] by auto lemma has_integral_add_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k) (cbox a b)" "(g has_integral l) (cbox a b)" shows "((\x. f x + g x) has_integral (k + l)) (cbox a b)" using assms unfolding has_integral_cbox by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add) lemma has_integral_add: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral k) S" and g: "(g has_integral l) S" shows "((\x. f x + g x) has_integral (k + l)) S" proof (cases "\a b. S = cbox a b") case True with has_integral_add_cbox assms show ?thesis by blast next let ?S = "\f x. if x \ S then f x else 0" case False then show ?thesis proof (subst has_integral_alt, clarsimp, goal_cases) case (1 e) then have e2: "e/2 > 0" by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - k) < e/2" using has_integral_altD[OF f False e2] by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ (cbox a b) \ \z. (?S g has_integral z) (cbox a b) \ norm (z - l) < e/2" using has_integral_altD[OF g False e2] by blast show ?case proof (rule_tac x="max Bf Bg" in exI, clarsimp simp add: max.strict_coboundedI1 \0 < Bf\) fix a b assume "ball 0 (max Bf Bg) \ cbox a (b::'n)" then have fs: "ball 0 Bf \ cbox a (b::'n)" and gs: "ball 0 Bg \ cbox a (b::'n)" by auto obtain w where w: "(?S f has_integral w) (cbox a b)" "norm (w - k) < e/2" using Bf[OF fs] by blast obtain z where z: "(?S g has_integral z) (cbox a b)" "norm (z - l) < e/2" using Bg[OF gs] by blast have *: "\x. (if x \ S then f x + g x else 0) = (?S f x) + (?S g x)" by auto show "\z. (?S(\x. f x + g x) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" proof (intro exI conjI) show "(?S(\x. f x + g x) has_integral (w + z)) (cbox a b)" by (simp add: has_integral_add_cbox[OF w(1) z(1), unfolded *[symmetric]]) show "norm (w + z - (k + l)) < e" by (metis dist_norm dist_triangle_add_half w(2) z(2)) qed qed qed qed lemma has_integral_diff: "(f has_integral k) S \ (g has_integral l) S \ ((\x. f x - g x) has_integral (k - l)) S" using has_integral_add[OF _ has_integral_neg, of f k S g l] by (auto simp: algebra_simps) lemma integral_0 [simp]: "integral S (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" by (rule integral_unique has_integral_0)+ lemma integral_add: "f integrable_on S \ g integrable_on S \ integral S (\x. f x + g x) = integral S f + integral S g" by (rule integral_unique) (metis integrable_integral has_integral_add) lemma integral_cmul [simp]: "integral S (\x. c *\<^sub>R f x) = c *\<^sub>R integral S f" proof (cases "f integrable_on S \ c = 0") case True with has_integral_cmul integrable_integral show ?thesis by fastforce next case False then have "\ (\x. c *\<^sub>R f x) integrable_on S" using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ S "inverse c"] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_mult: fixes K::real shows "f integrable_on X \ K * integral X f = integral X (\x. K * f x)" unfolding real_scaleR_def[symmetric] integral_cmul .. lemma integral_neg [simp]: "integral S (\x. - f x) = - integral S f" proof (cases "f integrable_on S") case True then show ?thesis by (simp add: has_integral_neg integrable_integral integral_unique) next case False then have "\ (\x. - f x) integrable_on S" using has_integral_neg [of "(\x. - f x)" _ S ] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_diff: "f integrable_on S \ g integrable_on S \ integral S (\x. f x - g x) = integral S f - integral S g" by (rule integral_unique) (metis integrable_integral has_integral_diff) lemma integrable_0: "(\x. 0) integrable_on S" unfolding integrable_on_def using has_integral_0 by auto lemma integrable_add: "f integrable_on S \ g integrable_on S \ (\x. f x + g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_add) lemma integrable_cmul: "f integrable_on S \ (\x. c *\<^sub>R f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_cmul) lemma integrable_on_scaleR_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c *\<^sub>R f x) integrable_on S \ f integrable_on S" using integrable_cmul[of "\x. c *\<^sub>R f x" S "1 / c"] integrable_cmul[of f S c] \c \ 0\ by auto lemma integrable_on_cmult_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c * f x) integrable_on S \ f integrable_on S" using integrable_on_scaleR_iff [of c f] assms by simp lemma integrable_on_cmult_left: assumes "f integrable_on S" shows "(\x. of_real c * f x) integrable_on S" using integrable_cmul[of f S "of_real c"] assms by (simp add: scaleR_conv_of_real) lemma integrable_neg: "f integrable_on S \ (\x. -f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_neg) lemma integrable_neg_iff: "(\x. -f(x)) integrable_on S \ f integrable_on S" using integrable_neg by fastforce lemma integrable_diff: "f integrable_on S \ g integrable_on S \ (\x. f x - g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_diff) lemma integrable_linear: "f integrable_on S \ bounded_linear h \ (h \ f) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_linear) lemma integral_linear: "f integrable_on S \ bounded_linear h \ integral S (h \ f) = h (integral S f)" by (meson has_integral_iff has_integral_linear) lemma integrable_on_cnj_iff: "(\x. cnj (f x)) integrable_on A \ f integrable_on A" using integrable_linear[OF _ bounded_linear_cnj, of f A] integrable_linear[OF _ bounded_linear_cnj, of "cnj \ f" A] by (auto simp: o_def) lemma integral_cnj: "cnj (integral A f) = integral A (\x. cnj (f x))" by (cases "f integrable_on A") (simp_all add: integral_linear[OF _ bounded_linear_cnj, symmetric] o_def integrable_on_cnj_iff not_integrable_integral) lemma integral_component_eq[simp]: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on S" shows "integral S (\x. f x \ k) = integral S f \ k" unfolding integral_linear[OF assms(1) bounded_linear_inner_left,unfolded o_def] .. lemma has_integral_sum: assumes "finite T" and "\a. a \ T \ ((f a) has_integral (i a)) S" shows "((\x. sum (\a. f a x) T) has_integral (sum i T)) S" using assms(1) subset_refl[of T] proof (induct rule: finite_subset_induct) case empty then show ?case by auto next case (insert x F) with assms show ?case by (simp add: has_integral_add) qed lemma integral_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ integral S (\x. \a\I. f a x) = (\a\I. integral S (f a))" by (simp add: has_integral_sum integrable_integral integral_unique) lemma integrable_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ (\x. \a\I. f a x) integrable_on S" unfolding integrable_on_def using has_integral_sum[of I] by metis lemma has_integral_eq: assumes "\x. x \ s \ f x = g x" and "(f has_integral k) s" shows "(g has_integral k) s" using has_integral_diff[OF assms(2), of "\x. f x - g x" 0] using has_integral_is_0[of s "\x. f x - g x"] using assms(1) by auto lemma integrable_eq: "\f integrable_on s; \x. x \ s \ f x = g x\ \ g integrable_on s" unfolding integrable_on_def using has_integral_eq[of s f g] has_integral_eq by blast lemma has_integral_cong: assumes "\x. x \ s \ f x = g x" shows "(f has_integral i) s = (g has_integral i) s" using has_integral_eq[of s f g] has_integral_eq[of s g f] assms by auto lemma integral_cong: assumes "\x. x \ s \ f x = g x" shows "integral s f = integral s g" unfolding integral_def by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) lemma integrable_on_cmult_left_iff [simp]: assumes "c \ 0" shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] by (simp add: scaleR_conv_of_real) then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" by (simp add: algebra_simps) with \c \ 0\ show ?rhs by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) qed (blast intro: integrable_on_cmult_left) lemma integrable_on_cmult_right: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "f integrable_on s" shows "(\x. f x * of_real c) integrable_on s" using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) lemma integrable_on_cmult_right_iff [simp]: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "c \ 0" shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) lemma integrable_on_cdivide: fixes f :: "_ \ 'b :: real_normed_field" assumes "f integrable_on s" shows "(\x. f x / of_real c) integrable_on s" by (simp add: integrable_on_cmult_right divide_inverse assms flip: of_real_inverse) lemma integrable_on_cdivide_iff [simp]: fixes f :: "_ \ 'b :: real_normed_field" assumes "c \ 0" shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" by (simp add: divide_inverse assms flip: of_real_inverse) lemma has_integral_null [intro]: "content(cbox a b) = 0 \ (f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] by (subst tendsto_cong[where g="\_. 0"]) (auto elim: eventually_mono intro: sum_content_null) lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \ (f has_integral 0) {a..b}" by (metis box_real(2) has_integral_null) lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" by (auto simp add: has_integral_null dest!: integral_unique) lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" by (metis has_integral_null integral_unique) lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" by (simp add: has_integral_integrable) lemma has_integral_empty[intro]: "(f has_integral 0) {}" by (meson ex_in_conv has_integral_is_0) lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" by (auto simp add: has_integral_empty has_integral_unique) lemma integrable_on_empty[intro]: "f integrable_on {}" unfolding integrable_on_def by auto lemma integral_empty[simp]: "integral {} f = 0" by (rule integral_unique) (rule has_integral_empty) lemma has_integral_refl[intro]: fixes a :: "'a::euclidean_space" shows "(f has_integral 0) (cbox a a)" and "(f has_integral 0) {a}" proof - show "(f has_integral 0) (cbox a a)" by (rule has_integral_null) simp then show "(f has_integral 0) {a}" by simp qed lemma integrable_on_refl[intro]: "f integrable_on cbox a a" unfolding integrable_on_def by auto lemma integral_refl [simp]: "integral (cbox a a) f = 0" by (rule integral_unique) auto lemma integral_singleton [simp]: "integral {a} f = 0" by auto lemma integral_blinfun_apply: assumes "f integrable_on s" shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) lemma blinfun_apply_integral: assumes "f integrable_on s" shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) lemma has_integral_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" proof safe fix b :: 'b assume "(f has_integral y) A" from has_integral_linear[OF this(1) bounded_linear_inner_left, of b] show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) next assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "(f has_integral y) A" by (simp add: euclidean_representation) qed lemma has_integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" by (subst has_integral_componentwise_iff) blast lemma integrable_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" proof assume "f integrable_on A" then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" by (subst (asm) has_integral_componentwise_iff) thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) next assume "(\b\Basis. (\x. f x \ b) integrable_on A)" then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" unfolding integrable_on_def by (subst (asm) bchoice_iff) blast hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) qed lemma integrable_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" by (subst integrable_componentwise_iff) blast lemma integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" proof - from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) (simp_all add: bounded_linear_scaleR_left) have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" by (subst integral_sum) (simp_all add: o_def) finally show ?thesis . qed lemma integrable_component: "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) subsection \Cauchy-type criterion for integrability\ proposition integrable_Cauchy: fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on cbox a b \ (\e>0. \\. gauge \ \ (\\1 \2. \1 tagged_division_of (cbox a b) \ \ fine \1 \ \2 tagged_division_of (cbox a b) \ \ fine \2 \ norm ((\(x,K)\\1. content K *\<^sub>R f x) - (\(x,K)\\2. content K *\<^sub>R f x)) < e))" (is "?l = (\e>0. \\. ?P e \)") proof (intro iffI allI impI) assume ?l then obtain y where y: "\e. e > 0 \ \\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" by (auto simp: integrable_on_def has_integral) show "\\. ?P e \" if "e > 0" for e proof - have "e/2 > 0" using that by auto with y obtain \ where "gauge \" and \: "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - y) < e/2" by meson show ?thesis apply (rule_tac x=\ in exI, clarsimp simp: \gauge \\) by (blast intro!: \ dist_triangle_half_l[where y=y,unfolded dist_norm]) qed next assume "\e>0. \\. ?P e \" then have "\n::nat. \\. ?P (1 / (n + 1)) \" by auto then obtain \ :: "nat \ 'n \ 'n set" where \: "\m. gauge (\ m)" "\m \1 \2. \\1 tagged_division_of cbox a b; \ m fine \1; \2 tagged_division_of cbox a b; \ m fine \2\ \ norm ((\(x,K) \ \1. content K *\<^sub>R f x) - (\(x,K) \ \2. content K *\<^sub>R f x)) < 1 / (m + 1)" by metis have "gauge (\x. \{\ i x |i. i \ {0..n}})" for n using \ by (intro gauge_Inter) auto then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{\ i x |i. i \ {0..n}}) fine p" by (meson fine_division_exists) then obtain p where p: "\z. p z tagged_division_of cbox a b" "\z. (\x. \{\ i x |i. i \ {0..z}}) fine p z" by meson have dp: "\i n. i\n \ \ i fine p n" using p unfolding fine_Inter using atLeastAtMost_iff by blast have "Cauchy (\n. sum (\(x,K). content K *\<^sub>R (f x)) (p n))" proof (rule CauchyI) fix e::real assume "0 < e" then obtain N where "N \ 0" and N: "inverse (real N) < e" using real_arch_inverse[of e] by blast show "\M. \m\M. \n\M. norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" proof (intro exI allI impI) fix m n assume mn: "N \ m" "N \ n" have "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < 1 / (real N + 1)" by (simp add: p(1) dp mn \) also have "... < e" using N \N \ 0\ \0 < e\ by (auto simp: field_simps) finally show "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" . qed qed then obtain y where y: "\no. \n\no. norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < r" if "r > 0" for r by (auto simp: convergent_eq_Cauchy[symmetric] dest: LIMSEQ_D) show ?l unfolding integrable_on_def has_integral proof (rule_tac x=y in exI, clarify) fix e :: real assume "e>0" then have e2: "e/2 > 0" by auto then obtain N1::nat where N1: "N1 \ 0" "inverse (real N1) < e/2" using real_arch_inverse by blast obtain N2::nat where N2: "\n. n \ N2 \ norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < e/2" using y[OF e2] by metis show "\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" proof (intro exI conjI allI impI) show "gauge (\ (N1+N2))" using \ by auto show "norm ((\(x,K) \ q. content K *\<^sub>R f x) - y) < e" if "q tagged_division_of cbox a b \ \ (N1+N2) fine q" for q proof (rule norm_triangle_half_r) have "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < 1 / (real (N1+N2) + 1)" by (rule \; simp add: dp p that) also have "... < e/2" using N1 \0 < e\ by (auto simp: field_simps intro: less_le_trans) finally show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < e/2" . show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - y) < e/2" using N2 le_add_same_cancel2 by blast qed qed qed qed subsection \Additivity of integral on abutting intervals\ lemma tagged_division_split_left_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. x \ k \ c}) = {}" by (metis tagged_division_split_left_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed lemma tagged_division_split_right_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. c \ x \ k}) = {}" by (metis tagged_division_split_right_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed proposition has_integral_split: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes fi: "(f has_integral i) (cbox a b \ {x. x\k \ c})" and fj: "(f has_integral j) (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" shows "(f has_integral (i + j)) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "0 < e" then have e: "e/2 > 0" by auto obtain \1 where \1: "gauge \1" and \1norm: "\\. \\ tagged_division_of cbox a b \ {x. x \ k \ c}; \1 fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - i) < e/2" apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done obtain \2 where \2: "gauge \2" and \2norm: "\\. \\ tagged_division_of cbox a b \ {x. c \ x \ k}; \2 fine \\ \ norm ((\(x, k) \ \. content k *\<^sub>R f x) - j) < e/2" apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done let ?\ = "\x. if x\k = c then (\1 x \ \2 x) else ball x \x\k - c\ \ \1 x \ \2 x" have "gauge ?\" using \1 \2 unfolding gauge_def by auto then show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - (i + j)) < e)" proof (rule_tac x="?\" in exI, safe) fix p assume p: "p tagged_division_of (cbox a b)" and "?\ fine p" have ab_eqp: "cbox a b = \{K. \x. (x, K) \ p}" using p by blast have xk_le_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have xk_ge_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have fin_finite: "finite {(x,f K) | x K. (x,K) \ s \ P x K}" if "finite s" for s and f :: "'a set \ 'a set" and P :: "'a \ 'a set \ bool" proof - from that have "finite ((\(x,K). (x, f K)) ` s)" by auto then show ?thesis by (rule rev_finite_subset) auto qed { fix \ :: "'a set \ 'a set" fix i :: "'a \ 'a set" assume "i \ (\(x, k). (x, \ k)) ` p - {(x, \ k) |x k. (x, k) \ p \ \ k \ {}}" then obtain x K where xk: "i = (x, \ K)" "(x,K) \ p" "(x, \ K) \ {(x, \ K) |x K. (x,K) \ p \ \ K \ {}}" by auto have "content (\ K) = 0" using xk using content_empty by auto then have "(\(x,K). content K *\<^sub>R f x) i = 0" unfolding xk split_conv by auto } note [simp] = this have "finite p" using p by blast let ?M1 = "{(x, K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \1_fine: "\1 fine ?M1" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" proof (rule \1norm [OF tagged_division_ofI \1_fine]) show "finite ?M1" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M1" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_le_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M1" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed moreover let ?M2 = "{(x,K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \2_fine: "\2 fine ?M2" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" proof (rule \2norm [OF tagged_division_ofI \2_fine]) show "finite ?M2" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M2" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_ge_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M2" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed ultimately have "norm (((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j)) < e/2 + e/2" using norm_add_less by blast moreover have "((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j) = (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" proof - have eq0: "\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" by auto have cont_eq: "\g. (\(x,l). content l *\<^sub>R f x) \ (\(x,l). (x,g l)) = (\(x,l). content (g l) *\<^sub>R f x)" by auto have *: "\\ :: 'a set \ 'a set. (\(x,K)\{(x, \ K) |x K. (x,K) \ p \ \ K \ {}}. content K *\<^sub>R f x) = (\(x,K)\(\(x,K). (x, \ K)) ` p. content K *\<^sub>R f x)" by (rule sum.mono_neutral_left) (auto simp: \finite p\) have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" by auto moreover have "\ = (\(x,K) \ p. content (K \ {x. x \ k \ c}) *\<^sub>R f x) + (\(x,K) \ p. content (K \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" unfolding * apply (subst (1 2) sum.reindex_nontrivial) apply (auto intro!: k p eq0 tagged_division_split_left_inj_content tagged_division_split_right_inj_content simp: cont_eq \finite p\) done moreover have "\x. x \ p \ (\(a,B). content (B \ {a. a \ k \ c}) *\<^sub>R f a) x + (\(a,B). content (B \ {a. c \ a \ k}) *\<^sub>R f a) x = (\(a,B). content B *\<^sub>R f a) x" proof clarify fix a B assume "(a, B) \ p" with p obtain u v where uv: "B = cbox u v" by blast then show "content (B \ {x. x \ k \ c}) *\<^sub>R f a + content (B \ {x. c \ x \ k}) *\<^sub>R f a = content B *\<^sub>R f a" by (auto simp: scaleR_left_distrib uv content_split[OF k,of u v c]) qed ultimately show ?thesis by (auto simp: sum.distrib[symmetric]) qed ultimately show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed subsection \A sort of converse, integrability on subintervals\ lemma has_integral_separate_sides: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes f: "(f has_integral i) (cbox a b)" and "e > 0" and k: "k \ Basis" obtains d where "gauge d" "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ norm ((sum (\(x,k). content k *\<^sub>R f x) p1 + sum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" proof - obtain \ where d: "gauge \" "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e" using has_integralD[OF f \e > 0\] by metis { fix p1 p2 assume tdiv1: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" and "\ fine p1" note p1=tagged_division_ofD[OF this(1)] assume tdiv2: "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" and "\ fine p2" note p2=tagged_division_ofD[OF this(1)] note tagged_division_Un_interval[OF tdiv1 tdiv2] note p12 = tagged_division_ofD[OF this] this { fix a b assume ab: "(a, b) \ p1 \ p2" have "(a, b) \ p1" using ab by auto obtain u v where uv: "b = cbox u v" using \(a, b) \ p1\ p1(4) by moura have "b \ {x. x\k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce moreover have "interior {x::'a. x \ k = c} = {}" proof (rule ccontr) assume "\ ?thesis" then obtain x where x: "x \ interior {x::'a. x\k = c}" by auto then obtain \ where "0 < \" and \: "ball x \ \ {x. x \ k = c}" using mem_interior by metis have x: "x\k = c" using x interior_subset by fastforce have *: "\i. i \ Basis \ \(x - (x + (\/2) *\<^sub>R k)) \ i\ = (if i = k then \/2 else 0)" using \0 < \\ k by (auto simp: inner_simps inner_not_same_Basis) have "(\i\Basis. \(x - (x + (\/2 ) *\<^sub>R k)) \ i\) = (\i\Basis. (if i = k then \/2 else 0))" using "*" by (blast intro: sum.cong) also have "\ < \" by (subst sum.delta) (use \0 < \\ in auto) finally have "x + (\/2) *\<^sub>R k \ ball x \" unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) then have "x + (\/2) *\<^sub>R k \ {x. x\k = c}" using \ by auto then show False using \0 < \\ x k by (auto simp: inner_simps) qed ultimately have "content b = 0" unfolding uv content_eq_0_interior using interior_mono by blast then have "content b *\<^sub>R f a = 0" by auto } then have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" by (subst sum.union_inter_neutral) (auto simp: p1 p2) also have "\ < e" using d(2) p12 by (simp add: fine_Un k \\ fine p1\ \\ fine p2\) finally have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . } then show ?thesis using d(1) that by auto qed lemma integrable_split [intro]: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on cbox a b" and k: "k \ Basis" shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis1) and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis2) proof - obtain y where y: "(f has_integral y) (cbox a b)" using f by blast define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a' b]) fix p assume "p tagged_division_of cbox a' b" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p1 p] d[of p2 p]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis1 by (simp add: interval_split[OF k] integrable_Cauchy) have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a b']) fix p assume "p tagged_division_of cbox a b'" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p p1] d[of p p2]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis2 by (simp add: interval_split[OF k] integrable_Cauchy) qed lemma operative_integralI: fixes f :: "'a::euclidean_space \ 'b::banach" shows "operative (lift_option (+)) (Some 0) (\i. if f integrable_on i then Some (integral i f) else None)" proof - interpret comm_monoid "lift_option plus" "Some (0::'b)" by (rule comm_monoid_lift_option) (rule add.comm_monoid_axioms) show ?thesis proof fix a b c fix k :: 'a assume k: "k \ Basis" show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = lift_option (+) (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" proof (cases "f integrable_on cbox a b") case True with k show ?thesis by (auto simp: integrable_split intro: integral_unique [OF has_integral_split[OF _ _ k]]) next case False have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" proof (rule ccontr) assume "\ ?thesis" then have "f integrable_on cbox a b" unfolding integrable_on_def apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) apply (auto intro: has_integral_split[OF _ _ k]) done then show False using False by auto qed then show ?thesis using False by auto qed next fix a b :: 'a assume "box a b = {}" then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" using has_integral_null_eq by (auto simp: integrable_on_null content_eq_0_interior) qed qed subsection \Bounds on the norm of Riemann sums and the integral itself\ lemma dsum_bound: assumes p: "p division_of (cbox a b)" and "norm c \ e" shows "norm (sum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" proof - have sumeq: "(\i\p. \content i\) = sum content p" by simp have e: "0 \ e" using assms(2) norm_ge_zero order_trans by blast have "norm (sum (\l. content l *\<^sub>R c) p) \ (\i\p. norm (content i *\<^sub>R c))" using norm_sum by blast also have "... \ e * (\i\p. \content i\)" by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg) also have "... \ e * content (cbox a b)" by (metis additive_content_division p eq_iff sumeq) finally show ?thesis . qed lemma rsum_bound: assumes p: "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" proof (cases "cbox a b = {}") case True show ?thesis using p unfolding True tagged_division_of_trivial by auto next case False then have e: "e \ 0" by (meson ex_in_conv assms(2) norm_ge_zero order_trans) have sum_le: "sum (content \ snd) p \ content (cbox a b)" unfolding additive_content_tagged_division[OF p, symmetric] split_def by (auto intro: eq_refl) have con: "\xk. xk \ p \ 0 \ content (snd xk)" using tagged_division_ofD(4) [OF p] content_pos_le by force have "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ (\i\p. norm (case i of (x, k) \ content k *\<^sub>R f x))" by (rule norm_sum) also have "... \ e * content (cbox a b)" proof - have "\xk. xk \ p \ norm (f (fst xk)) \ e" using assms(2) p tag_in_interval by force moreover have "(\i\p. \content (snd i)\ * e) \ e * content (cbox a b)" unfolding sum_distrib_right[symmetric] using con sum_le by (auto simp: mult.commute intro: mult_left_mono [OF _ e]) ultimately show ?thesis unfolding split_def norm_scaleR by (metis (no_types, lifting) mult_left_mono[OF _ abs_ge_zero] order_trans[OF sum_mono]) qed finally show ?thesis . qed lemma rsum_diff_bound: assumes "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x - g x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p - sum (\(x,k). content k *\<^sub>R g x) p) \ e * content (cbox a b)" using order_trans[OF _ rsum_bound[OF assms]] by (simp add: split_def scaleR_diff_right sum_subtractf eq_refl) lemma has_integral_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and f: "(f has_integral i) (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof (rule ccontr) assume "\ ?thesis" then have "norm i - B * content (cbox a b) > 0" by auto with f[unfolded has_integral] obtain \ where "gauge \" and \: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, K)\p. content K *\<^sub>R f x) - i) < norm i - B * content (cbox a b)" by metis then obtain p where p: "p tagged_division_of cbox a b" and "\ fine p" using fine_division_exists by blast have "\s B. norm s \ B \ \ norm (s - i) < norm i - B" unfolding not_less by (metis diff_left_mono dist_commute dist_norm norm_triangle_ineq2 order_trans) then show False using \ [OF p \\ fine p\] rsum_bound[OF p] assms by metis qed corollary integrable_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and "f integrable_on (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" by (metis integrable_integral has_integral_bound assms) subsection \Similar theorems about relationship among components\ lemma rsum_component_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes p: "p tagged_division_of (cbox a b)" and "\x. x \ cbox a b \ (f x)\i \ (g x)\i" shows "(\(x, K)\p. content K *\<^sub>R f x) \ i \ (\(x, K)\p. content K *\<^sub>R g x) \ i" unfolding inner_sum_left proof (rule sum_mono, clarify) fix x K assume ab: "(x, K) \ p" with p obtain u v where K: "K = cbox u v" by blast then show "(content K *\<^sub>R f x) \ i \ (content K *\<^sub>R g x) \ i" by (metis ab assms inner_scaleR_left measure_nonneg mult_left_mono tag_in_interval) qed lemma has_integral_component_le: fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" assumes "(f has_integral i) S" "(g has_integral j) S" and f_le_g: "\x. x \ S \ (f x)\k \ (g x)\k" shows "i\k \ j\k" proof - have ik_le_jk: "i\k \ j\k" if f_i: "(f has_integral i) (cbox a b)" and g_j: "(g has_integral j) (cbox a b)" and le: "\x\cbox a b. (f x)\k \ (g x)\k" for a b i and j :: 'b and f g :: "'a \ 'b" proof (rule ccontr) assume "\ ?thesis" then have *: "0 < (i\k - j\k) / 3" by auto obtain \1 where "gauge \1" and \1: "\p. \p tagged_division_of cbox a b; \1 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < (i \ k - j \ k) / 3" using f_i[unfolded has_integral,rule_format,OF *] by fastforce obtain \2 where "gauge \2" and \2: "\p. \p tagged_division_of cbox a b; \2 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R g x) - j) < (i \ k - j \ k) / 3" using g_j[unfolded has_integral,rule_format,OF *] by fastforce obtain p where p: "p tagged_division_of cbox a b" and "\1 fine p" "\2 fine p" using fine_division_exists[OF gauge_Int[OF \gauge \1\ \gauge \2\], of a b] unfolding fine_Int by metis then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" using le_less_trans[OF Basis_le_norm[OF k]] k \1 \2 by metis+ then show False unfolding inner_simps using rsum_component_le[OF p] le by (fastforce simp add: abs_real_def split: if_split_asm) qed show ?thesis proof (cases "\a b. S = cbox a b") case True with ik_le_jk assms show ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ i\k \ j\k" then have ij: "(i\k - j\k) / 3 > 0" by auto obtain B1 where "0 < B1" and B1: "\a b. ball 0 B1 \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast obtain B2 where "0 < B2" and B2: "\a b. ball 0 B2 \ cbox a b \ \z. ((\x. if x \ S then g x else 0) has_integral z) (cbox a b) \ norm (z - j) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ from bounded_subset_cbox_symmetric[OF this] obtain a b::'a where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (meson Un_subset_iff) then obtain w1 w2 where int_w1: "((\x. if x \ S then f x else 0) has_integral w1) (cbox a b)" and norm_w1: "norm (w1 - i) < (i \ k - j \ k) / 3" and int_w2: "((\x. if x \ S then g x else 0) has_integral w2) (cbox a b)" and norm_w2: "norm (w2 - j) < (i \ k - j \ k) / 3" using B1 B2 by blast have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" by (simp add: abs_real_def split: if_split_asm) have "\(w1 - i) \ k\ < (i \ k - j \ k) / 3" "\(w2 - j) \ k\ < (i \ k - j \ k) / 3" using Basis_le_norm k le_less_trans norm_w1 norm_w2 by blast+ moreover have "w1\k \ w2\k" using ik_le_jk int_w1 int_w2 f_le_g by auto ultimately show False unfolding inner_simps by(rule *) qed qed qed lemma integral_component_le: fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "f integrable_on S" "g integrable_on S" and "\x. x \ S \ (f x)\k \ (g x)\k" shows "(integral S f)\k \ (integral S g)\k" using has_integral_component_le assms by blast lemma has_integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ i\k" using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] using assms(3-) by auto lemma integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ (integral S f)\k" proof (cases "f integrable_on S") case True show ?thesis using True assms has_integral_component_nonneg by blast next case False then show ?thesis by (simp add: not_integrable_integral) qed lemma has_integral_component_neg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ (f x)\k \ 0" shows "i\k \ 0" using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) by auto lemma has_integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ i\k" using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) by (auto simp add: field_simps) lemma has_integral_component_ubound: fixes f::"'a::euclidean_space => 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "i\k \ B * content (cbox a b)" using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) by (auto simp add: field_simps) lemma integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" using assms has_integral_component_lbound by blast lemma integral_component_lbound_real: assumes "f integrable_on {a ::real..b}" and "\x\{a..b}. B \ f(x)\k" and "k \ Basis" shows "B * content {a..b} \ (integral {a..b} f)\k" using assms by (metis box_real(2) integral_component_lbound) lemma integral_component_ubound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" using assms has_integral_component_ubound by blast lemma integral_component_ubound_real: fixes f :: "real \ 'a::euclidean_space" assumes "f integrable_on {a..b}" and "\x\{a..b}. f x\k \ B" and "k \ Basis" shows "(integral {a..b} f)\k \ B * content {a..b}" using assms by (metis box_real(2) integral_component_ubound) subsection \Uniform limit of integrable functions is integrable\ lemma real_arch_invD: "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" by (subst(asm) real_arch_inverse) lemma integrable_uniform_limit: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "\e. e > 0 \ \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" shows "f integrable_on cbox a b" proof (cases "content (cbox a b) > 0") case False then show ?thesis using has_integral_null by (simp add: content_lt_nz integrable_on_def) next case True have "1 / (real n + 1) > 0" for n by auto then have "\g. (\x\cbox a b. norm (f x - g x) \ 1 / (real n + 1)) \ g integrable_on cbox a b" for n using assms by blast then obtain g where g_near_f: "\n x. x \ cbox a b \ norm (f x - g n x) \ 1 / (real n + 1)" and int_g: "\n. g n integrable_on cbox a b" by metis then obtain h where h: "\n. (g n has_integral h n) (cbox a b)" unfolding integrable_on_def by metis have "Cauchy h" unfolding Cauchy_def proof clarify fix e :: real assume "e>0" then have "e/4 / content (cbox a b) > 0" using True by (auto simp: field_simps) then obtain M where "M \ 0" and M: "1 / (real M) < e/4 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) show "\M. \m\M. \n\M. dist (h m) (h n) < e" proof (rule exI [where x=M], clarify) fix m n assume m: "M \ m" and n: "M \ n" have "e/4>0" using \e>0\ by auto then obtain gm gn where "gauge gm" "gauge gn" and gm: "\\. \ tagged_division_of cbox a b \ gm fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g m x) - h m) < e/4" and gn: "\\. \ tagged_division_of cbox a b \ gn fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g n x) - h n) < e/4" using h[unfolded has_integral] by meson then obtain \ where \: "\ tagged_division_of cbox a b" "(\x. gm x \ gn x) fine \" by (metis (full_types) fine_division_exists gauge_Int) have triangle3: "norm (i1 - i2) < e" if no: "norm(s2 - s1) \ e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" for s1 s2 i1 and i2::'b proof - have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed have finep: "gm fine \" "gn fine \" using fine_Int \ by auto have norm_le: "norm (g n x - g m x) \ 2 / real M" if x: "x \ cbox a b" for x proof - have "norm (f x - g n x) + norm (f x - g m x) \ 1 / (real n + 1) + 1 / (real m + 1)" using g_near_f[OF x, of n] g_near_f[OF x, of m] by simp also have "\ \ 1 / (real M) + 1 / (real M)" using \M \ 0\ m n by (intro add_mono; force simp: field_split_simps) also have "\ = 2 / real M" by auto finally show "norm (g n x - g m x) \ 2 / real M" using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] by (auto simp: algebra_simps simp add: norm_minus_commute) qed have "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ 2 / real M * content (cbox a b)" by (blast intro: norm_le rsum_diff_bound[OF \(1), where e="2 / real M"]) also have "... \ e/2" using M True by (auto simp: field_simps) finally have le_e2: "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ e/2" . then show "dist (h m) (h n) < e" unfolding dist_norm using gm gn \ finep by (auto intro!: triangle3) qed qed then obtain s where s: "h \ s" using convergent_eq_Cauchy[symmetric] by blast show ?thesis unfolding integrable_on_def has_integral proof (rule_tac x=s in exI, clarify) fix e::real assume e: "0 < e" then have "e/3 > 0" by auto then obtain N1 where N1: "\n\N1. norm (h n - s) < e/3" using LIMSEQ_D [OF s] by metis from e True have "e/3 / content (cbox a b) > 0" by (auto simp: field_simps) then obtain N2 :: nat where "N2 \ 0" and N2: "1 / (real N2) < e/3 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) obtain g' where "gauge g'" and g': "\\. \ tagged_division_of cbox a b \ g' fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" by (metis h has_integral \e/3 > 0\) have *: "norm (sf - s) < e" if no: "norm (sf - sg) \ e/3" "norm(h - s) < e/3" "norm (sg - h) < e/3" for sf sg h proof - have "norm (sf - s) \ norm (sf - sg) + norm (sg - h) + norm (h - s)" using norm_triangle_ineq[of "sf - sg" "sg - s"] using norm_triangle_ineq[of "sg - h" " h - s"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed { fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "g' fine \" then have norm_less: "norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" using g' by blast have "content (cbox a b) < e/3 * (of_nat N2)" using \N2 \ 0\ N2 using True by (auto simp: field_split_simps) moreover have "e/3 * of_nat N2 \ e/3 * (of_nat (N1 + N2) + 1)" using \e>0\ by auto ultimately have "content (cbox a b) < e/3 * (of_nat (N1 + N2) + 1)" by linarith then have le_e3: "1 / (real (N1 + N2) + 1) * content (cbox a b) \ e/3" unfolding inverse_eq_divide by (auto simp: field_simps) have ne3: "norm (h (N1 + N2) - s) < e/3" using N1 by auto have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - (\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x)) \ 1 / (real (N1 + N2) + 1) * content (cbox a b)" by (blast intro: g_near_f rsum_diff_bound[OF ptag]) then have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e" by (rule *[OF order_trans [OF _ le_e3] ne3 norm_less]) } then show "\d. gauge d \ (\\. \ tagged_division_of cbox a b \ d fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e)" by (blast intro: g' \gauge g'\) qed qed lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] subsection \Negligible sets\ definition "negligible (s:: 'a::euclidean_space set) \ (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" subsubsection \Negligibility of hyperplane\ lemma content_doublesplit: fixes a :: "'a::euclidean_space" assumes "0 < e" and k: "k \ Basis" obtains d where "0 < d" and "content (cbox a b \ {x. \x\k - c\ \ d}) < e" proof cases assume *: "a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j)" define a' where "a' d = (\j\Basis. (if j = k then max (a\j) (c - d) else a\j) *\<^sub>R j)" for d define b' where "b' d = (\j\Basis. (if j = k then min (b\j) (c + d) else b\j) *\<^sub>R j)" for d have "((\d. \j\Basis. (b' d - a' d) \ j) \ (\j\Basis. (b' 0 - a' 0) \ j)) (at_right 0)" by (auto simp: b'_def a'_def intro!: tendsto_min tendsto_max tendsto_eq_intros) also have "(\j\Basis. (b' 0 - a' 0) \ j) = 0" using k * by (intro prod_zero bexI[OF _ k]) (auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong) also have "((\d. \j\Basis. (b' d - a' d) \ j) \ 0) (at_right 0) = ((\d. content (cbox a b \ {x. \x\k - c\ \ d})) \ 0) (at_right 0)" proof (intro tendsto_cong eventually_at_rightI) fix d :: real assume d: "d \ {0<..<1}" have "cbox a b \ {x. \x\k - c\ \ d} = cbox (a' d) (b' d)" for d using * d k by (auto simp add: cbox_def set_eq_iff Int_def ball_conj_distrib abs_diff_le_iff a'_def b'_def) moreover have "j \ Basis \ a' d \ j \ b' d \ j" for j using * d k by (auto simp: a'_def b'_def) ultimately show "(\j\Basis. (b' d - a' d) \ j) = content (cbox a b \ {x. \x\k - c\ \ d})" by simp qed simp finally have "((\d. content (cbox a b \ {x. \x \ k - c\ \ d})) \ 0) (at_right 0)" . from order_tendstoD(2)[OF this \0] obtain d' where "0 < d'" and d': "\y. y > 0 \ y < d' \ content (cbox a b \ {x. \x \ k - c\ \ y}) < e" by (subst (asm) eventually_at_right[of _ 1]) auto show ?thesis by (rule that[of "d'/2"], insert \0 d'[of "d'/2"], auto) next assume *: "\ (a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j))" then have "(\j\Basis. b \ j < a \ j) \ (c < a \ k \ b \ k < c)" by (auto simp: not_le) show thesis proof cases assume "\j\Basis. b \ j < a \ j" then have [simp]: "cbox a b = {}" using box_ne_empty(1)[of a b] by auto show ?thesis by (rule that[of 1]) (simp_all add: \0) next assume "\ (\j\Basis. b \ j < a \ j)" with * have "c < a \ k \ b \ k < c" by auto then show thesis proof assume c: "c < a \ k" moreover have "x \ cbox a b \ c \ x \ k" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (a \ k - c)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(a \ k - c)/2"] show ?thesis by auto next assume c: "b \ k < c" moreover have "x \ cbox a b \ x \ k \ c" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (c - b \ k)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(c - b \ k)/2"] show ?thesis by auto qed qed qed proposition negligible_standard_hyperplane[intro]: fixes k :: "'a::euclidean_space" assumes k: "k \ Basis" shows "negligible {x. x\k = c}" unfolding negligible_def has_integral proof clarsimp fix a b and e::real assume "e > 0" with k obtain d where "0 < d" and d: "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" by (metis content_doublesplit) let ?i = "indicator {x::'a. x\k = c} :: 'a\real" show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K * ?i x\ < e)" proof (intro exI, safe) show "gauge (\x. ball x d)" using \0 < d\ by blast next fix \ assume p: "\ tagged_division_of (cbox a b)" "(\x. ball x d) fine \" have "content L = content (L \ {x. \x \ k - c\ \ d})" if "(x, L) \ \" "?i x \ 0" for x L proof - have xk: "x\k = c" using that by (simp add: indicator_def split: if_split_asm) have "L \ {x. \x \ k - c\ \ d}" proof fix y assume y: "y \ L" have "L \ ball x d" using p(2) that(1) by auto then have "norm (x - y) < d" by (simp add: dist_norm subset_iff y) then have "\(x - y) \ k\ < d" using k norm_bound_Basis_lt by blast then show "y \ {x. \x \ k - c\ \ d}" unfolding inner_simps xk by auto qed then show "content L = content (L \ {x. \x \ k - c\ \ d})" by (metis inf.orderE) qed then have *: "(\(x,K)\\. content K * ?i x) = (\(x,K)\\. content (K \ {x. \x\k - c\ \ d}) *\<^sub>R ?i x)" by (force simp add: split_paired_all intro!: sum.cong [OF refl]) note p'= tagged_division_ofD[OF p(1)] and p''=division_of_tagged_division[OF p(1)] have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * indicator {x. x \ k = c} x) < e" proof - have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) \ (\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}))" by (force simp add: indicator_def intro!: sum_mono) also have "\ < e" proof (subst sum.over_tagged_division_lemma[OF p(1)]) fix u v::'a assume "box u v = {}" then have *: "content (cbox u v) = 0" unfolding content_eq_0_interior by simp have "cbox u v \ {x. \x \ k - c\ \ d} \ cbox u v" by auto then have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" unfolding interval_doublesplit[OF k] by (rule content_subset) then show "content (cbox u v \ {x. \x \ k - c\ \ d}) = 0" unfolding * interval_doublesplit[OF k] by (blast intro: antisym) next have "(\l\snd ` \. content (l \ {x. \x \ k - c\ \ d})) = sum content ((\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}})" proof (subst (2) sum.reindex_nontrivial) fix x y assume "x \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "y \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "x \ y" and eq: "x \ {x. \x \ k - c\ \ d} = y \ {x. \x \ k - c\ \ d}" then obtain x' y' where "(x', x) \ \" "x \ {x. \x \ k - c\ \ d} \ {}" "(y', y) \ \" "y \ {x. \x \ k - c\ \ d} \ {}" by (auto) from p'(5)[OF \(x', x) \ \\ \(y', y) \ \\] \x \ y\ have "interior (x \ y) = {}" by auto moreover have "interior ((x \ {x. \x \ k - c\ \ d}) \ (y \ {x. \x \ k - c\ \ d})) \ interior (x \ y)" by (auto intro: interior_mono) ultimately have "interior (x \ {x. \x \ k - c\ \ d}) = {}" by (auto simp: eq) then show "content (x \ {x. \x \ k - c\ \ d}) = 0" using p'(4)[OF \(x', x) \ \\] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) qed (insert p'(1), auto intro!: sum.mono_neutral_right) also have "\ \ norm (\l\(\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}. content l *\<^sub>R 1::real)" by simp also have "\ \ 1 * content (cbox a b \ {x. \x \ k - c\ \ d})" using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto also have "\ < e" using d by simp finally show "(\K\snd ` \. content (K \ {x. \x \ k - c\ \ d})) < e" . qed finally show "(\(x, K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) < e" . qed then show "\\(x, K)\\. content K * ?i x\ < e" unfolding * by (simp add: sum_nonneg split: prod.split) qed qed corollary negligible_standard_hyperplane_cart: fixes k :: "'a::finite" shows "negligible {x. x$k = (0::real)}" by (simp add: cart_eq_inner_axis negligible_standard_hyperplane) subsubsection \Hence the main theorem about negligible sets\ lemma has_integral_negligible_cbox: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and 0: "\x. x \ S \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "e > 0" then have nn_gt0: "e/2 / ((real n+1) * (2 ^ n)) > 0" for n by simp then have "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n))" for n using negs [unfolded negligible_def has_integral] by auto then obtain \ where gd: "\n. gauge (\ n)" and \: "\n \. \\ tagged_division_of cbox a b; \ n fine \\ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n)" by metis show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e)" proof (intro exI, safe) show "gauge (\x. \ (nat \norm (f x)\) x)" using gd by (auto simp: gauge_def) show "norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e" if "\ tagged_division_of (cbox a b)" "(\x. \ (nat \norm (f x)\) x) fine \" for \ proof (cases "\ = {}") case True with \0 < e\ show ?thesis by simp next case False obtain N where "Max ((\(x, K). norm (f x)) ` \) \ real N" using real_arch_simple by blast then have N: "\x. x \ (\(x, K). norm (f x)) ` \ \ x \ real N" by (meson Max_ge that(1) dual_order.trans finite_imageI tagged_division_of_finite) have "\i. \q. q tagged_division_of (cbox a b) \ (\ i) fine q \ (\(x,K) \ \. K \ (\ i) x \ (x, K) \ q)" by (auto intro: tagged_division_finer[OF that(1) gd]) from choice[OF this] obtain q where q: "\n. q n tagged_division_of cbox a b" "\n. \ n fine q n" "\n x K. \(x, K) \ \; K \ \ n x\ \ (x, K) \ q n" by fastforce have "finite \" using that(1) by blast then have sum_le_inc: "\finite T; \x y. (x,y) \ T \ (0::real) \ g(x,y); \y. y\\ \ \x. (x,y) \ T \ f(y) \ g(x,y)\ \ sum f \ \ sum g T" for f g T by (rule sum_le_included[of \ T g snd f]; force) have "norm (\(x,K) \ \. content K *\<^sub>R f x) \ (\(x,K) \ \. norm (content K *\<^sub>R f x))" unfolding split_def by (rule norm_sum) also have "... \ (\(i, j) \ Sigma {..N + 1} q. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" proof (rule sum_le_inc, safe) show "finite (Sigma {..N+1} q)" by (meson finite_SigmaI finite_atMost tagged_division_of_finite q(1)) next fix x K assume xk: "(x, K) \ \" define n where "n = nat \norm (f x)\" have *: "norm (f x) \ (\(x, K). norm (f x)) ` \" using xk by auto have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" unfolding n_def by auto then have "n \ {0..N + 1}" using N[OF *] by auto moreover have "K \ \ (nat \norm (f x)\) x" using that(2) xk by auto moreover then have "(x, K) \ q (nat \norm (f x)\)" by (simp add: q(3) xk) moreover then have "(x, K) \ q n" using n_def by blast moreover have "norm (content K *\<^sub>R f x) \ (real n + 1) * (content K * indicator S x)" proof (cases "x \ S") case False then show ?thesis by (simp add: 0) next case True have *: "content K \ 0" using tagged_division_ofD(4)[OF that(1) xk] by auto moreover have "content K * norm (f x) \ content K * (real n + 1)" by (simp add: mult_left_mono nfx(2)) ultimately show ?thesis using nfx True by (auto simp: field_simps) qed ultimately show "\y. (y, x, K) \ (Sigma {..N + 1} q) \ norm (content K *\<^sub>R f x) \ (real y + 1) * (content K *\<^sub>R indicator S x)" by force qed auto also have "... = (\i\N + 1. \j\q i. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" using q(1) by (intro sum_Sigma_product [symmetric]) auto also have "... \ (\i\N + 1. (real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\)" by (rule sum_mono) (simp add: sum_distrib_left [symmetric]) also have "... \ (\i\N + 1. e/2/2 ^ i)" proof (rule sum_mono) show "(real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\ \ e/2/2 ^ i" if "i \ {..N + 1}" for i using \[of "q i" i] q by (simp add: divide_simps mult.left_commute) qed also have "... = e/2 * (\i\N + 1. (1/2) ^ i)" unfolding sum_distrib_left by (metis divide_inverse inverse_eq_divide power_one_over) also have "\ < e/2 * 2" proof (rule mult_strict_left_mono) have "sum (power (1/2)) {..N + 1} = sum (power (1/2::real)) {..0 < e\ in auto) finally show ?thesis by auto qed qed qed proposition has_integral_negligible: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and "\x. x \ (T - S) \ f x = 0" shows "(f has_integral 0) T" proof (cases "\a b. T = cbox a b") case True then have "((\x. if x \ T then f x else 0) has_integral 0) T" using assms by (auto intro!: has_integral_negligible_cbox) then show ?thesis by (rule has_integral_eq [rotated]) auto next case False let ?f = "(\x. if x \ T then f x else 0)" have "((\x. if x \ T then f x else 0) has_integral 0) T" apply (auto simp: False has_integral_alt [of ?f]) apply (rule_tac x=1 in exI, auto) apply (rule_tac x=0 in exI, simp add: has_integral_negligible_cbox [OF negs] assms) done then show ?thesis by (rule_tac f="?f" in has_integral_eq) auto qed lemma assumes "negligible S" shows integrable_negligible: "f integrable_on S" and integral_negligible: "integral S f = 0" using has_integral_negligible [OF assms] by (auto simp: has_integral_iff) lemma has_integral_spike: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" and fint: "(f has_integral y) T" shows "(g has_integral y) T" proof - have *: "(g has_integral y) (cbox a b)" if "(f has_integral y) (cbox a b)" "\x \ cbox a b - S. g x = f x" for a b f and g:: "'b \ 'a" and y proof - have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" using that by (intro has_integral_add has_integral_negligible) (auto intro!: \negligible S\) then show ?thesis by auto qed have \
: "\a b z. \\x. x \ T \ x \ S \ g x = f x; ((\x. if x \ T then f x else 0) has_integral z) (cbox a b)\ \ ((\x. if x \ T then g x else 0) has_integral z) (cbox a b)" by (auto dest!: *[where f="\x. if x\T then f x else 0" and g="\x. if x \ T then g x else 0"]) show ?thesis using fint gf apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (auto split: if_split_asm) apply (blast dest: *) using \
by meson qed lemma has_integral_spike_eq: assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" shows "(f has_integral y) T \ (g has_integral y) T" using has_integral_spike [OF \negligible S\] gf by metis lemma integrable_spike: assumes "f integrable_on T" "negligible S" "\x. x \ T - S \ g x = f x" shows "g integrable_on T" using assms unfolding integrable_on_def by (blast intro: has_integral_spike) lemma integral_spike: assumes "negligible S" and "\x. x \ T - S \ g x = f x" shows "integral T f = integral T g" using has_integral_spike_eq[OF assms] by (auto simp: integral_def integrable_on_def) subsection \Some other trivialities about negligible sets\ lemma negligible_subset: assumes "negligible s" "t \ s" shows "negligible t" unfolding negligible_def by (metis (no_types) Diff_iff assms contra_subsetD has_integral_negligible indicator_simps(2)) lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible (s - t)" using assms by (meson Diff_subset negligible_subset) lemma negligible_Int: assumes "negligible s \ negligible t" shows "negligible (s \ t)" using assms negligible_subset by force lemma negligible_Un: assumes "negligible S" and T: "negligible T" shows "negligible (S \ T)" proof - have "(indicat_real (S \ T) has_integral 0) (cbox a b)" if S0: "(indicat_real S has_integral 0) (cbox a b)" and "(indicat_real T has_integral 0) (cbox a b)" for a b proof (subst has_integral_spike_eq[OF T]) show "indicat_real S x = indicat_real (S \ T) x" if "x \ cbox a b - T" for x by (metis Diff_iff Un_iff indicator_def that) show "(indicat_real S has_integral 0) (cbox a b)" by (simp add: S0) qed with assms show ?thesis unfolding negligible_def by blast qed lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" using negligible_Un negligible_subset by blast lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] negligible_subset by blast lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" by (metis insert_is_Un negligible_Un_eq negligible_sing) lemma negligible_empty[iff]: "negligible {}" using negligible_insert by blast text\Useful in this form for backchaining\ lemma empty_imp_negligible: "S = {} \ negligible S" by simp lemma negligible_finite[intro]: assumes "finite s" shows "negligible s" using assms by (induct s) auto lemma negligible_Union[intro]: assumes "finite \" and "\t. t \ \ \ negligible t" shows "negligible(\\)" using assms by induct auto lemma negligible: "negligible S \ (\T. (indicat_real S has_integral 0) T)" proof (intro iffI allI) fix T assume "negligible S" then show "(indicator S has_integral 0) T" by (meson Diff_iff has_integral_negligible indicator_simps(2)) qed (simp add: negligible_def) subsection \Finite case of the spike theorem is quite commonly needed\ lemma has_integral_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "(f has_integral y) T" shows "(g has_integral y) T" using assms has_integral_spike negligible_finite by blast lemma has_integral_spike_finite_eq: assumes "finite S" and "\x. x \ T - S \ g x = f x" shows "((f has_integral y) T \ (g has_integral y) T)" by (metis assms has_integral_spike_finite) lemma integrable_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "f integrable_on T" shows "g integrable_on T" using assms has_integral_spike_finite by blast lemma has_integral_bound_spike_finite: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and f: "(f has_integral i) (cbox a b)" and leB: "\x. x \ cbox a b - S \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof - define g where "g \ (\x. if x \ S then 0 else f x)" then have "\x. x \ cbox a b - S \ norm (g x) \ B" using leB by simp moreover have "(g has_integral i) (cbox a b)" using has_integral_spike_finite [OF \finite S\ _ f] by (simp add: g_def) ultimately show ?thesis by (simp add: \0 \ B\ g_def has_integral_bound) qed corollary has_integral_bound_real: fixes f :: "real \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and "(f has_integral i) {a..b}" and "\x. x \ {a..b} - S \ norm (f x) \ B" shows "norm i \ B * content {a..b}" by (metis assms box_real(2) has_integral_bound_spike_finite) subsection \In particular, the boundary of an interval is negligible\ lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" proof - let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" have "negligible ?A" by (force simp add: negligible_Union[OF finite_imageI]) moreover have "cbox a b - box a b \ ?A" by (force simp add: mem_box) ultimately show ?thesis by (rule negligible_subset) qed lemma has_integral_spike_interior: assumes f: "(f has_integral y) (cbox a b)" and gf: "\x. x \ box a b \ g x = f x" shows "(g has_integral y) (cbox a b)" by (meson Diff_iff gf has_integral_spike[OF negligible_frontier_interval _ f]) lemma has_integral_spike_interior_eq: assumes "\x. x \ box a b \ g x = f x" shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" by (metis assms has_integral_spike_interior) lemma integrable_spike_interior: assumes "\x. x \ box a b \ g x = f x" and "f integrable_on cbox a b" shows "g integrable_on cbox a b" using assms has_integral_spike_interior_eq by blast subsection \Integrability of continuous functions\ lemma operative_approximableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" proof - interpret comm_monoid conj True by standard auto show ?thesis proof (standard, safe) fix a b :: 'b show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if "box a b = {}" for a b using assms that by (metis content_eq_0_interior integrable_on_null interior_cbox norm_zero right_minus_eq) { fix c g and k :: 'b assume fg: "\x\cbox a b. norm (f x - g x) \ e" and g: "g integrable_on cbox a b" assume k: "k \ Basis" show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" using fg g k by auto } show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if fg1: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" and g1: "g1 integrable_on cbox a b \ {x. x \ k \ c}" and fg2: "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" and g2: "g2 integrable_on cbox a b \ {x. c \ x \ k}" and k: "k \ Basis" for c k g1 g2 proof - let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" proof (intro exI conjI ballI) show "norm (f x - ?g x) \ e" if "x \ cbox a b" for x by (auto simp: that assms fg1 fg2) show "?g integrable_on cbox a b" proof - have "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" by(rule integrable_spike[OF _ negligible_standard_hyperplane[of k c]], use k g1 g2 in auto)+ with has_integral_split[OF _ _ k] show ?thesis unfolding integrable_on_def by blast qed qed qed qed qed lemma comm_monoid_set_F_and: "comm_monoid_set.F (\) True f s \ (finite s \ (\x\s. f x))" proof - interpret bool: comm_monoid_set \(\)\ True .. show ?thesis by (induction s rule: infinite_finite_induct) auto qed lemma approximable_on_division: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" and d: "d division_of (cbox a b)" and f: "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" proof - interpret operative conj True "\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i" using \0 \ e\ by (rule operative_approximableI) from f local.division [OF d] that show thesis by auto qed lemma integrable_continuous: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "continuous_on (cbox a b) f" shows "f integrable_on cbox a b" proof (rule integrable_uniform_limit) fix e :: real assume e: "e > 0" then obtain d where "0 < d" and d: "\x x'. \x \ cbox a b; x' \ cbox a b; dist x' x < d\ \ dist (f x') (f x) < e" using compact_uniformly_continuous[OF assms compact_cbox] unfolding uniformly_continuous_on_def by metis obtain p where ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x d) fine p" using fine_division_exists[OF gauge_ball[OF \0 < d\], of a b] . have *: "\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" proof (safe, unfold snd_conv) fix x l assume as: "(x, l) \ p" obtain a b where l: "l = cbox a b" using as ptag by blast then have x: "x \ cbox a b" using as ptag by auto show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" proof (intro exI conjI strip) show "(\y. f x) integrable_on l" unfolding integrable_on_def l by blast next fix y assume y: "y \ l" then have "y \ ball x d" using as finep by fastforce then show "norm (f y - f x) \ e" using d x y as l by (metis dist_commute dist_norm less_imp_le mem_ball ptag subsetCE tagged_division_ofD(3)) qed qed from e have "e \ 0" by auto from approximable_on_division[OF this division_of_tagged_division[OF ptag] *] show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" by metis qed lemma integrable_continuous_interval: fixes f :: "'b::ordered_euclidean_space \ 'a::banach" assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}" by (metis assms integrable_continuous interval_cbox) lemmas integrable_continuous_real = integrable_continuous_interval[where 'b=real] lemma integrable_continuous_closed_segment: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" shows "f integrable_on (closed_segment a b)" using assms by (auto intro!: integrable_continuous_interval simp: closed_segment_eq_real_ivl) subsection \Specialization of additivity to one dimension\ subsection \A useful lemma allowing us to factor out the content size\ lemma has_integral_factor_content: "(f has_integral i) (cbox a b) \ (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" proof (cases "content (cbox a b) = 0") case True have "\e p. p tagged_division_of cbox a b \ norm ((\(x, k)\p. content k *\<^sub>R f x)) \ e * content (cbox a b)" unfolding sum_content_null[OF True] True by force moreover have "i = 0" if "\e. e > 0 \ \d. gauge d \ (\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) \ e * content (cbox a b))" using that [of 1] by (force simp add: True sum_content_null[OF True] intro: fine_division_exists[of _ a b]) ultimately show ?thesis unfolding has_integral_null_eq[OF True] by (force simp add: ) next case False then have F: "0 < content (cbox a b)" using zero_less_measure_iff by blast let ?P = "\e opp. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" show ?thesis proof (subst has_integral, safe) fix e :: real assume e: "e > 0" show "?P (e * content (cbox a b)) (\)" if \
[rule_format]: "\\>0. ?P \ (<)" using \
[of "e * content (cbox a b)"] by (meson F e less_imp_le mult_pos_pos) show "?P e (<)" if \
[rule_format]: "\\>0. ?P (\ * content (cbox a b)) (\)" using \
[of "e/2 / content (cbox a b)"] using F e by (force simp add: algebra_simps) qed qed lemma has_integral_factor_content_real: "(f has_integral i) {a..b::real} \ (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a..b} ))" unfolding box_real[symmetric] by (rule has_integral_factor_content) subsection \Fundamental theorem of calculus\ lemma interval_bounds_real: fixes q b :: real assumes "a \ b" shows "Sup {a..b} = b" and "Inf {a..b} = a" using assms by auto theorem fundamental_theorem_of_calculus: fixes f :: "real \ 'a::banach" assumes "a \ b" and vecd: "\x. x \ {a..b} \ (f has_vector_derivative f' x) (at x within {a..b})" shows "(f' has_integral (f b - f a)) {a..b}" unfolding has_integral_factor_content box_real[symmetric] proof safe fix e :: real assume "e > 0" then have "\x. \d>0. x \ {a..b} \ (\y\{a..b}. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x))" using vecd[unfolded has_vector_derivative_def has_derivative_within_alt] by blast then obtain d where d: "\x. 0 < d x" "\x y. \x \ {a..b}; y \ {a..b}; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x)" by metis show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" proof (rule exI, safe) show "gauge (\x. ball x (d x))" using d(1) gauge_ball_dependent by blast next fix p assume ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x (d x)) fine p" have ba: "b - a = (\(x,K)\p. Sup K - Inf K)" "f b - f a = (\(x,K)\p. f(Sup K) - f(Inf K))" using additive_tagged_division_1[where f= "\x. x"] additive_tagged_division_1[where f= f] \a \ b\ ptag by auto have "norm (\(x, K) \ p. (content K *\<^sub>R f' x) - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, k) \ Sup k - Inf k))" proof (rule sum_norm_le,safe) fix x K assume "(x, K) \ p" then have "x \ K" and kab: "K \ cbox a b" using ptag by blast+ then obtain u v where k: "K = cbox u v" and "x \ K" and kab: "K \ cbox a b" using ptag \(x, K) \ p\ by auto have "u \ v" using \x \ K\ unfolding k by auto have ball: "\y\K. y \ ball x (d x)" using finep \(x, K) \ p\ by blast have "u \ K" "v \ K" by (simp_all add: \u \ v\ k) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm (f u - f x - (u - x) *\<^sub>R f' x - (f v - f x - (v - x) *\<^sub>R f' x))" by (auto simp add: algebra_simps) also have "... \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" by (rule norm_triangle_ineq4) finally have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" . also have "\ \ e * norm (u - x) + e * norm (v - x)" proof (rule add_mono) show "norm (f u - f x - (u - x) *\<^sub>R f' x) \ e * norm (u - x)" proof (rule d) show "norm (u - x) < d x" using \u \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \u \ K\ kab in auto) show "norm (f v - f x - (v - x) *\<^sub>R f' x) \ e * norm (v - x)" proof (rule d) show "norm (v - x) < d x" using \v \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \v \ K\ kab in auto) qed also have "\ \ e * (Sup K - Inf K)" using \x \ K\ by (auto simp: k interval_bounds_real[OF \u \ v\] field_simps) finally show "norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (Sup K - Inf K)" using \u \ v\ by (simp add: k) qed with \a \ b\ show "norm ((\(x, K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" by (auto simp: ba split_def sum_subtractf [symmetric] sum_distrib_left) qed qed lemma ident_has_integral: fixes a::real assumes "a \ b" shows "((\x. x) has_integral (b\<^sup>2 - a\<^sup>2)/2) {a..b}" proof - have "((\x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}" unfolding power2_eq_square by (rule fundamental_theorem_of_calculus [OF assms] derivative_eq_intros | simp)+ then show ?thesis by (simp add: field_simps) qed lemma integral_ident [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} (\x. x) = (if a \ b then (b\<^sup>2 - a\<^sup>2)/2 else 0)" by (metis assms ident_has_integral integral_unique) lemma ident_integrable_on: fixes a::real shows "(\x. x) integrable_on {a..b}" by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) lemma integral_sin [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} sin = cos a - cos b" proof - have "(sin has_integral (- cos b - - cos a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "((\a. - cos a) has_vector_derivative sin x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma integral_cos [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} cos = sin b - sin a" proof - have "(cos has_integral (sin b - sin a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "(sin has_vector_derivative cos x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" proof (cases "n = 0") case False have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" if "a \ {-pi..pi}" for a :: real using that False unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ qed auto then show ?thesis by simp qed auto lemma integral_sin_nx: "integral {-pi..pi} (\x. sin(x * real_of_int n)) = 0" using has_integral_sin_nx [of n] by (force simp: mult.commute) lemma has_integral_cos_nx: "((\x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}" proof (cases "n = 0") case True then show ?thesis using has_integral_const_real [of "1::real" "-pi" pi] by auto next case False have "((\x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})" if "x \ {-pi..pi}" for x :: real using that False unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ qed auto with False show ?thesis by (simp add: mult.commute) qed lemma integral_cos_nx: "integral {-pi..pi} (\x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)" using has_integral_cos_nx [of n] by (force simp: mult.commute) subsection \Taylor series expansion\ lemma mvt_integral: fixes f::"'a::real_normed_vector\'b::banach" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes line_in: "\t. t \ {0..1} \ x + t *\<^sub>R y \ S" shows "f (x + y) - f x = integral {0..1} (\t. f' (x + t *\<^sub>R y) y)" (is ?th1) proof - from assms have subset: "(\xa. x + xa *\<^sub>R y) ` {0..1} \ S" by auto note [derivative_intros] = has_derivative_subset[OF _ subset] has_derivative_in_compose[where f="(\xa. x + xa *\<^sub>R y)" and g = f] note [continuous_intros] = continuous_on_compose2[where f="(\xa. x + xa *\<^sub>R y)"] continuous_on_subset[OF _ subset] have "\t. t \ {0..1} \ ((\t. f (x + t *\<^sub>R y)) has_vector_derivative f' (x + t *\<^sub>R y) y) (at t within {0..1})" using assms by (auto simp: has_vector_derivative_def linear_cmul[OF has_derivative_linear[OF f'], symmetric] intro!: derivative_eq_intros) from fundamental_theorem_of_calculus[rule_format, OF _ this] show ?th1 by (auto intro!: integral_unique[symmetric]) qed lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative: assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and g0: "Dg 0 = g" and Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" and ivl: "a \ t" "t \ b" shows "((\t. \iR prod (Df i t) (Dg (p - Suc i) t)) has_vector_derivative prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) (at t within {a..b})" using assms proof cases assume p: "p \ 1" define p' where "p' = p - 2" from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" by auto let ?f = "\i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = (\i\(Suc p'). ?f i - ?f (Suc i))" by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) also note sum_telescope finally have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" unfolding p'[symmetric] by (simp add: assms) thus ?thesis using assms by (auto intro!: derivative_eq_intros has_vector_derivative) qed (auto intro!: derivative_eq_intros has_vector_derivative) lemma fixes f::"real\'a::banach" assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and ivl: "a \ b" defines "i \ \x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" shows Taylor_has_integral: "(i has_integral f b - (\iR Df i a)) {a..b}" and Taylor_integral: "f b = (\iR Df i a) + integral {a..b} i" and Taylor_integrable: "i integrable_on {a..b}" proof goal_cases case 1 interpret bounded_bilinear "scaleR::real\'a\'a" by (rule bounded_bilinear_scaleR) define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s define Dg where [abs_def]: "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s have g0: "Dg 0 = g" using \p > 0\ by (auto simp add: Dg_def field_split_simps g_def split: if_split_asm) { fix m assume "p > Suc m" hence "p - Suc m = Suc (p - Suc (Suc m))" by auto hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" by auto } note fact_eq = this have Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" unfolding Dg_def by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq field_split_simps) let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, OF \p > 0\ g0 Dg f0 Df] have deriv: "\t. a \ t \ t \ b \ (?sum has_vector_derivative g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})" by auto from fundamental_theorem_of_calculus[rule_format, OF \a \ b\ deriv] have "(i has_integral ?sum b - ?sum a) {a..b}" using atLeastatMost_empty'[simp del] by (simp add: i_def g_def Dg_def) also have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" and "{.. {i. p = Suc i} = {p - 1}" for p' using \p > 0\ by (auto simp: power_mult_distrib[symmetric]) then have "?sum b = f b" using Suc_pred'[OF \p > 0\] by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib if_distribR sum.If_cases f0) also have "{..x. p - x - 1) ` {.. (\x. p - x - 1) ` {..iR Df i a)" by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one) finally show c: ?case . case 2 show ?case using c integral_unique by (metis (lifting) add.commute diff_eq_eq integral_unique) case 3 show ?case using c by force qed subsection \Only need trivial subintervals if the interval itself is trivial\ proposition division_of_nontrivial: fixes \ :: "'a::euclidean_space set set" assumes sdiv: "\ division_of (cbox a b)" and cont0: "content (cbox a b) \ 0" shows "{k. k \ \ \ content k \ 0} division_of (cbox a b)" using sdiv proof (induction "card \" arbitrary: \ rule: less_induct) case less note \ = division_ofD[OF less.prems] { presume *: "{k \ \. content k \ 0} \ \ \ ?case" then show ?case using less.prems by fastforce } assume noteq: "{k \ \. content k \ 0} \ \" then obtain K c d where "K \ \" and contk: "content K = 0" and keq: "K = cbox c d" using \(4) by blast then have "card \ > 0" unfolding card_gt_0_iff using less by auto then have card: "card (\ - {K}) < card \" using less \K \ \\ by (simp add: \(1)) have closed: "closed (\(\ - {K}))" using less.prems by auto have "x islimpt \(\ - {K})" if "x \ K" for x unfolding islimpt_approachable proof (intro allI impI) fix e::real assume "e > 0" obtain i where i: "c\i = d\i" "i\Basis" using contk \(3) [OF \K \ \\] unfolding box_ne_empty keq by (meson content_eq_0 dual_order.antisym) then have xi: "x\i = d\i" using \x \ K\ unfolding keq mem_box by (metis antisym) define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i)/2 then c\i + min e (b\i - c\i)/2 else c\i - min e (c\i - a\i)/2 else x\j) *\<^sub>R j)" show "\x'\\(\ - {K}). x' \ x \ dist x' x < e" proof (intro bexI conjI) have "d \ cbox c d" using \(3)[OF \K \ \\] by (simp add: box_ne_empty(1) keq mem_box(2)) then have "d \ cbox a b" using \(2)[OF \K \ \\] by (auto simp: keq) then have di: "a \ i \ d \ i \ d \ i \ b \ i" using \i \ Basis\ mem_box(2) by blast then have xyi: "y\i \ x\i" unfolding y_def i xi using \e > 0\ cont0 \i \ Basis\ by (auto simp: content_eq_0 elim!: ballE[of _ _ i]) then show "y \ x" unfolding euclidean_eq_iff[where 'a='a] using i by auto have "norm (y-x) \ (\b\Basis. \(y - x) \ b\)" by (rule norm_le_l1) also have "... = \(y - x) \ i\ + (\b \ Basis - {i}. \(y - x) \ b\)" by (meson finite_Basis i(2) sum.remove) also have "... < e + sum (\i. 0) Basis" proof (rule add_less_le_mono) show "\(y-x) \ i\ < e" using di \e > 0\ y_def i xi by (auto simp: inner_simps) show "(\i\Basis - {i}. \(y-x) \ i\) \ (\i\Basis. 0)" unfolding y_def by (auto simp: inner_simps) qed finally have "norm (y-x) < e + sum (\i. 0) Basis" . then show "dist y x < e" unfolding dist_norm by auto have "y \ K" unfolding keq mem_box using i(1) i(2) xi xyi by fastforce moreover have "y \ \\" using subsetD[OF \(2)[OF \K \ \\] \x \ K\] \e > 0\ di i by (auto simp: \ mem_box y_def field_simps elim!: ballE[of _ _ i]) ultimately show "y \ \(\ - {K})" by auto qed qed then have "K \ \(\ - {K})" using closed closed_limpt by blast then have "\(\ - {K}) = cbox a b" unfolding \(6)[symmetric] by auto then have "\ - {K} division_of cbox a b" by (metis Diff_subset less.prems division_of_subset \(6)) then have "{ka \ \ - {K}. content ka \ 0} division_of (cbox a b)" using card less.hyps by blast moreover have "{ka \ \ - {K}. content ka \ 0} = {K \ \. content K \ 0}" using contk by auto ultimately show ?case by auto qed subsection \Integrability on subintervals\ lemma operative_integrableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. f integrable_on i)" proof - interpret comm_monoid conj True proof qed show ?thesis proof show "\a b. box a b = {} \ (f integrable_on cbox a b) = True" by (simp add: content_eq_0_interior integrable_on_null) show "\a b c k. k \ Basis \ (f integrable_on cbox a b) \ (f integrable_on cbox a b \ {x. x \ k \ c} \ f integrable_on cbox a b \ {x. c \ x \ k})" unfolding integrable_on_def by (auto intro!: has_integral_split) qed qed lemma integrable_subinterval: fixes f :: "'b::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and cd: "cbox c d \ cbox a b" shows "f integrable_on cbox c d" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) show ?thesis proof (cases "cbox c d = {}") case True then show ?thesis using division [symmetric] f by (auto simp: comm_monoid_set_F_and) next case False then show ?thesis by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1) qed qed lemma integrable_subinterval_real: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "{c..d} \ {a..b}" shows "f integrable_on {c..d}" by (metis assms box_real(2) integrable_subinterval) subsection \Combining adjacent intervals in 1 dimension\ lemma has_integral_combine: fixes a b c :: real and j :: "'a::banach" assumes "a \ c" and "c \ b" and ac: "(f has_integral i) {a..c}" and cb: "(f has_integral j) {c..b}" shows "(f has_integral (i + j)) {a..b}" proof - interpret operative_real "lift_option plus" "Some 0" "\i. if f integrable_on i then Some (integral i f) else None" using operative_integralI by (rule operative_realI) from \a \ c\ \c \ b\ ac cb coalesce_less_eq have *: "lift_option (+) (if f integrable_on {a..c} then Some (integral {a..c} f) else None) (if f integrable_on {c..b} then Some (integral {c..b} f) else None) = (if f integrable_on {a..b} then Some (integral {a..b} f) else None)" by (auto simp: split: if_split_asm) then have "f integrable_on cbox a b" using ac cb by (auto split: if_split_asm) with * show ?thesis using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm) qed lemma integral_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and ab: "f integrable_on {a..b}" shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" proof - have "(f has_integral integral {a..c} f) {a..c}" using ab \c \ b\ integrable_subinterval_real by fastforce moreover have "(f has_integral integral {c..b} f) {c..b}" using ab \a \ c\ integrable_subinterval_real by fastforce ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}" using \a \ c\ \c \ b\ has_integral_combine by blast then show ?thesis by (simp add: has_integral_integrable_integral) qed lemma integrable_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and "f integrable_on {a..c}" and "f integrable_on {c..b}" shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by (auto intro!: has_integral_combine) lemma integral_minus_sets: fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {c .. max a b} \ integral {c .. a} f - integral {c .. b} f = (if a \ b then - integral {a .. b} f else integral {b .. a} f)" using integral_combine[of c a b f] integral_combine[of c b a f] by (auto simp: algebra_simps max_def) lemma integral_minus_sets': fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {min a b .. c} \ integral {a .. c} f - integral {b .. c} f = (if a \ b then integral {a .. b} f else - integral {b .. a} f)" using integral_combine[of b a c f] integral_combine[of a b c f] by (auto simp: algebra_simps min_def) subsection \Reduce integrability to "local" integrability\ lemma integrable_on_little_subintervals: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v" shows "f integrable_on cbox a b" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) have "\x. \d>0. x\cbox a b \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v)" using assms by (metis zero_less_one) then obtain d where d: "\x. 0 < d x" "\x u v. \x \ cbox a b; x \ cbox u v; cbox u v \ ball x (d x); cbox u v \ cbox a b\ \ f integrable_on cbox u v" by metis obtain p where p: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" using fine_division_exists[OF gauge_ball_dependent,of d a b] d(1) by blast then have sndp: "snd ` p division_of cbox a b" by (metis division_of_tagged_division) have "f integrable_on k" if "(x, k) \ p" for x k using tagged_division_ofD(2-4)[OF p(1) that] fineD[OF p(2) that] d[of x] by auto then show ?thesis unfolding division [symmetric, OF sndp] comm_monoid_set_F_and by auto qed subsection \Second FTC or existence of antiderivative\ lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" unfolding integrable_on_def by blast lemma integral_has_vector_derivative_continuous_at: fixes f :: "real \ 'a::banach" assumes f: "f integrable_on {a..b}" and x: "x \ {a..b} - S" and "finite S" and fx: "continuous (at x within ({a..b} - S)) f" shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - S))" proof - let ?I = "\a b. integral {a..b} f" { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b} - S; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" (is "?lhs \ ?rhs") if y: "y \ {a..b} - S" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) then have Idiff: "?I a y - ?I a x = ?I x y" using False x by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}" proof (rule has_integral_diff) show "(f has_integral integral {x..y} f) {x..y}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (y - x) *\<^sub>R f x) {x..y}" using has_integral_const_real [of "f x" x y] False by simp qed have "?lhs \ e * content {x..y}" using yx False d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ ?rhs" using False by auto finally show ?thesis . next case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) then have Idiff: "?I a x - ?I a y = ?I y x" using True x y by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" proof (rule has_integral_diff) show "(f has_integral integral {y..x} f) {y..x}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (x - y) *\<^sub>R f x) {y..x}" using has_integral_const_real [of "f x" y x] True by simp qed have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * content {y..x}" using yx True d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ e * \y - x\" using True by auto finally have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" . then show ?thesis by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b} - S. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" using \d>0\ by blast } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed lemma integral_has_vector_derivative: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})" using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] by (fastforce simp: continuous_on_eq_continuous_within) lemma integral_has_real_derivative: assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {a..x} g) has_real_derivative g t) (at t within {a..b})" using integral_has_vector_derivative[of a b g t] assms by (auto simp: has_field_derivative_iff_has_vector_derivative) lemma antiderivative_continuous: fixes q b :: real assumes "continuous_on {a..b} f" obtains g where "\x. x \ {a..b} \ (g has_vector_derivative (f x::_::banach)) (at x within {a..b})" using integral_has_vector_derivative[OF assms] by auto subsection \Combined fundamental theorem of calculus\ lemma antiderivative_integral_continuous: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" obtains g where "\u\{a..b}. \v \ {a..b}. u \ v \ (f has_integral (g v - g u)) {u..v}" proof - obtain g where g: "\x. x \ {a..b} \ (g has_vector_derivative f x) (at x within {a..b})" using antiderivative_continuous[OF assms] by metis have "(f has_integral g v - g u) {u..v}" if "u \ {a..b}" "v \ {a..b}" "u \ v" for u v proof - have "\x. x \ cbox u v \ (g has_vector_derivative f x) (at x within cbox u v)" by (metis atLeastAtMost_iff atLeastatMost_subset_iff box_real(2) g has_vector_derivative_within_subset subsetCE that(1,2)) then show ?thesis by (metis box_real(2) that(3) fundamental_theorem_of_calculus) qed then show ?thesis using that by blast qed subsection \General "twiddling" for interval-to-interval function image\ lemma has_integral_twiddle: assumes "0 < r" and hg: "\x. h(g x) = x" and gh: "\x. g(h x) = x" and contg: "\x. continuous (at x) g" and g: "\u v. \w z. g ` cbox u v = cbox w z" and h: "\u v. \w z. h ` cbox u v = cbox w z" and r: "\u v. content(g ` cbox u v) = r * content (cbox u v)" and intfi: "(f has_integral i) (cbox a b)" shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" proof (cases "cbox a b = {}") case True then show ?thesis using intfi by auto next case False obtain w z where wz: "h ` cbox a b = cbox w z" using h by blast have inj: "inj g" "inj h" using hg gh injI by metis+ from h obtain ha hb where h_eq: "h ` cbox a b = cbox ha hb" by blast have "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" if "e > 0" for e proof - have "e * r > 0" using that \0 < r\ by simp with intfi[unfolded has_integral] obtain d where "gauge d" and d: "\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e * r" by metis define d' where "d' x = g -` d (g x)" for x show ?thesis proof (rule_tac x=d' in exI, safe) show "gauge d'" using \gauge d\ continuous_open_vimage[OF _ contg] by (auto simp: gauge_def d'_def) next fix p assume ptag: "p tagged_division_of h ` cbox a b" and finep: "d' fine p" note p = tagged_division_ofD[OF ptag] have gab: "g y \ cbox a b" if "y \ K" "(x, K) \ p" for x y K by (metis hg inj(2) inj_image_mem_iff p(3) subsetCE that that) have gimp: "(\(x,K). (g x, g ` K)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of proof safe show "finite ((\(x, k). (g x, g ` k)) ` p)" using ptag by auto show "d fine (\(x, k). (g x, g ` k)) ` p" using finep unfolding fine_def d'_def by auto next fix x k assume xk: "(x, k) \ p" show "g x \ g ` k" using p(2)[OF xk] by auto show "\u v. g ` k = cbox u v" using p(4)[OF xk] using assms(5-6) by auto fix x' K' u assume xk': "(x', K') \ p" and u: "u \ interior (g ` k)" "u \ interior (g ` K')" have "interior k \ interior K' \ {}" proof assume "interior k \ interior K' = {}" moreover have "u \ g ` (interior k \ interior K')" using interior_image_subset[OF \inj g\ contg] u unfolding image_Int[OF inj(1)] by blast ultimately show False by blast qed then have same: "(x, k) = (x', K')" using ptag xk' xk by blast then show "g x = g x'" by auto show "g u \ g ` K'"if "u \ k" for u using that same by auto show "g u \ g ` k"if "u \ K'" for u using that same by auto next fix x assume "x \ cbox a b" then have "h x \ \{k. \x. (x, k) \ p}" using p(6) by auto then obtain X y where "h x \ X" "(y, X) \ p" by blast then show "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" by clarsimp (metis (no_types, lifting) gh image_eqI pair_imageI) qed (use gab in auto) have *: "inj_on (\(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastforce have "(\(x,K)\(\(y,L). (g y, g ` L)) ` p. content K *\<^sub>R f x) = (\u\p. case case u of (x,K) \ (g x, g ` K) of (y,L) \ content L *\<^sub>R f y)" by (metis (mono_tags, lifting) "*" sum.reindex_cong) also have "... = (\(x,K)\p. r *\<^sub>R content K *\<^sub>R f (g x))" using r by (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4)) finally have "(\(x, K)\(\(x,K). (g x, g ` K)) ` p. content K *\<^sub>R f x) - i = r *\<^sub>R (\(x,K)\p. content K *\<^sub>R f (g x)) - i" by (simp add: scaleR_right.sum split_def) also have "\ = r *\<^sub>R ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" using \0 < r\ by (auto simp: scaleR_diff_right) finally show "norm ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using d[OF gimp] \0 < r\ by auto qed qed then show ?thesis by (auto simp: h_eq has_integral) qed subsection \Special case of a basic affine transformation\ lemma AE_lborel_inner_neq: assumes k: "k \ Basis" shows "AE x in lborel. x \ k \ c" proof - interpret finite_product_sigma_finite "\_. lborel" Basis proof qed simp have "emeasure lborel {x\space lborel. x \ k = c} = emeasure (\\<^sub>M j::'a\Basis. lborel) (\\<^sub>E j\Basis. if j = k then {c} else UNIV)" using k by (auto simp add: lborel_eq[where 'a='a] emeasure_distr intro!: arg_cong2[where f=emeasure]) (auto simp: space_PiM PiE_iff extensional_def split: if_split_asm) also have "\ = (\j\Basis. emeasure lborel (if j = k then {c} else UNIV))" by (intro measure_times) auto also have "\ = 0" by (intro prod_zero bexI[OF _ k]) auto finally show ?thesis by (subst AE_iff_measurable[OF _ refl]) auto qed lemma content_image_stretch_interval: fixes m :: "'a::euclidean_space \ real" defines "s f x \ (\k::'a\Basis. (f k * (x\k)) *\<^sub>R k)" shows "content (s m ` cbox a b) = \\k\Basis. m k\ * content (cbox a b)" proof cases have s[measurable]: "s f \ borel \\<^sub>M borel" for f by (auto simp: s_def[abs_def]) assume m: "\k\Basis. m k \ 0" then have s_comp_s: "s (\k. 1 / m k) \ s m = id" "s m \ s (\k. 1 / m k) = id" by (auto simp: s_def[abs_def] fun_eq_iff euclidean_representation) then have "inv (s (\k. 1 / m k)) = s m" "bij (s (\k. 1 / m k))" by (auto intro: inv_unique_comp o_bij) then have eq: "s m ` cbox a b = s (\k. 1 / m k) -` cbox a b" using bij_vimage_eq_inv_image[OF \bij (s (\k. 1 / m k))\, of "cbox a b"] by auto show ?thesis using m unfolding eq measure_def by (subst lborel_affine_euclidean[where c=m and t=0]) (simp_all add: emeasure_density measurable_sets_borel[OF s] abs_prod nn_integral_cmult s_def[symmetric] emeasure_distr vimage_comp s_comp_s enn2real_mult prod_nonneg) next assume "\ (\k\Basis. m k \ 0)" then obtain k where k: "k \ Basis" "m k = 0" by auto then have [simp]: "(\k\Basis. m k) = 0" by (intro prod_zero) auto have "emeasure lborel {x\space lborel. x \ s m ` cbox a b} = 0" proof (rule emeasure_eq_0_AE) show "AE x in lborel. x \ s m ` cbox a b" using AE_lborel_inner_neq[OF \k\Basis\] proof eventually_elim show "x \ k \ 0 \ x \ s m ` cbox a b " for x using k by (auto simp: s_def[abs_def] cbox_def) qed qed then show ?thesis by (simp add: measure_def) qed lemma interval_image_affinity_interval: "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" unfolding image_affinity_cbox by auto lemma content_image_affinity_cbox: "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = \m\ ^ DIM('a) * content (cbox a b)" (is "?l = ?r") proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "m \ 0") case True with \cbox a b \ {}\ have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis by (simp add: image_affinity_cbox True content_cbox' prod.distrib inner_diff_left) next case False with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis using False by (simp add: image_affinity_cbox content_cbox' prod.distrib[symmetric] inner_diff_left flip: prod_constant) qed qed lemma has_integral_affinity: fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m \ 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (1 / (\m\ ^ DIM('a))) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" proof (rule has_integral_twiddle) show "\w z. (\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox u v = cbox w z" "\w z. (\x. m *\<^sub>R x + c) ` cbox u v = cbox w z" for u v using interval_image_affinity_interval by blast+ show "content ((\x. m *\<^sub>R x + c) ` cbox u v) = \m\ ^ DIM('a) * content (cbox u v)" for u v using content_image_affinity_cbox by blast qed (use assms zero_less_power in \auto simp: field_simps\) lemma integrable_affinity: assumes "f integrable_on cbox a b" and "m \ 0" shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" using has_integral_affinity assms unfolding integrable_on_def by blast lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed subsection \Special case of stretching coordinate axes separately\ lemma has_integral_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "(f has_integral i) (cbox a b)" and "\k\Basis. m k \ 0" shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral ((1/ \prod m Basis\) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" apply (rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] using assms by auto lemma has_integral_stretch_real: fixes f :: "real \ 'b::real_normed_vector" assumes "(f has_integral i) {a..b}" and "m \ 0" shows "((\x. f (m * x)) has_integral (1 / \m\) *\<^sub>R i) ((\x. x / m) ` {a..b})" using has_integral_stretch [of f i a b "\b. m"] assms by simp lemma integrable_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "f integrable_on cbox a b" and "\k\Basis. m k \ 0" shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" using assms unfolding integrable_on_def by (force dest: has_integral_stretch) lemma vec_lambda_eq_sum: "(\ k. f k (x $ k)) = (\k\Basis. (f (axis_index k) (x \ k)) *\<^sub>R k)" (is "?lhs = ?rhs") proof - have "?lhs = (\ k. f k (x \ axis k 1))" by (simp add: cart_eq_inner_axis) also have "... = (\u\UNIV. f u (x \ axis u 1) *\<^sub>R axis u 1)" by (simp add: vec_eq_iff axis_def if_distrib cong: if_cong) also have "... = ?rhs" by (simp add: Basis_vec_def UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) finally show ?thesis . qed lemma has_integral_stretch_cart: fixes m :: "'n::finite \ real" assumes f: "(f has_integral i) (cbox a b)" and m: "\k. m k \ 0" shows "((\x. f(\ k. m k * x$k)) has_integral i /\<^sub>R \prod m UNIV\) ((\x. \ k. x$k / m k) ` (cbox a b))" proof - have *: "\k:: real^'n \ Basis. m (axis_index k) \ 0" using axis_index by (simp add: m) have eqp: "(\k:: real^'n \ Basis. m (axis_index k)) = prod m UNIV" by (simp add: Basis_vec_def UNION_singleton_eq_range prod.reindex axis_eq_axis inj_on_def) show ?thesis using has_integral_stretch [OF f *] vec_lambda_eq_sum [where f="\i x. m i * x"] vec_lambda_eq_sum [where f="\i x. x / m i"] by (simp add: field_simps eqp) qed lemma image_stretch_interval_cart: fixes m :: "'n::finite \ real" shows "(\x. \ k. m k * x$k) ` cbox a b = (if cbox a b = {} then {} else cbox (\ k. min (m k * a$k) (m k * b$k)) (\ k. max (m k * a$k) (m k * b$k)))" proof - have *: "(\k\Basis. min (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. min (m k * a $ k) (m k * b $ k))" "(\k\Basis. max (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. max (m k * a $ k) (m k * b $ k))" apply (simp_all add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) apply (simp_all add: vec_eq_iff axis_def if_distrib cong: if_cong) done show ?thesis by (simp add: vec_lambda_eq_sum [where f="\i x. m i * x"] image_stretch_interval eq_cbox *) qed subsection \even more special cases\ lemma uminus_interval_vector[simp]: fixes a b :: "'a::euclidean_space" shows "uminus ` cbox a b = cbox (-b) (-a)" proof - have "x \ uminus ` cbox a b" if "x \ cbox (- b) (- a)" for x proof - have "-x \ cbox a b" using that by (auto simp: mem_box) then show ?thesis by force qed then show ?thesis by (auto simp: mem_box) qed lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) (cbox a b)" shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" using has_integral_affinity[OF assms, of "-1" 0] by auto lemma has_integral_reflect_lemma_real[intro]: assumes "(f has_integral i) {a..b::real}" shows "((\x. f(-x)) has_integral i) {-b .. -a}" using assms unfolding box_real[symmetric] by (rule has_integral_reflect_lemma) lemma has_integral_reflect[simp]: "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" by (auto dest: has_integral_reflect_lemma) lemma has_integral_reflect_real[simp]: fixes a b::real shows "((\x. f (-x)) has_integral i) {-b..-a} \ (f has_integral i) {a..b}" by (metis has_integral_reflect interval_cbox) lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" unfolding integrable_on_def by auto lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule integrable_reflect) lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" unfolding integral_def by auto lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a..b::real} f" unfolding box_real[symmetric] by (rule integral_reflect) subsection \Stronger form of FCT; quite a tedious proof\ lemma split_minus[simp]: "(\(x, k). f x k) x - (\(x, k). g x k) x = (\(x, k). f x k - g x k) x" by (simp add: split_def) theorem fundamental_theorem_of_calculus_interior: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and contf: "continuous_on {a..b} f" and derf: "\x. x \ {a <..< b} \ (f has_vector_derivative f' x) (at x)" shows "(f' has_integral (f b - f a)) {a..b}" proof (cases "a = b") case True then have *: "cbox a b = {b}" "f b - f a = 0" by (auto simp add: order_antisym) with True show ?thesis by auto next case False with \a \ b\ have ab: "a < b" by arith show ?thesis unfolding has_integral_factor_content_real proof (intro allI impI) fix e :: real assume e: "e > 0" then have eba8: "(e * (b-a)) / 8 > 0" using ab by (auto simp add: field_simps) note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt, THEN conjunct2, rule_format] thm derf_exp have bounded: "\x. x \ {a<.. bounded_linear (\u. u *\<^sub>R f' x)" by (simp add: bounded_linear_scaleR_left) have "\x \ box a b. \d>0. \y. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" (is "\x \ box a b. ?Q x") \\The explicit quantifier is required by the following step\ proof fix x assume "x \ box a b" with e show "?Q x" using derf_exp [of x "e/2"] by auto qed then obtain d where d: "\x. 0 < d x" "\x y. \x \ box a b; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" unfolding bgauge_existence_lemma by metis have "bounded (f ` cbox a b)" using compact_cbox assms by (auto simp: compact_imp_bounded compact_continuous_image) then obtain B where "0 < B" and B: "\x. x \ f ` cbox a b \ norm x \ B" unfolding bounded_pos by metis obtain da where "0 < da" and da: "\c. \a \ c; {a..c} \ {a..b}; {a..c} \ ball a da\ \ norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ (e * (b-a)) / 4" proof - have "continuous (at a within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm (x-a); norm (x-a) < k\ \ norm (f x - f a) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \ e * (b-a) / 8" proof (cases "f' a = 0") case True with ab e that show ?thesis by auto next case False show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' a)) *\<^sub>R f' a) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' a)" using False ab e by (auto simp add: field_simps) qed qed have "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" if "a \ c" "{a..c} \ {a..b}" and bmin: "{a..c} \ ball a (min k l)" for c proof - have minkl: "\a - x\ < min k l" if "x \ {a..c}" for x using bmin dist_real_def that by auto then have lel: "\c - a\ \ \l\" using that by force have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((c - a) *\<^sub>R f' a) \ norm (l *\<^sub>R f' a)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((c - a) *\<^sub>R f' a) \ e * (b-a) / 8" . next have "norm (f c - f a) < e * (b-a) / 8" proof (cases "a = c") case True then show ?thesis using eba8 by auto next case False show ?thesis by (rule k) (use minkl \a \ c\ that False in auto) qed then show "norm (f c - f a) \ e * (b-a) / 8" by simp qed finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" unfolding content_real[OF \a \ c\] by auto qed then show ?thesis by (rule_tac da="min k l" in that) (auto simp: l \0 < k\) qed obtain db where "0 < db" and db: "\c. \c \ b; {c..b} \ {a..b}; {c..b} \ ball b db\ \ norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ (e * (b-a)) / 4" proof - have "continuous (at b within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm(x-b); norm(x-b) < k\ \ norm (f b - f x) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm norm_minus_commute by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' b) \ (e * (b-a)) / 8" proof (cases "f' b = 0") case True thus ?thesis using ab e that by auto next case False show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' b)) *\<^sub>R f' b) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' b)" using False ab e by (auto simp add: field_simps) qed qed have "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" if "c \ b" "{c..b} \ {a..b}" and bmin: "{c..b} \ ball b (min k l)" for c proof - have minkl: "\b - x\ < min k l" if "x \ {c..b}" for x using bmin dist_real_def that by auto then have lel: "\b - c\ \ \l\" using that by force have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((b - c) *\<^sub>R f' b) \ norm (l *\<^sub>R f' b)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((b - c) *\<^sub>R f' b) \ e * (b-a) / 8" . next have "norm (f b - f c) < e * (b-a) / 8" proof (cases "b = c") case True with eba8 show ?thesis by auto next case False show ?thesis by (rule k) (use minkl \c \ b\ that False in auto) qed then show "norm (f b - f c) \ e * (b-a) / 8" by simp qed finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" unfolding content_real[OF \c \ b\] by auto qed then show ?thesis by (rule_tac db="min k l" in that) (auto simp: l \0 < k\) qed let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" proof (rule exI, safe) show "gauge ?d" using ab \db > 0\ \da > 0\ d(1) by (auto intro: gauge_ball_dependent) next fix p assume ptag: "p tagged_division_of {a..b}" and fine: "?d fine p" let ?A = "{t. fst t \ {a, b}}" note p = tagged_division_ofD[OF ptag] have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" using ptag fine by auto have le_xz: "\w x y z::real. y \ z/2 \ w - x \ z/2 \ w + y \ x + z" by arith have non: False if xk: "(x,K) \ p" and "x \ a" "x \ b" and less: "e * (Sup K - Inf K)/2 < norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" for x K proof - obtain u v where k: "K = cbox u v" using p(4) xk by blast then have "u \ v" and uv: "{u, v} \ cbox u v" using p(2)[OF xk] by auto then have result: "e * (v - u)/2 < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using less[unfolded k box_real interval_bounds_real content_real] by auto then have "x \ box a b" using p(2) p(3) \x \ a\ \x \ b\ xk by fastforce with d have *: "\y. norm (y-x) < d x \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" by metis have xd: "norm (u - x) < d x" "norm (v - x) < d x" using fineD[OF fine xk] \x \ a\ \x \ b\ uv by (auto simp add: k subset_eq dist_commute dist_real_def) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm ((f u - f x - (u - x) *\<^sub>R f' x) - (f v - f x - (v - x) *\<^sub>R f' x))" by (rule arg_cong[where f=norm]) (auto simp: scaleR_left.diff) also have "\ \ e/2 * norm (u - x) + e/2 * norm (v - x)" by (metis norm_triangle_le_diff add_mono * xd) also have "\ \ e/2 * norm (v - u)" using p(2)[OF xk] by (auto simp add: field_simps k) also have "\ < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using result by (simp add: \u \ v\) finally have "e * (v - u)/2 < e * (v - u)/2" using uv by auto then show False by auto qed have "norm (\(x, K)\p - ?A. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\(x, K)\p - ?A. norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" by (auto intro: sum_norm_le) also have "... \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k)/2)" using non by (fastforce intro: sum_mono) finally have I: "norm (\(x, k)\p - ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" by (simp add: sum_divide_distrib) have II: "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - (\n\p \ ?A. e * (case n of (x, k) \ Sup k - Inf k)) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" proof - have ge0: "0 \ e * (Sup k - Inf k)" if xkp: "(x, k) \ p \ ?A" for x k proof - obtain u v where uv: "k = cbox u v" by (meson Int_iff xkp p(4)) with p(2) that uv have "cbox u v \ {}" by blast then show "0 \ e * ((Sup k) - (Inf k))" unfolding uv using e by (auto simp add: field_simps) qed let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" let ?C = "{t \ p. fst t \ {a, b} \ content (snd t) \ 0}" have "norm (\(x, k)\p \ {t. fst t \ {a, b}}. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" proof - have *: "\S f e. sum f S = sum f (p \ ?C) \ norm (sum f (p \ ?C)) \ e \ norm (sum f S) \ e" by auto have 1: "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" if "(x,K) \ p \ {t. fst t \ {a, b}} - p \ ?C" for x K proof - have xk: "(x,K) \ p" and k0: "content K = 0" using that by auto then obtain u v where uv: "K = cbox u v" using p(4) by blast then have "u = v" using xk k0 p(2) by force then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" using xk unfolding uv by auto qed have 2: "norm(\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b-a)/2" proof - have norm_le: "norm (sum f S) \ e" if "\x y. \x \ S; y \ S\ \ x = y" "\x. x \ S \ norm (f x) \ e" "e > 0" for S f and e :: real proof (cases "S = {}") case True with that show ?thesis by auto next case False then obtain x where "x \ S" by auto then have "S = {x}" using that(1) by auto then show ?thesis using \x \ S\ that(2) by auto qed have *: "p \ ?C = ?B a \ ?B b" by blast then have "norm (\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) = norm (\(x,K) \ ?B a \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" by simp also have "... = norm ((\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) + (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" using p(1) ab e by (subst sum.union_disjoint) auto also have "... \ e * (b - a) / 4 + e * (b - a) / 4" proof (rule norm_triangle_le [OF add_mono]) have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(a, K) \ p" "(a, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pa obtain v v' where v: "K = cbox a v" "a \ v" and v': "K' = cbox a v'" "a \ v'" by blast let ?v = "min v v'" have "box a ?v \ K \ K'" unfolding v v' by (auto simp add: mem_box) then have "interior (box a (min v v')) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have "(a + ?v)/2 \ box a ?v" using ne0 unfolding v v' content_eq_0 not_le by (auto simp add: mem_box) ultimately have "(a + ?v)/2 \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(a, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(a, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox a v" and "a \ v" using pa[OF \(a, c) \ p\] by metis then have "a \ {a..v}" "v \ b" using p(3)[OF \(a, c) \ p\] by auto moreover have "{a..v} \ ball a da" using fineD[OF \?d fine p\ \(a, c) \ p\] by (simp add: v split: if_split_asm) ultimately show ?thesis unfolding v interval_bounds_real[OF \a \ v\] box_real using da \a \ v\ by auto qed qed (use ab e in auto) next have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(b, K) \ p" "(b, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pb obtain v v' where v: "K = cbox v b" "v \ b" and v': "K' = cbox v' b" "v' \ b" by blast let ?v = "max v v'" have "box ?v b \ K \ K'" unfolding v v' by (auto simp: mem_box) then have "interior (box (max v v') b) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have " ((b + ?v)/2) \ box ?v b" using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box) ultimately have "((b + ?v)/2) \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(b, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(b, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox v b" and "v \ b" using \(b, c) \ p\ pb by blast then have "v \ a""b \ {v.. b}" using p(3)[OF \(b, c) \ p\] by auto moreover have "{v..b} \ ball b db" using fineD[OF \?d fine p\ \(b, c) \ p\] box_real(2) v False by force ultimately show ?thesis using db v by auto qed qed (use ab e in auto) qed also have "... = e * (b-a)/2" by simp finally show "norm (\(x,k)\p \ ?C. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" . qed show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \ e * (b-a)/2" apply (rule * [OF sum.mono_neutral_right[OF pA(2)]]) using 1 2 by (auto simp: split_paired_all) qed also have "... = (\n\p. e * (case n of (x, k) \ Sup k - Inf k))/2" unfolding sum_distrib_left[symmetric] by (subst additive_tagged_division_1[OF \a \ b\ ptag]) auto finally have norm_le: "norm (\(x,K)\p \ {t. fst t \ {a, b}}. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, K) \ Sup K - Inf K))/2" . have le2: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2)/2 \ x - s1 \ s2/2" by auto show ?thesis apply (rule le2 [OF sum_nonneg]) using ge0 apply force by (metis (no_types, lifting) Diff_Diff_Int Diff_subset norm_le p(1) sum.subset_diff) qed note * = additive_tagged_division_1[OF assms(1) ptag, symmetric] have "norm (\(x,K)\p \ ?A \ (p - ?A). content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (\(x,K)\p \ ?A \ (p - ?A). Sup K - Inf K)" unfolding sum_distrib_left unfolding sum.union_disjoint[OF pA(2-)] using le_xz norm_triangle_le I II by blast then show "norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b}" by (simp only: content_real[OF \a \ b\] *[of "\x. x"] *[of f] sum_subtractf[symmetric] split_minus pA(1) [symmetric]) qed qed qed subsection \Stronger form with finite number of exceptional points\ lemma fundamental_theorem_of_calculus_interior_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induction arbitrary: a b) case empty then show ?case using fundamental_theorem_of_calculus_interior by force next case (insert x S) show ?case proof (cases "x \ {a<..continuous_on {a..b} f\ \a < x\ \x < b\ continuous_on_subset by (force simp: intro!: insert)+ then have "(f' has_integral f x - f a + (f b - f x)) {a..b}" using \a < x\ \x < b\ has_integral_combine less_imp_le by blast then show ?thesis by simp qed qed corollary fundamental_theorem_of_calculus_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" and vec: "\x. x \ {a..b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a..b} f" shows "(f' has_integral (f b - f a)) {a..b}" by (rule fundamental_theorem_of_calculus_interior_strong [OF \finite S\]) (force simp: assms)+ proposition indefinite_integral_continuous_left: fixes f:: "real \ 'a::banach" assumes intf: "f integrable_on {a..b}" and "a < c" "c \ b" "e > 0" obtains d where "d > 0" and "\t. c - d < t \ t \ c \ norm (integral {a..c} f - integral {a..t} f) < e" proof - obtain w where "w > 0" and w: "\t. \c - w < t; t < c\ \ norm (f c) * norm(c - t) < e/3" proof (cases "f c = 0") case False hence e3: "0 < e/3 / norm (f c)" using \e>0\ by simp moreover have "norm (f c) * norm (c - t) < e/3" if "t < c" and "c - e/3 / norm (f c) < t" for t proof - have "norm (c - t) < e/3 / norm (f c)" using that by auto then show "norm (f c) * norm (c - t) < e/3" by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff) qed ultimately show ?thesis using that by auto next case True then show ?thesis using \e > 0\ that by auto qed let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" have e3: "e/3 > 0" using \e > 0\ by auto have "f integrable_on {a..c}" using \a < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) then obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of {a..c}; d1 fine p\ \ norm (?SUM p - integral {a..c} f) < e/3" using integrable_integral has_integral_real e3 by metis define d where [abs_def]: "d x = ball x w \ d1 x" for x have "gauge d" unfolding d_def using \w > 0\ \gauge d1\ by auto then obtain k where "0 < k" and k: "ball c k \ d c" by (meson gauge_def open_contains_ball) let ?d = "min k (c - a)/2" show thesis proof (intro that[of ?d] allI impI, safe) show "?d > 0" using \0 < k\ \a < c\ by auto next fix t assume t: "c - ?d < t" "t \ c" show "norm (integral ({a..c}) f - integral ({a..t}) f) < e" proof (cases "t < c") case False with \t \ c\ show ?thesis by (simp add: \e > 0\) next case True have "f integrable_on {a..t}" using \t < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) then obtain d2 where d2: "gauge d2" "\p. p tagged_division_of {a..t} \ d2 fine p \ norm (?SUM p - integral {a..t} f) < e/3" using integrable_integral has_integral_real e3 by metis define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x have "gauge d3" using \gauge d1\ \gauge d2\ unfolding d3_def gauge_def by auto then obtain p where ptag: "p tagged_division_of {a..t}" and pfine: "d3 fine p" by (metis box_real(2) fine_division_exists) note p' = tagged_division_ofD[OF ptag] have pt: "(x,K)\p \ x \ t" for x K by (meson atLeastAtMost_iff p'(2) p'(3) subsetCE) with pfine have "d2 fine p" unfolding fine_def d3_def by fastforce then have d2_fin: "norm (?SUM p - integral {a..t} f) < e/3" using d2(2) ptag by auto have eqs: "{a..c} \ {x. x \ t} = {a..t}" "{a..c} \ {x. x \ t} = {t..c}" using t by (auto simp add: field_simps) have "p \ {(c, {t..c})} tagged_division_of {a..c}" proof (intro tagged_division_Un_interval_real) show "{(c, {t..c})} tagged_division_of {a..c} \ {x. t \ x \ 1}" using \t \ c\ by (auto simp: eqs tagged_division_of_self_real) qed (auto simp: eqs ptag) moreover have "d1 fine p \ {(c, {t..c})}" unfolding fine_def proof safe fix x K y assume "(x,K) \ p" and "y \ K" then show "y \ d1 x" by (metis Int_iff d3_def subsetD fineD pfine) next fix x assume "x \ {t..c}" then have "dist c x < k" using t(1) by (auto simp add: field_simps dist_real_def) with k show "x \ d1 c" unfolding d_def by auto qed ultimately have d1_fin: "norm (?SUM(p \ {(c, {t..c})}) - integral {a..c} f) < e/3" using d1 by metis have SUMEQ: "?SUM(p \ {(c, {t..c})}) = (c - t) *\<^sub>R f c + ?SUM p" proof - have "?SUM(p \ {(c, {t..c})}) = (content{t..c} *\<^sub>R f c) + ?SUM p" proof (subst sum.union_disjoint) show "p \ {(c, {t..c})} = {}" using \t < c\ pt by force qed (use p'(1) in auto) also have "... = (c - t) *\<^sub>R f c + ?SUM p" using \t \ c\ by auto finally show ?thesis . qed have "c - k < t" using \k>0\ t(1) by (auto simp add: field_simps) moreover have "k \ w" proof (rule ccontr) assume "\ k \ w" then have "c + (k + w) / 2 \ d c" by (auto simp add: field_simps not_le not_less dist_real_def d_def) then have "c + (k + w) / 2 \ ball c k" using k by blast then show False using \0 < w\ \\ k \ w\ dist_real_def by auto qed ultimately have cwt: "c - w < t" by (auto simp add: field_simps) have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) - integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c" by auto have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3" unfolding eq proof (intro norm_triangle_lt add_strict_mono) show "norm (- ((c - t) *\<^sub>R f c + ?SUM p - integral {a..c} f)) < e/3" by (metis SUMEQ d1_fin norm_minus_cancel) show "norm (?SUM p - integral {a..t} f) < e/3" using d2_fin by blast show "norm ((c - t) *\<^sub>R f c) < e/3" using w cwt \t < c\ by simp (simp add: field_simps) qed then show ?thesis by simp qed qed qed lemma indefinite_integral_continuous_right: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "a \ c" and "c < b" and "e > 0" obtains d where "0 < d" and "\t. c \ t \ t < c + d \ norm (integral {a..c} f - integral {a..t} f) < e" proof - have intm: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" using assms by auto from indefinite_integral_continuous_left[OF intm \e>0\] obtain d where "0 < d" and d: "\t. \- c - d < t; t \ -c\ \ norm (integral {- b..- c} (\x. f (-x)) - integral {- b..t} (\x. f (-x))) < e" by metis let ?d = "min d (b - c)" show ?thesis proof (intro that[of "?d"] allI impI) show "0 < ?d" using \0 < d\ \c < b\ by auto fix t :: real assume t: "c \ t \ t < c + ?d" have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f" "integral {a..t} f = integral {a..b} f - integral {t..b} f" using assms t by (auto simp: algebra_simps integral_combine) have "(- c) - d < (- t)" "- t \ - c" using t by auto from d[OF this] show "norm (integral {a..c} f - integral {a..t} f) < e" by (auto simp add: algebra_simps norm_minus_commute *) qed qed lemma indefinite_integral_continuous_1: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {a..x} f)" proof - have "\d>0. \x'\{a..b}. dist x' x < d \ dist (integral {a..x'} f) (integral {a..x} f) < e" if x: "x \ {a..b}" and "e > 0" for x e :: real proof (cases "a = b") case True with that show ?thesis by force next case False with x have "a < b" by force with x consider "x = a" | "x = b" | "a < x" "x < b" by force then show ?thesis proof cases case 1 then show ?thesis by (force simp: dist_norm algebra_simps intro: indefinite_integral_continuous_right [OF assms _ \a < b\ \e > 0\]) next case 2 then show ?thesis by (force simp: dist_norm norm_minus_commute algebra_simps intro: indefinite_integral_continuous_left [OF assms \a < b\ _ \e > 0\]) next case 3 obtain d1 where "0 < d1" and d1: "\t. \x - d1 < t; t \ x\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_left [OF assms \a < x\ _ \e > 0\]) obtain d2 where "0 < d2" and d2: "\t. \x \ t; t < x + d2\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_right [OF assms _ \x < b\ \e > 0\]) show ?thesis proof (intro exI ballI conjI impI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by simp show "dist (integral {a..y} f) (integral {a..x} f) < e" if "y \ {a..b}" "dist y x < min d1 d2" for y proof (cases "y < x") case True with that d1 show ?thesis by (auto simp: dist_commute dist_norm) next case False with that d2 show ?thesis by (auto simp: dist_commute dist_norm) qed qed qed qed then show ?thesis by (auto simp: continuous_on_iff) qed lemma indefinite_integral_continuous_1': fixes f::"real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {x..b} f)" proof - have "integral {a..b} f - integral {a..x} f = integral {x..b} f" if "x \ {a..b}" for x using integral_combine[OF _ _ assms, of x] that by (auto simp: algebra_simps) with _ show ?thesis by (rule continuous_on_eq) (auto intro!: continuous_intros indefinite_integral_continuous_1 assms) qed theorem integral_has_vector_derivative': fixes f :: "real \ 'b::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {u..b} f) has_vector_derivative - f x) (at x within {a..b})" proof - have *: "integral {x..b} f = integral {a .. b} f - integral {a .. x} f" if "a \ x" "x \ b" for x using integral_combine[of a x b for x, OF that integrable_continuous_real[OF assms(1)]] by (simp add: algebra_simps) show ?thesis using \x \ _\ * by (rule has_vector_derivative_transform) (auto intro!: derivative_eq_intros assms integral_has_vector_derivative) qed lemma integral_has_real_derivative': assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {x..b} g) has_real_derivative -g t) (at t within {a..b})" using integral_has_vector_derivative'[OF assms] by (auto simp: has_field_derivative_iff_has_vector_derivative) subsection \This doesn't directly involve integration, but that gives an easy proof\ lemma has_derivative_zero_unique_strong_interval: fixes f :: "real \ 'a::banach" assumes "finite k" and contf: "continuous_on {a..b} f" and "f a = y" and fder: "\x. x \ {a..b} - k \ (f has_derivative (\h. 0)) (at x within {a..b})" and x: "x \ {a..b}" shows "f x = y" proof - have "a \ b" "a \ x" using assms by auto have "((\x. 0::'a) has_integral f x - f a) {a..x}" proof (rule fundamental_theorem_of_calculus_interior_strong[OF \finite k\ \a \ x\]; clarify?) have "{a..x} \ {a..b}" using x by auto then show "continuous_on {a..x} f" by (rule continuous_on_subset[OF contf]) show "(f has_vector_derivative 0) (at z)" if z: "z \ {a<.. k" for z unfolding has_vector_derivative_def proof (simp add: at_within_open[OF z, symmetric]) show "(f has_derivative (\x. 0)) (at z within {a<..Generalize a bit to any convex set\ lemma has_derivative_zero_unique_strong_convex: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "convex S" "finite K" and contf: "continuous_on S f" and "c \ S" "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof (cases "x = c") case True with \f c = y\ show ?thesis by blast next case False let ?\ = "\u. (1 - u) *\<^sub>R c + u *\<^sub>R x" have contf': "continuous_on {0 ..1} (f \ ?\)" proof (rule continuous_intros continuous_on_subset[OF contf])+ show "(\u. (1 - u) *\<^sub>R c + u *\<^sub>R x) ` {0..1} \ S" using \convex S\ \x \ S\ \c \ S\ by (auto simp add: convex_alt algebra_simps) qed have "t = u" if "?\ t = ?\ u" for t u proof - from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c" by (auto simp add: algebra_simps) then show ?thesis using \x \ c\ by auto qed then have eq: "(SOME t. ?\ t = ?\ u) = u" for u by blast then have "(?\ -` K) \ (\z. SOME t. ?\ t = z) ` K" by (clarsimp simp: image_iff) (metis (no_types) eq) then have fin: "finite (?\ -` K)" by (rule finite_surj[OF \finite K\]) have derf': "((\u. f (?\ u)) has_derivative (\h. 0)) (at t within {0..1})" if "t \ {0..1} - {t. ?\ t \ K}" for t proof - have df: "(f has_derivative (\h. 0)) (at (?\ t) within ?\ ` {0..1})" using \convex S\ \x \ S\ \c \ S\ that by (auto simp add: convex_alt algebra_simps intro: has_derivative_subset [OF derf]) have "(f \ ?\ has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})" by (rule derivative_eq_intros df | simp)+ then show ?thesis unfolding o_def . qed have "(f \ ?\) 1 = y" apply (rule has_derivative_zero_unique_strong_interval[OF fin contf']) unfolding o_def using \f c = y\ derf' by auto then show ?thesis by auto qed text \Also to any open connected set with finite set of exceptions. Could generalize to locally convex set with limpt-free set of exceptions.\ lemma has_derivative_zero_unique_strong_connected: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite K" and contf: "continuous_on S f" and "c \ S" and "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof - have "\e>0. ball x e \ (S \ f -` {f x})" if "x \ S" for x proof - obtain e where "0 < e" and e: "ball x e \ S" using \x \ S\ \open S\ open_contains_ball by blast have "ball x e \ {u \ S. f u \ {f x}}" proof safe fix y assume y: "y \ ball x e" then show "y \ S" using e by auto show "f y = f x" proof (rule has_derivative_zero_unique_strong_convex[OF convex_ball \finite K\]) show "continuous_on (ball x e) f" using contf continuous_on_subset e by blast show "(f has_derivative (\h. 0)) (at u within ball x e)" if "u \ ball x e - K" for u by (metis Diff_iff contra_subsetD derf e has_derivative_subset that) qed (use y e \0 < e\ in auto) qed then show "\e>0. ball x e \ (S \ f -` {f x})" using \0 < e\ by blast qed then have "openin (top_of_set S) (S \ f -` {y})" by (auto intro!: open_openin_trans[OF \open S\] simp: open_contains_ball) moreover have "closedin (top_of_set S) (S \ f -` {y})" by (force intro!: continuous_closedin_preimage [OF contf]) ultimately have "(S \ f -` {y}) = {} \ (S \ f -` {y}) = S" using \connected S\ by (simp add: connected_clopen) then show ?thesis using \x \ S\ \f c = y\ \c \ S\ by auto qed lemma has_derivative_zero_connected_constant: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite k" and "continuous_on S f" and "\x\(S - k). (f has_derivative (\h. 0)) (at x within S)" obtains c where "\x. x \ S \ f(x) = c" proof (cases "S = {}") case True then show ?thesis by (metis empty_iff that) next case False then obtain c where "c \ S" by (metis equals0I) then show ?thesis by (metis has_derivative_zero_unique_strong_connected assms that) qed lemma DERIV_zero_connected_constant: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes "connected S" and "open S" and "finite K" and "continuous_on S f" and "\x\(S - K). DERIV f x :> 0" obtains c where "\x. x \ S \ f(x) = c" using has_derivative_zero_connected_constant [OF assms(1-4)] assms by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def) subsection \Integrating characteristic function of an interval\ lemma has_integral_restrict_open_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "(f has_integral i) (cbox c d)" and cb: "cbox c d \ cbox a b" shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" proof (cases "cbox c d = {}") case True then have "box c d = {}" by (metis bot.extremum_uniqueI box_subset_cbox) then show ?thesis using True intf by auto next case False then obtain p where pdiv: "p division_of cbox a b" and inp: "cbox c d \ p" using cb partial_division_extend_1 by blast define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x interpret operative "lift_option plus" "Some (0 :: 'b)" "\i. if g integrable_on i then Some (integral i g) else None" by (fact operative_integralI) note operat = division [OF pdiv, symmetric] show ?thesis proof (cases "(g has_integral i) (cbox a b)") case True then show ?thesis by (simp add: g_def) next case False have iterate:"F (\i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0" proof (intro neutral ballI) fix x assume x: "x \ p - {cbox c d}" then have "x \ p" by auto then obtain u v where uv: "x = cbox u v" using pdiv by blast have "interior x \ interior (cbox c d) = {}" using pdiv inp x by blast then have "(g has_integral 0) x" unfolding uv using has_integral_spike_interior[where f="\x. 0"] by (metis (no_types, lifting) disjoint_iff_not_equal g_def has_integral_0_eq interior_cbox) then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" by auto qed interpret comm_monoid_set "lift_option plus" "Some (0 :: 'b)" by (intro comm_monoid_set.intro comm_monoid_lift_option add.comm_monoid_axioms) have intg: "g integrable_on cbox c d" using integrable_spike_interior[where f=f] by (meson g_def has_integral_integrable intf) moreover have "integral (cbox c d) g = i" proof (rule has_integral_unique[OF has_integral_spike_interior intf]) show "\x. x \ box c d \ f x = g x" by (auto simp: g_def) show "(g has_integral integral (cbox c d) g) (cbox c d)" by (rule integrable_integral[OF intg]) qed ultimately have "F (\A. if g integrable_on A then Some (integral A g) else None) p = Some i" by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral) then have "(g has_integral i) (cbox a b)" by (metis integrable_on_def integral_unique operat option.inject option.simps(3)) with False show ?thesis by blast qed qed lemma has_integral_restrict_closed_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) (cbox c d)" and "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" proof - note has_integral_restrict_open_subinterval[OF assms] note * = has_integral_spike[OF negligible_frontier_interval _ this] show ?thesis by (rule *[of c d]) (use box_subset_cbox[of c d] in auto) qed lemma has_integral_restrict_closed_subintervals_eq: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" (is "?l = ?r") proof (cases "cbox c d = {}") case False let ?g = "\x. if x \ cbox c d then f x else 0" show ?thesis proof assume ?l then have "?g integrable_on cbox c d" using assms has_integral_integrable integrable_subinterval by blast then have "f integrable_on cbox c d" by (rule integrable_eq) auto moreover then have "i = integral (cbox c d) f" by (meson \((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)\ assms has_integral_restrict_closed_subinterval has_integral_unique integrable_integral) ultimately show ?r by auto next assume ?r then show ?l by (rule has_integral_restrict_closed_subinterval[OF _ assms]) qed qed auto text \Hence we can apply the limit process uniformly to all integrals.\ lemma has_integral': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" (is "?l \ (\e>0. ?r e)") proof (cases "\a b. S = cbox a b") case False then show ?thesis by (simp add: has_integral_alt) next case True then obtain a b where S: "S = cbox a b" by blast obtain B where " 0 < B" and B: "\x. x \ cbox a b \ norm x \ B" using bounded_cbox[unfolded bounded_pos] by blast show ?thesis proof safe fix e :: real assume ?l and "e > 0" have "((\x. if x \ S then f x else 0) has_integral i) (cbox c d)" if "ball 0 (B+1) \ cbox c d" for c d unfolding S using B that by (force intro: \?l\[unfolded S] has_integral_restrict_closed_subinterval) then show "?r e" by (meson \0 < B\ \0 < e\ add_pos_pos le_less_trans zero_less_one norm_pths(2)) next assume as: "\e>0. ?r e" then obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b)" by (meson zero_less_one) define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" "i \ Basis" for x i using that and Basis_le_norm[OF \i\Basis\, of x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (meson B mem_box(2) subsetI) have "c \ i \ x \ i \ x \ i \ d \ i" if x: "norm (0 - x) < C" and i: "i \ Basis" for x i using Basis_le_norm[OF i, of x] x i by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain y where y: "(f has_integral y) (cbox a b)" using c_d has_integral_restrict_closed_subintervals_eq S by blast have "y = i" proof (rule ccontr) assume "y \ i" then have "0 < norm (y - i)" by auto from as[rule_format,OF this] obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z-i) < norm (y-i)" by auto define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" and "i \ Basis" for x i using that Basis_le_norm[of i x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (simp add: B mem_box(2) subset_eq) have "c \ i \ x \ i \ x \ i \ d \ i" if "norm (0 - x) < C" and "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain z where z: "(f has_integral z) (cbox a b)" "norm (z-i) < norm (y-i)" using has_integral_restrict_closed_subintervals_eq[OF c_d] S by blast moreover then have "z = y" by (blast intro: has_integral_unique[OF _ y]) ultimately show False by auto qed then show ?l using y by (auto simp: S) qed qed lemma has_integral_le: fixes f :: "'n::euclidean_space \ real" assumes fg: "(f has_integral i) S" "(g has_integral j) S" and le: "\x. x \ S \ f x \ g x" shows "i \ j" using has_integral_component_le[OF _ fg, of 1] le by auto lemma integral_le: fixes f :: "'n::euclidean_space \ real" assumes "f integrable_on S" and "g integrable_on S" and "\x. x \ S \ f x \ g x" shows "integral S f \ integral S g" by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) lemma has_integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) S" and "\x. x \ S \ 0 \ f x" shows "0 \ i" using has_integral_component_nonneg[of 1 f i S] unfolding o_def using assms by auto lemma integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes f: "f integrable_on S" and 0: "\x. x \ S \ 0 \ f x" shows "0 \ integral S f" by (rule has_integral_nonneg[OF f[unfolded has_integral_integral] 0]) text \Hence a general restriction property.\ lemma has_integral_restrict [simp]: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) S" proof - have *: "\x. (if x \ T then if x \ S then f x else 0 else 0) = (if x\S then f x else 0)" using assms by auto show ?thesis apply (subst(2) has_integral') apply (subst has_integral') apply (simp add: *) done qed corollary has_integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" by auto lemma has_integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) (S \ T)" proof - have "((\x. if x \ T then if x \ S then f x else 0 else 0) has_integral i) UNIV = ((\x. if x \ S \ T then f x else 0) has_integral i) UNIV" by (rule has_integral_cong) auto then show ?thesis using has_integral_restrict_UNIV by fastforce qed lemma integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral T (\x. if x \ S then f x else 0) = integral (S \ T) f" by (metis (no_types, lifting) has_integral_cong has_integral_restrict_Int integrable_integral integral_unique not_integrable_integral) lemma integrable_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(\x. if x \ S then f x else 0) integrable_on T \ f integrable_on (S \ T)" using has_integral_restrict_Int by fastforce lemma has_integral_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" and "\x. x \ S \ f x = 0" and "S \ T" shows "(f has_integral i) T" proof - have "(\x. if x \ S then f x else 0) = (\x. if x \ T then f x else 0)" using assms by fastforce with f show ?thesis by (simp only: has_integral_restrict_UNIV [symmetric, of f]) qed lemma integrable_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\x. x \ S \ f x = 0" and "S \ t" shows "f integrable_on t" using assms unfolding integrable_on_def by (auto intro:has_integral_on_superset) lemma integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "integral UNIV (\x. if x \ S then f x else 0) = integral S f" by (simp add: integral_restrict_Int) lemma integrable_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes k: "k \ Basis" and as: "S \ T" "(f has_integral i) S" "(f has_integral j) T" "\x. x\T \ 0 \ f(x)\k" shows "i\k \ j\k" proof - have \
: "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" by (simp_all add: assms) show ?thesis using as by (force intro!: has_integral_component_le[OF k \
]) qed subsection\Integrals on set differences\ lemma has_integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes S: "(f has_integral i) S" and T: "(f has_integral j) T" and neg: "negligible (T - S)" shows "(f has_integral (i - j)) (S - T)" proof - show ?thesis unfolding has_integral_restrict_UNIV [symmetric, of f] proof (rule has_integral_spike [OF neg]) have eq: "(\x. (if x \ S then f x else 0) - (if x \ T then f x else 0)) = (\x. if x \ T - S then - f x else if x \ S - T then f x else 0)" by (force simp add: ) have "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" using S T has_integral_restrict_UNIV by auto from has_integral_diff [OF this] show "((\x. if x \ T - S then - f x else if x \ S - T then f x else 0) has_integral i-j) UNIV" by (simp add: eq) qed force qed lemma integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "f integrable_on S" "f integrable_on T" "negligible(T - S)" shows "integral (S - T) f = integral S f - integral T f" by (rule integral_unique) (simp add: assms has_integral_setdiff integrable_integral) lemma integrable_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) S" "(f has_integral j) T" "negligible (T - S)" shows "f integrable_on (S - T)" using has_integral_setdiff [OF assms] by (simp add: has_integral_iff) lemma negligible_setdiff [simp]: "T \ S \ negligible (T - S)" by (metis Diff_eq_empty_iff negligible_empty) lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") proof assume R: ?r show ?l unfolding negligible_def proof safe fix a b have "negligible (s \ cbox a b)" by (simp add: R) then show "(indicator s has_integral 0) (cbox a b)" by (meson Diff_iff Int_iff has_integral_negligible indicator_simps(2)) qed qed (simp add: negligible_Int) lemma negligible_translation: assumes "negligible S" shows "negligible ((+) c ` S)" proof - have inj: "inj ((+) c)" by simp show ?thesis using assms proof (clarsimp simp: negligible_def) fix a b assume "\x y. (indicator S has_integral 0) (cbox x y)" then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) have eq: "indicator ((+) c ` S) = (\x. indicator S (x - c))" by (force simp add: indicator_def) show "(indicator ((+) c ` S) has_integral 0) (cbox a b)" using has_integral_affinity [OF *, of 1 "-c"] cbox_translation [of "c" "-c+a" "-c+b"] by (simp add: eq) (simp add: ac_simps) qed qed lemma negligible_translation_rev: assumes "negligible ((+) c ` S)" shows "negligible S" by (metis negligible_translation [OF assms, of "-c"] translation_galois) lemma has_integral_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f has_integral y) S \ (f has_integral y) T" proof - have "((\x. if x \ S then f x else 0) has_integral y) UNIV = ((\x. if x \ T then f x else 0) has_integral y) UNIV" proof (rule has_integral_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto then show ?thesis by (simp add: has_integral_restrict_UNIV) qed corollary integral_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "integral S f = integral T f" using has_integral_spike_set_eq [OF assms] by (metis eq_integralD integral_unique) lemma integrable_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f integrable_on T" using has_integral_spike_set_eq [OF neg] f by blast lemma integrable_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((S - T) \ (T - S))" shows "f integrable_on S \ f integrable_on T" by (blast intro: integrable_spike_set assms negligible_subset) lemma integrable_on_insert_iff: "f integrable_on (insert x X) \ f integrable_on X" for f::"_ \ 'a::banach" by (rule integrable_spike_set_eq) (auto simp: insert_Diff_if) lemma has_integral_interior: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (interior S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF empty_imp_negligible negligible_subset]) (use interior_subset in \auto simp: frontier_def closure_def\) lemma has_integral_closure: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (closure S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF negligible_subset empty_imp_negligible]) (auto simp: closure_Un_frontier ) lemma has_integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(f has_integral y) (box a b) \ (f has_integral y) (cbox a b)" unfolding interior_cbox [symmetric] by (metis frontier_cbox has_integral_interior negligible_frontier_interval) lemma integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "f integrable_on box a b \ f integrable_on cbox a b" by (simp add: has_integral_open_interval integrable_on_def) lemma integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral(box a b) f = integral(cbox a b) f" by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral) subsection \More lemmas that are useful later\ lemma has_integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "(f has_integral i) s" and "(f has_integral j) t" and "\x\t. 0 \ f x" shows "i \ j" using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] using assms by auto lemma integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "k \ Basis" and "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x \ k" shows "(integral s f)\k \ (integral t f)\k" by (meson assms has_integral_subset_component_le integrable_integral) lemma integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x" shows "integral s f \ integral t f" using assms has_integral_subset_le by blast lemma has_integral_alt': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" (is "?l = ?r") proof assume rhs: ?r show ?l proof (subst has_integral', intro allI impI) fix e::real assume "e > 0" from rhs[THEN conjunct2,rule_format,OF this] show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e)" by (simp add: has_integral_iff rhs) qed next let ?\ = "\e a b. \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e" assume ?l then have lhs: "\B>0. \a b. ball 0 B \ cbox a b \ ?\ e a b" if "e > 0" for e using that has_integral'[of f] by auto let ?f = "\x. if x \ s then f x else 0" show ?r proof (intro conjI allI impI) fix a b :: 'n from lhs[OF zero_less_one] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ ?\ 1 a b" by blast let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" show "?f integrable_on cbox a b" proof (rule integrable_subinterval[of _ ?a ?b]) have "?a \ i \ x \ i \ x \ i \ ?b \ i" if "norm (0 - x) < B" "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp add:field_simps) then have "ball 0 B \ cbox ?a ?b" by (auto simp: mem_box dist_norm) then show "?f integrable_on cbox ?a ?b" unfolding integrable_on_def using B by blast show "cbox a b \ cbox ?a ?b" by (force simp: mem_box) qed fix e :: real assume "e > 0" with lhs show "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" by (metis (no_types, lifting) has_integral_integrable_integral) qed qed subsection \Continuity of the integral (for a 1-dimensional interval)\ lemma integrable_alt: fixes f :: "'n::euclidean_space \ 'a::banach" shows "f integrable_on s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" (is "?l = ?r") proof let ?F = "\x. if x \ s then f x else 0" assume ?l then obtain y where intF: "\a b. ?F integrable_on cbox a b" and y: "\e. 0 < e \ \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e" unfolding integrable_on_def has_integral_alt'[of f] by auto show ?r proof (intro conjI allI impI intF) fix e::real assume "e > 0" then have "e/2 > 0" by auto obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e/2" using \0 < e/2\ y by blast show "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" proof (intro conjI exI impI allI, rule \0 < B\) fix a b c d::'n assume sub: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" show "norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" using sub by (auto intro: norm_triangle_half_l dest: B) qed qed next let ?F = "\x. if x \ s then f x else 0" assume rhs: ?r let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" have "Cauchy (\n. integral (?cube n) ?F)" unfolding Cauchy_def proof (intro allI impI) fix e::real assume "e > 0" with rhs obtain B where "0 < B" and B: "\a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" by blast obtain N where N: "B \ real N" using real_arch_simple by blast have "ball 0 B \ ?cube n" if n: "n \ N" for n proof - have "sum ((*\<^sub>R) (- real n)) Basis \ i \ x \ i \ x \ i \ sum ((*\<^sub>R) (real n)) Basis \ i" if "norm x < B" "i \ Basis" for x i::'n using Basis_le_norm[of i x] n N that by (auto simp add: field_simps sum_negf) then show ?thesis by (auto simp: mem_box dist_norm) qed then show "\M. \m\M. \n\M. dist (integral (?cube m) ?F) (integral (?cube n) ?F) < e" by (fastforce simp add: dist_norm intro!: B) qed then obtain i where i: "(\n. integral (?cube n) ?F) \ i" using convergent_eq_Cauchy by blast have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - i) < e" if "e > 0" for e proof - have *: "e/2 > 0" using that by auto then obtain N where N: "\n. N \ n \ norm (i - integral (?cube n) ?F) < e/2" using i[THEN LIMSEQ_D, simplified norm_minus_commute] by meson obtain B where "0 < B" and B: "\a b c d. \ball 0 B \ cbox a b; ball 0 B \ cbox c d\ \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e/2" using rhs * by meson let ?B = "max (real N) B" show ?thesis proof (intro exI conjI allI impI) show "0 < ?B" using \B > 0\ by auto fix a b :: 'n assume "ball 0 ?B \ cbox a b" moreover obtain n where n: "max (real N) B \ real n" using real_arch_simple by blast moreover have "ball 0 B \ ?cube n" proof fix x :: 'n assume x: "x \ ball 0 B" have "\norm (0 - x) < B; i \ Basis\ \ sum ((*\<^sub>R) (-n)) Basis \ i\ x \ i \ x \ i \ sum ((*\<^sub>R) n) Basis \ i" for i using Basis_le_norm[of i x] n by (auto simp add: field_simps sum_negf) then show "x \ ?cube n" using x by (auto simp: mem_box dist_norm) qed ultimately show "norm (integral (cbox a b) ?F - i) < e" using norm_triangle_half_l [OF B N] by force qed qed then show ?l unfolding integrable_on_def has_integral_alt'[of f] using rhs by blast qed lemma integrable_altD: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" using assms[unfolded integrable_alt[of f]] by auto lemma integrable_alt_subset: fixes f :: "'a::euclidean_space \ 'b::banach" shows "f integrable_on S \ (\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm(integral (cbox a b) (\x. if x \ S then f x else 0) - integral (cbox c d) (\x. if x \ S then f x else 0)) < e)" (is "_ = ?rhs") proof - let ?g = "\x. if x \ S then f x else 0" have "f integrable_on S \ (\a b. ?g integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e)" by (rule integrable_alt) also have "\ = ?rhs" proof - { fix e :: "real" assume e: "\e. e>0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" and "e > 0" obtain B where "B > 0" and B: "\a b c d. \ball 0 B \ cbox a b; cbox a b \ cbox c d\ \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e/2" using \e > 0\ e [of "e/2"] by force have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" proof (intro exI allI conjI impI) fix a b c d :: "'a" let ?\ = "\i\Basis. max (a \ i) (c \ i) *\<^sub>R i" let ?\ = "\i\Basis. min (b \ i) (d \ i) *\<^sub>R i" show "norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" if ball: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" proof - have B': "norm (integral (cbox a b \ cbox c d) ?g - integral (cbox x y) ?g) < e/2" if "cbox a b \ cbox c d \ cbox x y" for x y using B [of ?\ ?\ x y] ball that by (simp add: Int_interval [symmetric]) show ?thesis using B' [of a b] B' [of c d] norm_triangle_half_r by blast qed qed (use \B > 0\ in auto)} then show ?thesis by force qed finally show ?thesis . qed lemma integrable_on_subcbox: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on S" and sub: "cbox a b \ S" shows "f integrable_on cbox a b" proof - have "(\x. if x \ S then f x else 0) integrable_on cbox a b" by (simp add: intf integrable_altD(1)) then show ?thesis by (metis (mono_tags) sub integrable_restrict_Int le_inf_iff order_refl subset_antisym) qed subsection \A straddling criterion for integrability\ lemma integrable_straddle_interval: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" shows "f integrable_on cbox a b" proof - have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ d fine p1 \ p2 tagged_division_of cbox a b \ d fine p2 \ \(\(x,K)\p1. content K *\<^sub>R f x) - (\(x,K)\p2. content K *\<^sub>R f x)\ < e)" if "e > 0" for e proof - have e: "e/3 > 0" using that by auto then obtain g h i j where ij: "\i - j\ < e/3" and "(g has_integral i) (cbox a b)" and "(h has_integral j) (cbox a b)" and fgh: "\x. x \ cbox a b \ g x \ f x \ f x \ h x" using assms real_norm_def by metis then obtain d1 d2 where "gauge d1" "gauge d2" and d1: "\p. \p tagged_division_of cbox a b; d1 fine p\ \ \(\(x,K)\p. content K *\<^sub>R g x) - i\ < e/3" and d2: "\p. \p tagged_division_of cbox a b; d2 fine p\ \ \(\(x,K) \ p. content K *\<^sub>R h x) - j\ < e/3" by (metis e has_integral real_norm_def) have "\(\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)\ < e" if p1: "p1 tagged_division_of cbox a b" and 11: "d1 fine p1" and 21: "d2 fine p1" and p2: "p2 tagged_division_of cbox a b" and 12: "d1 fine p2" and 22: "d2 fine p2" for p1 p2 proof - have *: "\g1 g2 h1 h2 f1 f2. \\g2 - i\ < e/3; \g1 - i\ < e/3; \h2 - j\ < e/3; \h1 - j\ < e/3; g1 - h2 \ f1 - f2; f1 - f2 \ h1 - g2\ \ \f1 - f2\ < e" using \e > 0\ ij by arith have 0: "(\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p1. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p2. content k *\<^sub>R h x) - (\(x, k)\p2. content k *\<^sub>R f x)" "(\(x, k)\p2. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p1. content k *\<^sub>R h x) - (\(x, k)\p1. content k *\<^sub>R f x)" unfolding sum_subtractf[symmetric] apply (auto intro!: sum_nonneg) apply (meson fgh measure_nonneg mult_left_mono tag_in_interval that sum_nonneg)+ done show ?thesis proof (rule *) show "\(\(x,K) \ p2. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p2 12]) show "\(\(x,K) \ p1. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p1 11]) show "\(\(x,K) \ p2. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p2 22]) show "\(\(x,K) \ p1. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p1 21]) qed (use 0 in auto) qed then show ?thesis by (rule_tac x="\x. d1 x \ d2 x" in exI) (auto simp: fine_Int intro: \gauge d1\ \gauge d2\ d1 d2) qed then show ?thesis by (simp add: integrable_Cauchy) qed lemma integrable_straddle: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) s \ (h has_integral j) s \ \i - j\ < e \ (\x\s. g x \ f x \ f x \ h x)" shows "f integrable_on s" proof - let ?fs = "(\x. if x \ s then f x else 0)" have "?fs integrable_on cbox a b" for a b proof (rule integrable_straddle_interval) fix e::real assume "e > 0" then have *: "e/4 > 0" by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/4" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/4" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/4" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson define c where "c = (\i\Basis. min (a\i) (- (max Bg Bh)) *\<^sub>R i)" define d where "d = (\i\Basis. max (b\i) (max Bg Bh) *\<^sub>R i)" have "\norm (0 - x) < Bg; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBg: "ball 0 Bg \ cbox c d" by (auto simp: mem_box dist_norm) have "\norm (0 - x) < Bh; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBh: "ball 0 Bh \ cbox c d" by (auto simp: mem_box dist_norm) have ab_cd: "cbox a b \ cbox c d" by (auto simp: c_def d_def subset_box_imp) have **: "\ch cg ag ah::real. \\ah - ag\ \ \ch - cg\; \cg - i\ < e/4; \ch - j\ < e/4\ \ \ag - ah\ < e" using ij by arith show "\g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. g x \ (if x \ s then f x else 0) \ (if x \ s then f x else 0) \ h x)" proof (intro exI ballI conjI) have eq: "\x f g. (if x \ s then f x else 0) - (if x \ s then g x else 0) = (if x \ s then f x - g x else (0::real))" by auto have int_hg: "(\x. if x \ s then h x - g x else 0) integrable_on cbox a b" "(\x. if x \ s then h x - g x else 0) integrable_on cbox c d" by (metis (no_types) integrable_diff g h has_integral_integrable integrable_altD(1))+ show "(?gs has_integral integral (cbox a b) ?gs) (cbox a b)" "(?hs has_integral integral (cbox a b) ?hs) (cbox a b)" by (intro integrable_integral int_g int_h)+ then have "integral (cbox a b) ?gs \ integral (cbox a b) ?hs" using fgh by (force intro: has_integral_le) then have "0 \ integral (cbox a b) ?hs - integral (cbox a b) ?gs" by simp then have "\integral (cbox a b) ?hs - integral (cbox a b) ?gs\ \ \integral (cbox c d) ?hs - integral (cbox c d) ?gs\" apply (simp add: integral_diff [symmetric] int_g int_h) apply (subst abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF int_h int_g]]) using fgh apply (force simp: eq intro!: integral_subset_le [OF ab_cd int_hg])+ done then show "\integral (cbox a b) ?gs - integral (cbox a b) ?hs\ < e" using ** Bg ballBg Bh ballBh by blast show "\x. x \ cbox a b \ ?gs x \ ?fs x" "\x. x \ cbox a b \ ?fs x \ ?hs x" using fgh by auto qed qed then have int_f: "?fs integrable_on cbox a b" for a b by simp have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ abs (integral (cbox a b) ?fs - integral (cbox c d) ?fs) < e" if "0 < e" for e proof - have *: "e/3 > 0" using that by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/3" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where "Bg > 0" and Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/3" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where "Bh > 0" and Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/3" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson { fix a b c d :: 'n assume as: "ball 0 (max Bg Bh) \ cbox a b" "ball 0 (max Bg Bh) \ cbox c d" have **: "ball 0 Bg \ ball (0::'n) (max Bg Bh)" "ball 0 Bh \ ball (0::'n) (max Bg Bh)" by auto have *: "\ga gc ha hc fa fc. \\ga - i\ < e/3; \gc - i\ < e/3; \ha - j\ < e/3; \hc - j\ < e/3; ga \ fa; fa \ ha; gc \ fc; fc \ hc\ \ \fa - fc\ < e" using ij by arith have "abs (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" proof (rule *) show "\integral (cbox a b) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox c d) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox a b) ?hs - j\ < e/3" using "**" Bh as by blast show "\integral (cbox c d) ?hs - j\ < e/3" using "**" Bh as by blast qed (use int_f int_g int_h fgh in \simp_all add: integral_le\) } then show ?thesis apply (rule_tac x="max Bg Bh" in exI) using \Bg > 0\ by auto qed then show ?thesis unfolding integrable_alt[of f] real_norm_def by (blast intro: int_f) qed subsection \Adding integrals over several sets\ lemma has_integral_Un: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" "(f has_integral j) T" and neg: "negligible (S \ T)" shows "(f has_integral (i + j)) (S \ T)" unfolding has_integral_restrict_UNIV[symmetric, of f] proof (rule has_integral_spike[OF neg]) let ?f = "\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)" show "(?f has_integral i + j) UNIV" by (simp add: f has_integral_add) qed auto lemma integral_Un [simp]: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" "f integrable_on T" "negligible (S \ T)" shows "integral (S \ T) f = integral S f + integral T f" by (simp add: has_integral_Un assms integrable_integral integral_unique) lemma integrable_Un: fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "negligible (A \ B)" "f integrable_on A" "f integrable_on B" shows "f integrable_on (A \ B)" proof - from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" by (auto simp: integrable_on_def) from has_integral_Un[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) qed lemma integrable_Un': fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "f integrable_on A" "f integrable_on B" "negligible (A \ B)" "C = A \ B" shows "f integrable_on C" using integrable_Un[of A B f] assms by simp lemma has_integral_Union: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "finite \" and int: "\S. S \ \ \ (f has_integral (i S)) S" and neg: "pairwise (\S S'. negligible (S \ S')) \" shows "(f has_integral (sum i \)) (\\)" proof - let ?\ = "((\(a,b). a \ b) ` {(a,b). a \ \ \ b \ {y. y \ \ \ a \ y}})" have "((\x. if x \ \\ then f x else 0) has_integral sum i \) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (\ \ \)" by (simp add: \) moreover have "{(a, b). a \ \ \ b \ {y \ \. a \ y}} \ \ \ \" by auto ultimately show "finite ?\" by (blast intro: finite_subset[of _ "\ \ \"]) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ \\ then f x else 0) = (\A\\. if x \ A then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix S assume "S \ \" "x \ S" moreover then have "\b\\. x \ b \ b = S" using that by blast ultimately show "f x = (\A\\. if x \ A then f x else 0)" by (simp add: sum.delta[OF \]) qed next show "((\x. \A\\. if x \ A then f x else 0) has_integral (\A\\. i A)) UNIV" using int by (simp add: has_integral_restrict_UNIV has_integral_sum [OF \]) qed then show ?thesis using has_integral_restrict_UNIV by blast qed text \In particular adding integrals over a division, maybe not of an interval.\ lemma has_integral_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" and "\k. k \ \ \ (f has_integral (i k)) k" shows "(f has_integral (sum i \)) S" proof - note \ = division_ofD[OF assms(1)] have neg: "negligible (S \ s')" if "S \ \" "s' \ \" "S \ s'" for S s' proof - obtain a c b \ where obt: "S = cbox a b" "s' = cbox c \" by (meson \S \ \\ \s' \ \\ \(4)) from \(5)[OF that] show ?thesis unfolding obt interior_cbox by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval) qed show ?thesis unfolding \(6)[symmetric] by (auto intro: \ neg assms has_integral_Union pairwiseI) qed lemma integral_combine_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" "\k. k \ \ \ f integrable_on k" shows "integral S f = sum (\i. integral i f) \" by (meson assms integral_unique has_integral_combine_division has_integral_integrable_integral) lemma has_integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and \: "\ division_of K" and "K \ S" shows "(f has_integral (sum (\i. integral i f) \)) K" proof - have "f integrable_on L" if "L \ \" for L proof - have "L \ S" using \K \ S\ \ that by blast then show "f integrable_on L" using that by (metis (no_types) f \ division_ofD(4) integrable_on_subcbox) qed then show ?thesis by (meson \ has_integral_combine_division has_integral_integrable_integral) qed lemma integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\ division_of S" shows "integral S f = sum (\i. integral i f) \" using assms has_integral_combine_division_topdown by blast lemma integrable_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of S" and f: "\i. i \ \ \ f integrable_on i" shows "f integrable_on S" using f unfolding integrable_on_def by (metis has_integral_combine_division[OF \]) lemma integrable_on_subdivision: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of i" and f: "f integrable_on S" and "i \ S" shows "f integrable_on i" proof - have "f integrable_on i" if "i \ \" for i proof - have "i \ S" using assms that by auto then show "f integrable_on i" using that by (metis (no_types) \ f division_ofD(4) integrable_on_subcbox) qed then show ?thesis using \ integrable_combine_division by blast qed subsection \Also tagged divisions\ lemma has_integral_combine_tagged_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of S" and "\x k. (x,k) \ p \ (f has_integral (i k)) k" shows "(f has_integral (\(x,k)\p. i k)) S" proof - have *: "(f has_integral (\k\snd`p. integral k f)) S" proof - have "snd ` p division_of S" by (simp add: assms(1) division_of_tagged_division) with assms show ?thesis by (metis (mono_tags, lifting) has_integral_combine_division has_integral_integrable_integral imageE prod.collapse) qed also have "(\k\snd`p. integral k f) = (\(x, k)\p. integral k f)" by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null) (simp add: content_eq_0_interior) finally show ?thesis using assms by (auto simp add: has_integral_iff intro!: sum.cong) qed lemma integral_combine_tagged_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes p: "p tagged_division_of (cbox a b)" and f: "\x k. (x,k)\p \ f integrable_on k" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" by (simp add: has_integral_combine_tagged_division[OF p] integral_unique f integrable_integral) lemma has_integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and p: "p tagged_division_of (cbox a b)" shows "(f has_integral (sum (\(x,K). integral K f) p)) (cbox a b)" proof - have "(f has_integral integral K f) K" if "(x,K) \ p" for x K by (metis assms integrable_integral integrable_on_subcbox tagged_division_ofD(3,4) that) then show ?thesis by (simp add: has_integral_combine_tagged_division p) qed lemma integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on cbox a b" and "p tagged_division_of (cbox a b)" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" using assms by (auto intro: integral_unique [OF has_integral_combine_tagged_division_topdown]) subsection \Henstock's lemma\ lemma Henstock_lemma_part1: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on cbox a b" and "e > 0" and "gauge d" and less_e: "\p. \p tagged_division_of (cbox a b); d fine p\ \ norm (sum (\(x,K). content K *\<^sub>R f x) p - integral(cbox a b) f) < e" and p: "p tagged_partial_division_of (cbox a b)" "d fine p" shows "norm (sum (\(x,K). content K *\<^sub>R f x - integral K f) p) \ e" (is "?lhs \ e") proof (rule field_le_epsilon) fix k :: real assume "k > 0" let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" note p' = tagged_partial_division_ofD[OF p(1)] have "\(snd ` p) \ cbox a b" using p'(3) by fastforce then obtain q where q: "snd ` p \ q" and qdiv: "q division_of cbox a b" by (meson p(1) partial_division_extend_interval partial_division_of_tagged_division) note q' = division_ofD[OF qdiv] define r where "r = q - snd ` p" have "snd ` p \ r = {}" unfolding r_def by auto have "finite r" using q' unfolding r_def by auto have "\p. p tagged_division_of i \ d fine p \ norm (?SUM p - integral i f) < k / (real (card r) + 1)" if "i\r" for i proof - have gt0: "k / (real (card r) + 1) > 0" using \k > 0\ by simp have i: "i \ q" using that unfolding r_def by auto then obtain u v where uv: "i = cbox u v" using q'(4) by blast then have "cbox u v \ cbox a b" using i q'(2) by auto then have "f integrable_on cbox u v" by (rule integrable_subinterval[OF intf]) with integrable_integral[OF this, unfolded has_integral[of f]] obtain dd where "gauge dd" and dd: "\\. \\ tagged_division_of cbox u v; dd fine \\ \ norm (?SUM \ - integral (cbox u v) f) < k / (real (card r) + 1)" using gt0 by auto with gauge_Int[OF \gauge d\ \gauge dd\] obtain qq where qq: "qq tagged_division_of cbox u v" "(\x. d x \ dd x) fine qq" using fine_division_exists by blast with dd[of qq] show ?thesis by (auto simp: fine_Int uv) qed then obtain qq where qq: "\i. i \ r \ qq i tagged_division_of i \ d fine qq i \ norm (?SUM (qq i) - integral i f) < k / (real (card r) + 1)" by metis let ?p = "p \ \(qq ` r)" have "norm (?SUM ?p - integral (cbox a b) f) < e" proof (rule less_e) show "d fine ?p" by (metis (mono_tags, hide_lams) qq fine_Un fine_Union imageE p(2)) note ptag = tagged_partial_division_of_Union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \finite r\]]) show "\i. i \ r \ qq i tagged_division_of i" using qq by auto show "\i1 i2. \i1 \ r; i2 \ r; i1 \ i2\ \ interior i1 \ interior i2 = {}" by (simp add: q'(5) r_def) show "interior (\(snd ` p)) \ interior (\r) = {}" proof (rule Int_interior_Union_intervals [OF \finite r\]) show "open (interior (\(snd ` p)))" by blast show "\T. T \ r \ \a b. T = cbox a b" by (simp add: q'(4) r_def) have "interior T \ interior (\(snd ` p)) = {}" if "T \ r" for T proof (rule Int_interior_Union_intervals) show "\U. U \ snd ` p \ \a b. U = cbox a b" using q q'(4) by blast show "\U. U \ snd ` p \ interior T \ interior U = {}" by (metis DiffE q q'(5) r_def subsetD that) qed (use p' in auto) then show "\T. T \ r \ interior (\(snd ` p)) \ interior T = {}" by (metis Int_commute) qed qed moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" using qdiv q unfolding Union_Un_distrib[symmetric] r_def by auto ultimately show "?p tagged_division_of (cbox a b)" by fastforce qed then have "norm (?SUM p + (?SUM (\(qq ` r))) - integral (cbox a b) f) < e" proof (subst sum.union_inter_neutral[symmetric, OF \finite p\], safe) show "content L *\<^sub>R f x = 0" if "(x, L) \ p" "(x, L) \ qq K" "K \ r" for x K L proof - obtain u v where uv: "L = cbox u v" using \(x,L) \ p\ p'(4) by blast have "L \ K" using qq[OF that(3)] tagged_division_ofD(3) \(x,L) \ qq K\ by metis have "L \ snd ` p" using \(x,L) \ p\ image_iff by fastforce then have "L \ q" "K \ q" "L \ K" using that(1,3) q(1) unfolding r_def by auto with q'(5) have "interior L = {}" using interior_mono[OF \L \ K\] by blast then show "content L *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed show "finite (\(qq ` r))" by (meson finite_UN qq \finite r\ tagged_division_of_finite) qed moreover have "content M *\<^sub>R f x = 0" if x: "(x,M) \ qq K" "(x,M) \ qq L" and KL: "qq K \ qq L" and r: "K \ r" "L \ r" for x M K L proof - note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] obtain u v where uv: "M = cbox u v" using \(x, M) \ qq L\ \L \ r\ kl(2) by blast have empty: "interior (K \ L) = {}" by (metis DiffD1 interior_Int q'(5) r_def KL r) have "interior M = {}" by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r) then show "content M *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e" apply (subst (asm) sum.Union_comp) using qq by (force simp: split_paired_all)+ moreover have "content M *\<^sub>R f x = 0" if "K \ r" "L \ r" "K \ L" "qq K = qq L" "(x, M) \ qq K" for K L x M using tagged_division_ofD(6) qq that by (metis (no_types, lifting)) ultimately have less_e: "norm (?SUM p + sum (?SUM \ qq) r - integral (cbox a b) f) < e" proof (subst (asm) sum.reindex_nontrivial [OF \finite r\]) qed (auto simp: split_paired_all sum.neutral) have norm_le: "norm (cp - ip) \ e + k" if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i" for ir ip i cr cp::'a proof - from that show ?thesis using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] unfolding that(3)[symmetric] norm_minus_cancel by (auto simp add: algebra_simps) qed have "?lhs = norm (?SUM p - (\(x, k)\p. integral k f))" unfolding split_def sum_subtractf .. also have "\ \ e + k" proof (rule norm_le[OF less_e]) have lessk: "k * real (card r) / (1 + real (card r)) < k" using \k>0\ by (auto simp add: field_simps) have "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) \ (\x\r. k / (real (card r) + 1))" unfolding sum_subtractf[symmetric] by (force dest: qq intro!: sum_norm_le) also have "... < k" by (simp add: lessk add.commute mult.commute) finally show "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) < k" . next from q(1) have [simp]: "snd ` p \ q = q" by auto have "integral l f = 0" if inp: "(x, l) \ p" "(y, m) \ p" and ne: "(x, l) \ (y, m)" and "l = m" for x l y m proof - obtain u v where uv: "l = cbox u v" using inp p'(4) by blast have "content (cbox u v) = 0" unfolding content_eq_0_interior using that p(1) uv by (auto dest: tagged_partial_division_ofD) then show ?thesis using uv by blast qed then have "(\(x, K)\p. integral K f) = (\K\snd ` p. integral K f)" apply (subst sum.reindex_nontrivial [OF \finite p\]) unfolding split_paired_all split_def by auto then show "(\(x, k)\p. integral k f) + (\k\r. integral k f) = integral (cbox a b) f" unfolding integral_combine_division_topdown[OF intf qdiv] r_def using q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p", symmetric] by simp qed finally show "?lhs \ e + k" . qed lemma Henstock_lemma_part2: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes fed: "f integrable_on cbox a b" "e > 0" "gauge d" and less_e: "\\. \\ tagged_division_of (cbox a b); d fine \\ \ norm (sum (\(x,k). content k *\<^sub>R f x) \ - integral (cbox a b) f) < e" and tag: "p tagged_partial_division_of (cbox a b)" and "d fine p" shows "sum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" proof - have "finite p" using tag tagged_partial_division_ofD by blast then show ?thesis unfolding split_def proof (rule sum_norm_allsubsets_bound) fix Q assume Q: "Q \ p" then have fine: "d fine Q" by (simp add: \d fine p\ fine_subset) show "norm (\x\Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \ e" apply (rule Henstock_lemma_part1[OF fed less_e, unfolded split_def]) using Q tag tagged_partial_division_subset by (force simp add: fine)+ qed qed lemma Henstock_lemma: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes intf: "f integrable_on cbox a b" and "e > 0" obtains \ where "gauge \" and "\p. \p tagged_partial_division_of (cbox a b); \ fine p\ \ sum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" proof - have *: "e/(2 * (real DIM('n) + 1)) > 0" using \e > 0\ by simp with integrable_integral[OF intf, unfolded has_integral] obtain \ where "gauge \" and \: "\\. \\ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - integral (cbox a b) f) < e/(2 * (real DIM('n) + 1))" by metis show thesis proof (rule that [OF \gauge \\]) fix p assume p: "p tagged_partial_division_of cbox a b" "\ fine p" have "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) \ 2 * real DIM('n) * (e/(2 * (real DIM('n) + 1)))" using Henstock_lemma_part2[OF intf * \gauge \\ \ p] by metis also have "... < e" using \e > 0\ by (auto simp add: field_simps) finally show "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) < e" . qed qed subsection \Monotone convergence (bounded interval first)\ lemma bounded_increasing_convergent: fixes f :: "nat \ real" shows "\bounded (range f); \n. f n \ f (Suc n)\ \ \l. f \ l" using Bseq_mono_convergent[of f] incseq_Suc_iff[of f] by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) lemma monotone_convergence_interval: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on cbox a b" and le: "\k x. x \ cbox a b \ (f k x) \ f (Suc k) x" and fg: "\x. x \ cbox a b \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral (cbox a b) (f k)))" shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) \ integral (cbox a b) g) sequentially" proof (cases "content (cbox a b) = 0") case True then show ?thesis by auto next case False have fg1: "(f k x) \ (g x)" if x: "x \ cbox a b" for x k proof - have "\\<^sub>F j in sequentially. f k x \ f j x" proof (rule eventually_sequentiallyI [of k]) show "\j. k \ j \ f k x \ f j x" using le x by (force intro: transitive_stepwise_le) qed then show "f k x \ g x" using tendsto_lowerbound [OF fg] x trivial_limit_sequentially by blast qed have int_inc: "\n. integral (cbox a b) (f n) \ integral (cbox a b) (f (Suc n))" by (metis integral_le intf le) then obtain i where i: "(\k. integral (cbox a b) (f k)) \ i" using bounded_increasing_convergent bou by blast have "\k. \\<^sub>F x in sequentially. integral (cbox a b) (f k) \ integral (cbox a b) (f x)" unfolding eventually_sequentially by (force intro: transitive_stepwise_le int_inc) then have i': "\k. (integral(cbox a b) (f k)) \ i" using tendsto_le [OF trivial_limit_sequentially i] by blast have "(g has_integral i) (cbox a b)" unfolding has_integral real_norm_def proof clarify fix e::real assume e: "e > 0" have "\k. (\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))" using intf e by (auto simp: has_integral_integral has_integral) then obtain c where c: "\x. gauge (c x)" "\x \. \\ tagged_division_of cbox a b; c x fine \\ \ abs ((\(u,K)\\. content K *\<^sub>R f x u) - integral (cbox a b) (f x)) < e/2 ^ (x + 2)" by metis have "\r. \k\r. 0 \ i - (integral (cbox a b) (f k)) \ i - (integral (cbox a b) (f k)) < e/4" proof - have "e/4 > 0" using e by auto show ?thesis using LIMSEQ_D [OF i \e/4 > 0\] i' by auto qed then obtain r where r: "\k. r \ k \ 0 \ i - integral (cbox a b) (f k)" "\k. r \ k \ i - integral (cbox a b) (f k) < e/4" by metis have "\n\r. \k\n. 0 \ (g x) - (f k x) \ (g x) - (f k x) < e/(4 * content(cbox a b))" if "x \ cbox a b" for x proof - have "e/(4 * content (cbox a b)) > 0" by (simp add: False content_lt_nz e) with fg that LIMSEQ_D obtain N where "\n\N. norm (f n x - g x) < e/(4 * content (cbox a b))" by metis then show "\n\r. \k\n. 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" apply (rule_tac x="N + r" in exI) using fg1[OF that] by (auto simp add: field_simps) qed then obtain m where r_le_m: "\x. x \ cbox a b \ r \ m x" and m: "\x k. \x \ cbox a b; m x \ k\ \ 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" by metis define d where "d x = c (m x) x" for x show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R g x) - i) < e)" proof (rule exI, safe) show "gauge d" using c(1) unfolding gauge_def d_def by auto next fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "d fine \" note p'=tagged_division_ofD[OF ptag] obtain s where s: "\x. x \ \ \ m (fst x) \ s" by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) have *: "\a - d\ < e" if "\a - b\ \ e/4" "\b - c\ < e/2" "\c - d\ < e/4" for a b c d using that norm_triangle_lt[of "a - b" "b - c" "3* e/4"] norm_triangle_lt[of "a - b + (b - c)" "c - d" e] by (auto simp add: algebra_simps) show "\(\(x, k)\\. content k *\<^sub>R g x) - i\ < e" proof (rule *) have "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ (\i\\. \(case i of (x, K) \ content K *\<^sub>R g x) - (case i of (x, K) \ content K *\<^sub>R f (m x) x)\)" by (metis (mono_tags) sum_subtractf sum_abs) also have "... \ (\(x, k)\\. content k * (e/(4 * content (cbox a b))))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xk: "(x,K) \ \" with ptag have x: "x \ cbox a b" by blast then have "abs (content K * (g x - f (m x) x)) \ content K * (e/(4 * content (cbox a b)))" by (metis m[OF x] mult_nonneg_nonneg abs_of_nonneg less_eq_real_def measure_nonneg mult_left_mono order_refl) then show "\content K * g x - content K * f (m x) x\ \ content K * e/(4 * content (cbox a b))" by (simp add: algebra_simps) qed also have "... = (e/(4 * content (cbox a b))) * (\(x, k)\\. content k)" by (simp add: sum_distrib_left sum_divide_distrib split_def mult.commute) also have "... \ e/4" by (metis False additive_content_tagged_division [OF ptag] nonzero_mult_divide_mult_cancel_right order_refl times_divide_eq_left) finally show "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ e/4" . next have "norm ((\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))) \ norm (\j = 0..s. \(x,K)\{xk \ \. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))" apply (subst sum.group) using s by (auto simp: sum_subtractf split_def p'(1)) also have "\ < e/2" proof - have "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) \ (\i = 0..s. e/2 ^ (i + 2))" proof (rule sum_norm_le) fix t assume "t \ {0..s}" have "norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) = norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f t x - integral k (f t))" by (force intro!: sum.cong arg_cong[where f=norm]) also have "... \ e/2 ^ (t + 2)" proof (rule Henstock_lemma_part1 [OF intf]) show "{xk \ \. m (fst xk) = t} tagged_partial_division_of cbox a b" proof (rule tagged_partial_division_subset[of \]) show "\ tagged_partial_division_of cbox a b" using ptag tagged_division_of_def by blast qed auto show "c t fine {xk \ \. m (fst xk) = t}" using \d fine \\ by (auto simp: fine_def d_def) qed (use c e in auto) finally show "norm (\(x,K)\{xk \ \. m (fst xk) = t}. content K *\<^sub>R f (m x) x - integral K (f (m x))) \ e/2 ^ (t + 2)" . qed also have "... = (e/2/2) * (\i = 0..s. (1/2) ^ i)" by (simp add: sum_distrib_left field_simps) also have "\ < e/2" by (simp add: sum_gp mult_strict_left_mono[OF _ e]) finally show "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e/2" . qed finally show "\(\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))\ < e/2" by simp next have comb: "integral (cbox a b) (f y) = (\(x, k)\\. integral k (f y))" for y using integral_combine_tagged_division_topdown[OF intf ptag] by metis have f_le: "\y m n. \y \ cbox a b; n\m\ \ f m y \ f n y" using le by (auto intro: transitive_stepwise_le) have "(\(x, k)\\. integral k (f r)) \ (\(x, K)\\. integral K (f (m x)))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f r) \ integral K (f (m x))" proof (rule integral_le) show "f r integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f r y \ f (m x) y" if "y \ K" for y using that r_le_m[of x] p'(2-3)[OF xK] f_le by auto qed qed moreover have "(\(x, K)\\. integral K (f (m x))) \ (\(x, k)\\. integral k (f s))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f (m x)) \ integral K (f s)" proof (rule integral_le) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f s integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) y \ f s y" if "y \ K" for y using that s xK f_le p'(3) by fastforce qed qed moreover have "0 \ i - integral (cbox a b) (f r)" "i - integral (cbox a b) (f r) < e/4" using r by auto ultimately show "\(\(x,K)\\. integral K (f (m x))) - i\ < e/4" using comb i'[of s] by auto qed qed qed with i integral_unique show ?thesis by blast qed lemma monotone_convergence_increasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes int_f: "\k. (f k) integrable_on S" and "\k x. x \ S \ (f k x) \ (f (Suc k) x)" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral S (f k)))" shows "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" proof - have lem: "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" if f0: "\k x. x \ S \ 0 \ f k x" and int_f: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f k x \ f (Suc k) x" and lim: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" for f :: "nat \ 'n::euclidean_space \ real" and g S proof - have fg: "(f k x) \ (g x)" if "x \ S" for x k proof - have "\xa. k \ xa \ f k x \ f xa x" using le by (force intro: transitive_stepwise_le that) then show ?thesis using tendsto_lowerbound [OF lim [OF that]] eventually_sequentiallyI by force qed obtain i where i: "(\k. integral S (f k)) \ i" using bounded_increasing_convergent [OF bou] le int_f integral_le by blast have i': "(integral S (f k)) \ i" for k proof - have "\k. \x. x \ S \ \n\k. f k x \ f n x" using le by (force intro: transitive_stepwise_le) then show ?thesis using tendsto_lowerbound [OF i eventually_sequentiallyI trivial_limit_sequentially] by (meson int_f integral_le) qed let ?f = "(\k x. if x \ S then f k x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have int: "?f k integrable_on cbox a b" for a b k by (simp add: int_f integrable_altD(1)) have int': "\k a b. f k integrable_on cbox a b \ S" using int by (simp add: Int_commute integrable_restrict_Int) have g: "?g integrable_on cbox a b \ (\k. integral (cbox a b) (?f k)) \ integral (cbox a b) ?g" for a b proof (rule monotone_convergence_interval) have "norm (integral (cbox a b) (?f k)) \ norm (integral S (f k))" for k proof - have "0 \ integral (cbox a b) (?f k)" by (metis (no_types) integral_nonneg Int_iff f0 inf_commute integral_restrict_Int int') moreover have "0 \ integral S (f k)" by (simp add: integral_nonneg f0 int_f) moreover have "integral (S \ cbox a b) (f k) \ integral S (f k)" by (metis f0 inf_commute int' int_f integral_subset_le le_inf_iff order_refl) ultimately show ?thesis by (simp add: integral_restrict_Int) qed moreover obtain B where "\x. x \ range (\k. integral S (f k)) \ norm x \ B" using bou unfolding bounded_iff by blast ultimately show "bounded (range (\k. integral (cbox a b) (?f k)))" unfolding bounded_iff by (blast intro: order_trans) qed (use int le lim in auto) moreover have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?g - i) < e" if "0 < e" for e proof - have "e/4>0" using that by auto with LIMSEQ_D [OF i] obtain N where N: "\n. n \ N \ norm (integral S (f n) - i) < e/4" by metis with int_f[of N, unfolded has_integral_integral has_integral_alt'[of "f N"]] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (?f N) - integral S (f N)) < e/4" by (meson \0 < e/4\) have "norm (integral (cbox a b) ?g - i) < e" if ab: "ball 0 B \ cbox a b" for a b proof - obtain M where M: "\n. n \ M \ abs (integral (cbox a b) (?f n) - integral (cbox a b) ?g) < e/2" using \e > 0\ g by (fastforce simp add: dest!: LIMSEQ_D [where r = "e/2"]) have *: "\\ \ g. \\\ - i\ < e/2; \\ - g\ < e/2; \ \ \; \ \ i\ \ \g - i\ < e" unfolding real_inner_1_right by arith show "norm (integral (cbox a b) ?g - i) < e" unfolding real_norm_def proof (rule *) show "\integral (cbox a b) (?f N) - i\ < e/2" proof (rule abs_triangle_half_l) show "\integral (cbox a b) (?f N) - integral S (f N)\ < e/2/2" using B[OF ab] by simp show "abs (i - integral S (f N)) < e/2/2" using N by (simp add: abs_minus_commute) qed show "\integral (cbox a b) (?f (M + N)) - integral (cbox a b) ?g\ < e/2" by (metis le_add1 M[of "M + N"]) show "integral (cbox a b) (?f N) \ integral (cbox a b) (?f (M + N))" proof (intro ballI integral_le[OF int int]) fix x assume "x \ cbox a b" have "(f m x) \ (f n x)" if "x \ S" "n \ m" for m n proof (rule transitive_stepwise_le [OF \n \ m\ order_refl]) show "\u y z. \f u x \ f y x; f y x \ f z x\ \ f u x \ f z x" using dual_order.trans by blast qed (simp add: le \x \ S\) then show "(?f N)x \ (?f (M+N))x" by auto qed have "integral (cbox a b \ S) (f (M + N)) \ integral S (f (M + N))" by (metis Int_lower1 f0 inf_commute int' int_f integral_subset_le) then have "integral (cbox a b) (?f (M + N)) \ integral S (f (M + N))" by (metis (no_types) inf_commute integral_restrict_Int) also have "... \ i" using i'[of "M + N"] by auto finally show "integral (cbox a b) (?f (M + N)) \ i" . qed qed then show ?thesis using \0 < B\ by blast qed ultimately have "(g has_integral i) S" unfolding has_integral_alt' by auto then show ?thesis using has_integral_integrable_integral i integral_unique by metis qed have sub: "\k. integral S (\x. f k x - f 0 x) = integral S (f k) - integral S (f 0)" by (simp add: integral_diff int_f) have *: "\x m n. x \ S \ n\m \ f m x \ f n x" using assms(2) by (force intro: transitive_stepwise_le) have gf: "(\x. g x - f 0 x) integrable_on S \ ((\k. integral S (\x. f (Suc k) x - f 0 x)) \ integral S (\x. g x - f 0 x)) sequentially" proof (rule lem) show "\k. (\x. f (Suc k) x - f 0 x) integrable_on S" by (simp add: integrable_diff int_f) show "(\k. f (Suc k) x - f 0 x) \ g x - f 0 x" if "x \ S" for x proof - have "(\n. f (Suc n) x) \ g x" using LIMSEQ_ignore_initial_segment[OF fg[OF \x \ S\], of 1] by simp then show ?thesis by (simp add: tendsto_diff) qed show "bounded (range (\k. integral S (\x. f (Suc k) x - f 0 x)))" proof - obtain B where B: "\k. norm (integral S (f k)) \ B" using bou by (auto simp: bounded_iff) then have "norm (integral S (\x. f (Suc k) x - f 0 x)) \ B + norm (integral S (f 0))" for k unfolding sub by (meson add_le_cancel_right norm_triangle_le_diff) then show ?thesis unfolding bounded_iff by blast qed qed (use * in auto) then have "(\x. integral S (\xa. f (Suc x) xa - f 0 xa) + integral S (f 0)) \ integral S (\x. g x - f 0 x) + integral S (f 0)" by (auto simp add: tendsto_add) moreover have "(\x. g x - f 0 x + f 0 x) integrable_on S" using gf integrable_add int_f [of 0] by metis ultimately show ?thesis by (simp add: integral_diff int_f LIMSEQ_imp_Suc sub) qed lemma has_integral_monotone_convergence_increasing: fixes f :: "nat \ 'a::euclidean_space \ real" assumes f: "\k. (f k has_integral x k) s" assumes "\k x. x \ s \ f k x \ f (Suc k) x" assumes "\x. x \ s \ (\k. f k x) \ g x" assumes "x \ x'" shows "(g has_integral x') s" proof - have x_eq: "x = (\i. integral s (f i))" by (simp add: integral_unique[OF f]) then have x: "range(\k. integral s (f k)) = range x" by auto have *: "g integrable_on s \ (\k. integral s (f k)) \ integral s g" proof (intro monotone_convergence_increasing allI ballI assms) show "bounded (range(\k. integral s (f k)))" using x convergent_imp_bounded assms by metis qed (use f in auto) then have "integral s g = x'" by (intro LIMSEQ_unique[OF _ \x \ x'\]) (simp add: x_eq) with * show ?thesis by (simp add: has_integral_integral) qed lemma monotone_convergence_decreasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f (Suc k) x \ f k x" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" shows "g integrable_on S \ (\k. integral S (f k)) \ integral S g" proof - have *: "range(\k. integral S (\x. - f k x)) = (*\<^sub>R) (- 1) ` (range(\k. integral S (f k)))" by force have "(\x. - g x) integrable_on S \ (\k. integral S (\x. - f k x)) \ integral S (\x. - g x)" proof (rule monotone_convergence_increasing) show "\k. (\x. - f k x) integrable_on S" by (blast intro: integrable_neg intf) show "\k x. x \ S \ - f k x \ - f (Suc k) x" by (simp add: le) show "\x. x \ S \ (\k. - f k x) \ - g x" by (simp add: fg tendsto_minus) show "bounded (range(\k. integral S (\x. - f k x)))" using "*" bou bounded_scaling by auto qed then show ?thesis by (force dest: integrable_neg tendsto_minus) qed lemma integral_norm_bound_integral: fixes f :: "'n::euclidean_space \ 'a::banach" assumes int_f: "f integrable_on S" and int_g: "g integrable_on S" and le_g: "\x. x \ S \ norm (f x) \ g x" shows "norm (integral S f) \ integral S g" proof - have norm: "norm \ \ y + e" if "norm \ \ x" and "\x - y\ < e/2" and "norm (\ - \) < e/2" for e x y and \ \ :: 'a proof - have "norm (\ - \) < e/2" by (metis norm_minus_commute that(3)) moreover have "x \ y + e/2" using that(2) by linarith ultimately show ?thesis using that(1) le_less_trans[OF norm_triangle_sub[of \ \]] by (auto simp: less_imp_le) qed have lem: "norm (integral(cbox a b) f) \ integral (cbox a b) g" if f: "f integrable_on cbox a b" and g: "g integrable_on cbox a b" and nle: "\x. x \ cbox a b \ norm (f x) \ g x" for f :: "'n \ 'a" and g a b proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto with integrable_integral[OF f,unfolded has_integral[of f]] obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" by meson moreover from integrable_integral[OF g,unfolded has_integral[of g]] e obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2" by meson ultimately have "gauge (\x. \ x \ \ x)" using gauge_Int by blast with fine_division_exists obtain \ where p: "\ tagged_division_of cbox a b" "(\x. \ x \ \ x) fine \" by metis have "\ fine \" "\ fine \" using fine_Int p(2) by blast+ show "norm (integral (cbox a b) f) \ integral (cbox a b) g + e" proof (rule norm) have "norm (content K *\<^sub>R f x) \ content K *\<^sub>R g x" if "(x, K) \ \" for x K proof- have K: "x \ K" "K \ cbox a b" using \(x, K) \ \\ p(1) by blast+ obtain u v where "K = cbox u v" using \(x, K) \ \\ p(1) by blast moreover have "content K * norm (f x) \ content K * g x" - by (metis K subsetD dual_order.antisym measure_nonneg mult_zero_left nle not_le real_mult_le_cancel_iff2) + by (meson K(1) K(2) content_pos_le mult_left_mono nle subsetD) then show ?thesis by simp qed then show "norm (\(x, k)\\. content k *\<^sub>R f x) \ (\(x, k)\\. content k *\<^sub>R g x)" by (simp add: sum_norm_le split_def) show "norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" using \\ fine \\ \ p(1) by simp show "\(\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g\ < e/2" using \\ fine \\ \ p(1) by simp qed qed show ?thesis proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto let ?f = "(\x. if x \ S then f x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b using int_f int_g integrable_altD by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?f has_integral z) (cbox a b) \ norm (z - integral S f) < e/2" using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ cbox a b \ \z. (?g has_integral z) (cbox a b) \ norm (z - integral S g) < e/2" using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast obtain a b::'n where ab: "ball 0 Bf \ ball 0 Bg \ cbox a b" using ball_max_Un by (metis bounded_ball bounded_subset_cbox_symmetric) have "ball 0 Bf \ cbox a b" using ab by auto with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2" by meson have "ball 0 Bg \ cbox a b" using ab by auto with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2" by meson show "norm (integral S f) \ integral S g + e" proof (rule norm) show "norm (integral (cbox a b) ?f) \ integral (cbox a b) ?g" by (simp add: le_g lem[OF f g, of a b]) show "\integral (cbox a b) ?g - integral S g\ < e/2" using int_gw integral_unique w by auto show "norm (integral (cbox a b) ?f - integral S f) < e/2" using int_fz integral_unique z by blast qed qed qed lemma continuous_on_imp_absolutely_integrable_on: fixes f::"real \ 'a::banach" shows "continuous_on {a..b} f \ norm (integral {a..b} f) \ integral {a..b} (\x. norm (f x))" by (intro integral_norm_bound_integral integrable_continuous_real continuous_on_norm) auto lemma integral_bound: fixes f::"real \ 'a::banach" assumes "a \ b" assumes "continuous_on {a .. b} f" assumes "\t. t \ {a .. b} \ norm (f t) \ B" shows "norm (integral {a .. b} f) \ B * (b - a)" proof - note continuous_on_imp_absolutely_integrable_on[OF assms(2)] also have "integral {a..b} (\x. norm (f x)) \ integral {a..b} (\_. B)" by (rule integral_le) (auto intro!: integrable_continuous_real continuous_intros assms) also have "\ = B * (b - a)" using assms by simp finally show ?thesis . qed lemma integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "f integrable_on S" and g: "g integrable_on S" and fg: "\x. x \ S \ norm(f x) \ (g x)\k" shows "norm (integral S f) \ (integral S g)\k" proof - have "norm (integral S f) \ integral S ((\x. x \ k) \ g)" using integral_norm_bound_integral[OF f integrable_linear[OF g]] by (simp add: bounded_linear_inner_left fg) then show ?thesis unfolding o_def integral_component_eq[OF g] . qed lemma has_integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "(f has_integral i) S" and g: "(g has_integral j) S" and "\x. x \ S \ norm (f x) \ (g x)\k" shows "norm i \ j\k" using integral_norm_bound_integral_component[of f S g k] unfolding integral_unique[OF f] integral_unique[OF g] using assms by auto lemma uniformly_convergent_improper_integral: fixes f :: "'b \ real \ 'a :: {banach}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "\y x. y \ A \ x \ a \ norm (f y x) \ g x" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI', goal_cases) case (1 \) from G have "Cauchy G" by (auto intro!: convergent_Cauchy) with 1 obtain M where M: "dist (G (real m)) (G (real n)) < \" if "m \ M" "n \ M" for m n by (force simp: Cauchy_def) define M' where "M' = max (nat \a\) M" show ?case proof (rule exI[of _ M'], safe, goal_cases) case (1 x m n) have M': "M' \ a" "M' \ M" unfolding M'_def by linarith+ have int_g: "(g has_integral (G (real n) - G (real m))) {real m..real n}" using 1 M' by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: DERIV_subset[OF deriv]) have int_f: "f x integrable_on {a'..real n}" if "a' \ a" for a' using that 1 by (cases "a' \ real n") (auto intro: integrable) have "dist (integral {a..real m} (f x)) (integral {a..real n} (f x)) = norm (integral {a..real n} (f x) - integral {a..real m} (f x))" by (simp add: dist_norm norm_minus_commute) also have "integral {a..real m} (f x) + integral {real m..real n} (f x) = integral {a..real n} (f x)" using M' and 1 by (intro integral_combine int_f) auto hence "integral {a..real n} (f x) - integral {a..real m} (f x) = integral {real m..real n} (f x)" by (simp add: algebra_simps) also have "norm \ \ integral {real m..real n} g" using le 1 M' int_f int_g by (intro integral_norm_bound_integral) auto also from int_g have "integral {real m..real n} g = G (real n) - G (real m)" by (simp add: has_integral_iff) also have "\ \ dist (G m) (G n)" by (simp add: dist_norm) also from 1 and M' have "\ < \" by (intro M) auto finally show ?case . qed qed lemma uniformly_convergent_improper_integral': fixes f :: "'b \ real \ 'a :: {banach, real_normed_algebra}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof - from le obtain a'' where le: "\y x. y \ A \ x \ a'' \ norm (f y x) \ g x" by (auto simp: eventually_at_top_linorder) define a' where "a' = max a a''" have "uniformly_convergent_on A (\b x. integral {a'..real b} (f x))" proof (rule uniformly_convergent_improper_integral) fix t assume t: "t \ a'" hence "(G has_field_derivative g t) (at t within {a..})" by (intro deriv) (auto simp: a'_def) moreover have "{a'..} \ {a..}" unfolding a'_def by auto ultimately show "(G has_field_derivative g t) (at t within {a'..})" by (rule DERIV_subset) qed (insert le, auto simp: a'_def intro: integrable G) hence "uniformly_convergent_on A (\b x. integral {a..a'} (f x) + integral {a'..real b} (f x))" (is ?P) by (intro uniformly_convergent_add) auto also have "eventually (\x. \y\A. integral {a..a'} (f y) + integral {a'..x} (f y) = integral {a..x} (f y)) sequentially" by (intro eventually_mono [OF eventually_ge_at_top[of "nat \a'\"]] ballI integral_combine) (auto simp: a'_def intro: integrable) hence "?P \ ?thesis" by (intro uniformly_convergent_cong) simp_all finally show ?thesis . qed subsection \differentiation under the integral sign\ lemma integral_continuous_on_param: fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" shows "continuous_on U (\x. integral (cbox a b) (f x))" proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding continuous_on_def proof (safe intro!: tendstoI) fix e'::real and x assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) assume "x \ U" from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] obtain X0 where X0: "x \ X0" "open X0" and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" unfolding split_beta fst_conv snd_conv dist_norm by metis have "\\<^sub>F y in at x within U. y \ X0 \ U" using X0(1) X0(2) eventually_at_topological by auto then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" proof eventually_elim case (elim y) have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = norm (integral (cbox a b) (\t. f y t - f x t))" using elim \x \ U\ unfolding dist_norm by (subst integral_diff) (auto intro!: integrable_continuous continuous_intros) also have "\ \ e * content (cbox a b)" using elim \x \ U\ by (intro integrable_bound) (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] integrable_continuous continuous_intros) also have "\ < e'" using \0 < e'\ \e > 0\ by (auto simp: e_def field_split_simps) finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . qed qed qed (auto intro!: continuous_on_const) lemma leibniz_rule: fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes [intro]: "x0 \ U" assumes "convex U" shows "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" (is "(?F has_derivative ?dF) _") proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) note [continuous_intros] = continuous_on_compose2[OF cont_f1] fix e'::real assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] obtain X0 where X0: "x0 \ X0" "open X0" and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" unfolding split_beta fst_conv snd_conv by (metis dist_norm) note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] moreover have "\\<^sub>F x in at x0 within U. x \ X0" using \open X0\ \x0 \ X0\ eventually_at_topological by blast moreover have "\\<^sub>F x in at x0 within U. x \ x0" by (auto simp: eventually_at_filter) moreover have "\\<^sub>F x in at x0 within U. x \ U" by (auto simp: eventually_at_filter) ultimately show "\\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" proof eventually_elim case (elim x) from elim have "0 < norm (x - x0)" by simp have "closed_segment x0 x \ U" by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) from elim have [intro]: "x \ U" by auto have "?F x - ?F x0 - ?dF (x - x0) = integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" (is "_ = ?id") using \x \ x0\ by (subst blinfun_apply_integral integral_diff, auto intro!: integrable_diff integrable_f2 continuous_intros intro: integrable_continuous)+ also { fix t assume t: "t \ (cbox a b)" have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" using \closed_segment x0 x \ U\ \closed_segment x0 x \ X0\ by (force simp: closed_segment_def algebra_simps) from t have deriv: "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" if "y \ X0 \ U" for y unfolding has_vector_derivative_def[symmetric] using that \x \ X0\ by (intro has_derivative_subset[OF fx]) auto have "\x. x \ X0 \ U \ onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" using fx_bound t by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) from differentiable_bound_linearization[OF seg deriv this] X0 have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" by (auto simp add: ac_simps) } then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" by (intro integral_norm_bound_integral) (auto intro!: continuous_intros integrable_diff integrable_f2 intro: integrable_continuous) also have "\ = content (cbox a b) * e * norm (x - x0)" by simp also have "\ < e' * norm (x - x0)" proof (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) show "content (cbox a b) * e < e'" using \e' > 0\ by (simp add: divide_simps e_def not_less) qed finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . then show ?case by (auto simp: divide_simps) qed qed (rule blinfun.bounded_linear_right) qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) lemma has_vector_derivative_eq_has_derivative_blinfun: "(f has_vector_derivative f') (at x within U) \ (f has_derivative blinfun_scaleR_left f') (at x within U)" by (simp add: has_vector_derivative_def) lemma leibniz_rule_vector_derivative: fixes f::"real \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding has_vector_derivative_eq_has_derivative_blinfun proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_scaleR_left (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))) = blinfun_apply (blinfun_scaleR_left (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_vector_derivative_def\) qed lemma has_field_derivative_eq_has_derivative_blinfun: "(f has_field_derivative f') (at x within U) \ (f has_derivative blinfun_mult_right f') (at x within U)" by (simp add: has_field_derivative_def) lemma leibniz_rule_field_derivative: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_field_derivative_eq_has_derivative_blinfun proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_mult_right (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))) = blinfun_apply (blinfun_mult_right (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_field_derivative_def\) qed lemma leibniz_rule_field_differentiable: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes "\x. x \ U \ (f x) integrable_on cbox a b" assumes "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes "x0 \ U" "convex U" shows "(\x. integral (cbox a b) (f x)) field_differentiable at x0 within U" using leibniz_rule_field_derivative[OF assms] by (auto simp: field_differentiable_def) subsection \Exchange uniform limit and integral\ lemma uniform_limit_integral_cbox: fixes f::"'a \ 'b::euclidean_space \ 'c::banach" assumes u: "uniform_limit (cbox a b) f g F" assumes c: "\n. continuous_on (cbox a b) (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) (cbox a b)" "(g has_integral J) (cbox a b)" "(I \ J) F" proof - have fi[simp]: "f n integrable_on (cbox a b)" for n by (auto intro!: integrable_continuous assms) then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" by atomize_elim (auto simp: integrable_on_def intro!: choice) moreover have gi[simp]: "g integrable_on (cbox a b)" by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) then obtain J where J: "(g has_integral J) (cbox a b)" by blast moreover have "(I \ J) F" proof cases assume "content (cbox a b) = 0" hence "I = (\_. 0)" "J = 0" by (auto intro!: has_integral_unique I J) thus ?thesis by simp next assume content_nonzero: "content (cbox a b) \ 0" show ?thesis proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' = e/2" with \e > 0\ have "e' > 0" by simp then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" using u content_nonzero by (auto simp: uniform_limit_iff dist_norm zero_less_measure_iff) then show "\\<^sub>F n in F. dist (I n) J < e" proof eventually_elim case (elim n) have "I n = integral (cbox a b) (f n)" "J = integral (cbox a b) g" using I[of n] J by (simp_all add: integral_unique) then have "dist (I n) J = norm (integral (cbox a b) (\x. f n x - g x))" by (simp add: integral_diff dist_norm) also have "\ \ integral (cbox a b) (\x. (e' / content (cbox a b)))" using elim by (intro integral_norm_bound_integral) (auto intro!: integrable_diff) also have "\ < e" using \0 < e\ by (simp add: e'_def) finally show ?case . qed qed qed ultimately show ?thesis .. qed lemma uniform_limit_integral: fixes f::"'a \ 'b::ordered_euclidean_space \ 'c::banach" assumes u: "uniform_limit {a..b} f g F" assumes c: "\n. continuous_on {a..b} (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) {a..b}" "(g has_integral J) {a..b}" "(I \ J) F" by (metis interval_cbox assms uniform_limit_integral_cbox) subsection \Integration by parts\ lemma integration_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" proof - interpret bounded_bilinear prod by fact have "((\x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral (prod (f b) (g b) - prod (f a) (g a))) {a..b}" using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) (auto intro!: continuous_intros continuous_on has_vector_derivative) from has_integral_diff[OF this int] show ?thesis by (simp add: algebra_simps) qed lemma integration_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integration_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" proof - from int obtain I where "((\x. prod (f x) (g' x)) has_integral I) {a..b}" unfolding integrable_on_def by blast hence "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp from integration_by_parts_interior_strong[OF assms(1-7) this] show ?thesis unfolding integrable_on_def by blast qed lemma integrable_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) subsection \Integration by substitution\ lemma has_integral_substitution_general: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof - let ?F = "\x. integral {c..g x} f" have cont_int: "continuous_on {a..b} ?F" by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous_1 f integrable_continuous_real)+ have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b})" if "x \ {a..b} - s" for x proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl]) show "(g has_vector_derivative g' x) (at x within {a..b})" using deriv has_field_derivative_iff_has_vector_derivative that by blast show "((\x. integral {c..x} f) has_vector_derivative f (g x)) (at (g x) within g ` {a..b})" using that le subset by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f) qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x)" if "x \ {a..b} - (s \ {a,b})" for x using deriv[of x] that by (simp add: at_within_Icc_at o_def) have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using le cont_int s deriv cont_int by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast from this[of a] this[of b] le have cd: "c \ g a" "g b \ d" "c \ g b" "g a \ d" by auto have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f" proof cases assume "g a \ g b" note le = le this from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with le show ?thesis by (cases "g a = g b") (simp_all add: algebra_simps) next assume less: "\g a \ g b" then have "g a \ g b" by simp note le = le this from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with less show ?thesis by (simp_all add: algebra_simps) qed finally show ?thesis . qed lemma has_integral_substitution_strong: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" "g a \ g b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" using has_integral_substitution_general[OF s le(1) subset f g deriv] le(2) by (cases "g a = g b") auto lemma has_integral_substitution: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes "a \ b" "g a \ g b" "g ` {a..b} \ {c..d}" and "continuous_on {c..d} f" and "\x. x \ {a..b} \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" by (intro has_integral_substitution_strong[of "{}" a b g c d] assms) (auto intro: DERIV_continuous_on assms) lemma integral_shift: fixes f :: "real \ 'a::euclidean_space" assumes cont: "continuous_on {a + c..b + c} f" shows "integral {a..b} (f \ (\x. x + c)) = integral {a + c..b + c} f" proof (cases "a \ b") case True have "((\x. 1 *\<^sub>R f (x + c)) has_integral integral {a+c..b+c} f) {a..b}" using True cont by (intro has_integral_substitution[where c = "a + c" and d = "b + c"]) (auto intro!: derivative_eq_intros) thus ?thesis by (simp add: has_integral_iff o_def) qed auto subsection \Compute a double integral using iterated integrals and switching the order of integration\ lemma continuous_on_imp_integrable_on_Pair1: fixes f :: "_ \ 'b::banach" assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \ cbox a b" shows "(\y. f (x, y)) integrable_on (cbox c d)" proof - have "f \ (\y. (x, y)) integrable_on (cbox c d)" proof (intro integrable_continuous continuous_on_compose [OF _ continuous_on_subset [OF con]]) show "continuous_on (cbox c d) (Pair x)" by (simp add: continuous_on_Pair) show "Pair x ` cbox c d \ cbox (a,c) (b,d)" using x by blast qed then show ?thesis by (simp add: o_def) qed lemma integral_integrable_2dim: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) f" shows "(\x. integral (cbox c d) (\y. f (x,y))) integrable_on cbox a b" proof (cases "content(cbox c d) = 0") case True then show ?thesis by (simp add: True integrable_const) next case False have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" by (simp add: assms compact_cbox compact_uniformly_continuous) { fix x::'a and e::real assume x: "x \ cbox a b" and e: "0 < e" then have e2_gt: "0 < e/2 / content (cbox c d)" and e2_less: "e/2 / content (cbox c d) * content (cbox c d) < e" by (auto simp: False content_lt_nz e) then obtain dd where dd: "\x x'. \x\cbox (a, c) (b, d); x'\cbox (a, c) (b, d); norm (x' - x) < dd\ \ norm (f x' - f x) \ e/(2 * content (cbox c d))" "dd>0" using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e/(2 * content (cbox c d))"] by (auto simp: dist_norm intro: less_imp_le) have "\delta>0. \x'\cbox a b. norm (x' - x) < delta \ norm (integral (cbox c d) (\u. f (x', u) - f (x, u))) < e" using dd e2_gt assms x apply (rule_tac x=dd in exI) apply clarify apply (rule le_less_trans [OF integrable_bound e2_less]) apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) done } note * = this show ?thesis proof (rule integrable_continuous) show "continuous_on (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" by (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) qed qed lemma integral_split: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on (cbox a b)" and k: "k \ Basis" shows "integral (cbox a b) f = integral (cbox a b \ {x. x\k \ c}) f + integral (cbox a b \ {x. x\k \ c}) f" using k f by (auto simp: has_integral_integral intro: integral_unique [OF has_integral_split]) lemma integral_swap_operativeI: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes f: "continuous_on s f" and e: "0 < e" shows "operative conj True (\k. \a b c d. cbox (a,c) (b,d) \ k \ cbox (a,c) (b,d) \ s \ norm(integral (cbox (a,c) (b,d)) f - integral (cbox a b) (\x. integral (cbox c d) (\y. f((x,y))))) \ e * content (cbox (a,c) (b,d)))" proof (standard, auto) fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b assume *: "box (a, c) (b, d) = {}" and cb1: "cbox (u, w) (v, z) \ cbox (a, c) (b, d)" and cb2: "cbox (u, w) (v, z) \ s" then have c0: "content (cbox (a, c) (b, d)) = 0" using * unfolding content_eq_0_interior by simp have c0': "content (cbox (u, w) (v, z)) = 0" by (fact content_0_subset [OF c0 cb1]) show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u,w) (v,z))" using content_cbox_pair_eq0_D [OF c0'] by (force simp add: c0') next fix a::'a and c::'b and b::'a and d::'b and M::real and i::'a and j::'b and u::'a and v::'a and w::'b and z::'b assume ij: "(i,j) \ Basis" and n1: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. x \ (i,j) \ M} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and n2: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. M \ x \ (i,j)} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and subs: "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" "cbox (u,w) (v,z) \ s" have fcont: "continuous_on (cbox (u, w) (v, z)) f" using assms(1) continuous_on_subset subs(2) by blast then have fint: "f integrable_on cbox (u, w) (v, z)" by (metis integrable_continuous) consider "i \ Basis" "j=0" | "j \ Basis" "i=0" using ij by (auto simp: Euclidean_Space.Basis_prod_def) then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z))" (is ?normle) proof cases case 1 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v \ {x. x \ i \ M}) * content (cbox w z)) + e * (content (cbox u v \ {x. M \ x \ i}) * content (cbox w z))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v \ {x. x \ i \ M}) (\x. integral (cbox w z) (\y. f (x, y))) + integral (cbox u v \ {x. M \ x \ i}) (\x. integral (cbox w z) (\y. f (x, y)))" using 1 f subs integral_integrable_2dim continuous_on_subset by (blast intro: integral_split) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 1 subs apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\u. M\u"] setcomp_dot1 [of "\u. u\M"] Sigma_Int_Paircomp1) apply (simp_all add: interval_split ij flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split intro!: n2 [rule_format]) done next case 2 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v) * content (cbox w z \ {x. x \ j \ M})) + e * (content (cbox u v) * content (cbox w z \ {x. M \ x \ j}))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have "(\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) integrable_on cbox u v" "(\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y))) integrable_on cbox u v" using 2 subs apply (simp_all add: interval_split) apply (rule integral_integrable_2dim [OF continuous_on_subset [OF f]]; auto simp: cbox_Pair_eq interval_split [symmetric])+ done with 2 have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v) (\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) + integral (cbox u v) (\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y)))" by (simp add: integral_add [symmetric] integral_split [symmetric] continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 2 subs apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\u. M\u"] setcomp_dot2 [of "\u. u\M"] Sigma_Int_Paircomp2) apply (simp_all add: interval_split ij flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split intro!: n2 [rule_format]) done qed qed lemma integral_Pair_const: "integral (cbox (a,c) (b,d)) (\x. k) = integral (cbox a b) (\x. integral (cbox c d) (\y. k))" by (simp add: content_Pair) lemma integral_prod_continuous: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a, c) (b, d)) f" (is "continuous_on ?CBOX f") shows "integral (cbox (a, c) (b, d)) f = integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" proof (cases "content ?CBOX = 0") case True then show ?thesis by (auto simp: content_Pair) next case False then have cbp: "content ?CBOX > 0" using content_lt_nz by blast have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) = 0" proof (rule dense_eq0_I, simp) fix e :: real assume "0 < e" with \content ?CBOX > 0\ have "0 < e/content ?CBOX" by simp have f_int_acbd: "f integrable_on ?CBOX" by (rule integrable_continuous [OF assms]) { fix p assume p: "p division_of ?CBOX" then have "finite p" by blast define e' where "e' = e/content ?CBOX" with \0 < e\ \0 < e/content ?CBOX\ have "0 < e'" by simp interpret operative conj True "\k. \a' b' c' d'. cbox (a', c') (b', d') \ k \ cbox (a', c') (b', d') \ ?CBOX \ norm (integral (cbox (a', c') (b', d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f ((x, y))))) \ e' * content (cbox (a', c') (b', d'))" using assms \0 < e'\ by (rule integral_swap_operativeI) have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e' * content ?CBOX" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e' * content (cbox (u, w) (v, z))" using that division [of p "(a, c)" "(b, d)"] p \finite p\ by (auto simp add: comm_monoid_set_F_and) with False have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u, w) (v, z)) / content ?CBOX" using that by (simp add: e'_def) } note op_acbd = this { fix k::real and \ and u::'a and v w and z::'b and t1 t2 l assume k: "0 < k" and nf: "\x y u v. \x \ cbox a b; y \ cbox c d; u \ cbox a b; v\cbox c d; norm (u-x, v-y) < k\ \ norm (f(u,v) - f(x,y)) < e/(2 * (content ?CBOX))" and p_acbd: "\ tagged_division_of cbox (a,c) (b,d)" and fine: "(\x. ball x k) fine \" "((t1,t2), l) \ \" and uwvz_sub: "cbox (u,w) (v,z) \ l" "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" have t: "t1 \ cbox a b" "t2 \ cbox c d" by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ have ls: "l \ ball (t1,t2) k" using fine by (simp add: fine_def Ball_def) { fix x1 x2 assume xuvwz: "x1 \ cbox u v" "x2 \ cbox w z" then have x: "x1 \ cbox a b" "x2 \ cbox c d" using uwvz_sub by auto have "norm (x1 - t1, x2 - t2) = norm (t1 - x1, t2 - x2)" by (simp add: norm_Pair norm_minus_commute) also have "norm (t1 - x1, t2 - x2) < k" using xuvwz ls uwvz_sub unfolding ball_def by (force simp add: cbox_Pair_eq dist_norm ) finally have "norm (f (x1,x2) - f (t1,t2)) \ e/content ?CBOX/2" using nf [OF t x] by simp } note nf' = this have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" using f_int_acbd uwvz_sub integrable_on_subcbox by blast have f_int_uv: "\x. x \ cbox u v \ (\y. f (x,y)) integrable_on cbox w z" using assms continuous_on_subset uwvz_sub by (blast intro: continuous_on_imp_integrable_on_Pair1) have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\x. f (t1,t2))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) apply (auto simp: integrable_diff f_int_uwvz integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) done have int_integrable: "(\x. integral (cbox w z) (\y. f (x, y))) integrable_on cbox u v" using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast have normint_wz: "\x. x \ cbox u v \ norm (integral (cbox w z) (\y. f (x, y)) - integral (cbox w z) (\y. f (t1, t2))) \ e * content (cbox w z) / content (cbox (a, c) (b, d))/2" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) apply (auto simp: integrable_diff f_int_uv integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) done have "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)) - integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox w z) / content ?CBOX/2 * content (cbox u v)" using cbp \0 < e/content ?CBOX\ apply (intro integrable_bound [OF _ _ normint_wz]) apply (auto simp: field_split_simps integrable_diff int_integrable integrable_const) done also have "... \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp add: content_Pair field_split_simps) finally have 2: "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y))) - integral (cbox u v) (\x. integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp only: integral_diff [symmetric] int_integrable integrable_const) have norm_xx: "\x' = y'; norm(x - x') \ e/2; norm(y - y') \ e/2\ \ norm(x - y) \ e" for x::'c and y x' y' e using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] field_sum_of_halves by (simp add: norm_minus_commute) have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX" by (rule norm_xx [OF integral_Pair_const 1 2]) } note * = this have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" if "\x\?CBOX. \x'\?CBOX. norm (x' - x) < k \ norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k proof - obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)" and fine: "(\x. ball x k) fine p" using fine_division_exists \0 < k\ by blast show ?thesis using that fine ptag \0 < k\ by (auto simp: * intro: op_acbd [OF division_of_tagged_division [OF ptag]]) qed then show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" using compact_uniformly_continuous [OF assms compact_cbox] apply (simp add: uniformly_continuous_on_def dist_norm) apply (drule_tac x="e/2 / content?CBOX" in spec) using cbp \0 < e\ by (auto simp: zero_less_mult_iff) qed then show ?thesis by simp qed lemma integral_swap_2dim: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox (a, c) (b, d)) (\(x, y). f x y) = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" proof - have "((\(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" proof (rule has_integral_twiddle [of 1 prod.swap prod.swap "\(x,y). f y x" "integral (cbox (c, a) (d, b)) (\(x, y). f y x)", simplified]) show "\u v. content (prod.swap ` cbox u v) = content (cbox u v)" by (metis content_Pair mult.commute old.prod.exhaust swap_cbox_Pair) show "((\(x, y). f y x) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (cbox (c, a) (d, b))" by (simp add: assms integrable_continuous integrable_integral swap_continuous) qed (use isCont_swap in \fastforce+\) then show ?thesis by force qed theorem integral_swap_continuous: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" proof - have "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\(x, y). f x y)" using integral_prod_continuous [OF assms] by auto also have "... = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" by (rule integral_swap_2dim [OF assms]) also have "... = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" using integral_prod_continuous [OF swap_continuous] assms by auto finally show ?thesis . qed subsection \Definite integrals for exponential and power function\ lemma has_integral_exp_minus_to_infinity: assumes a: "a > 0" shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" proof - define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" { fix k :: nat assume k: "of_nat k \ c" from k a have "((\x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def by (subst has_integral_restrict) simp_all } note has_integral_f = this have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) have integral_f: "integral {c..} (f k) = (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" for k using integral_unique[OF has_integral_f[of k]] by simp have A: "(\x. exp (-a*x)) integrable_on {c..} \ (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" proof (intro monotone_convergence_increasing allI ballI) fix k ::nat have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) hence "(f k) integrable_on {c..of_real k}" by (rule integrable_eq) (simp add: f_def) then show "f k integrable_on {c..}" by (rule integrable_on_superset) (auto simp: f_def) next fix x assume x: "x \ {c..}" have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) also have "{nat \x\..} \ {k. x \ real k}" by auto also have "inf (principal \) (principal {k. \x \ real k}) = principal ({k. x \ real k} \ {k. \x \ real k})" by simp also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast finally have "inf sequentially (principal {k. \x \ real k}) = bot" by (simp add: inf.coboundedI1 bot_unique) with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def by (intro filterlim_If) auto next have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k proof (cases "c > of_nat k") case False hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" by (simp add: integral_f) also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = exp (- (a * c)) / a - exp (- (a * real k)) / a" using False a by (intro abs_of_nonneg) (simp_all add: field_simps) also have "\ \ exp (- a * c) / a" using a by simp finally show ?thesis . qed (insert a, simp_all add: integral_f) thus "bounded (range(\k. integral {c..} (f k)))" by (intro boundedI[of _ "exp (-a*c)/a"]) auto qed (auto simp: f_def) have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ (insert a, simp_all) moreover from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" by eventually_elim linarith hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" by eventually_elim (simp add: integral_f) ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" by (rule Lim_transform_eventually) from LIMSEQ_unique[OF conjunct2[OF A] this] have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp with conjunct1[OF A] show ?thesis by (simp add: has_integral_integral) qed lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast lemma has_integral_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" proof (cases "c = 0") case False define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" { fix k :: nat have "(f k has_integral F k) {0..c}" proof (cases "inverse (of_nat (Suc k)) \ c") case True { fix x assume x: "x \ inverse (1 + real k)" have "0 < inverse (1 + real k)" by simp also note x finally have "x > 0" . } note x = this hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" using True a by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const simp: has_field_derivative_iff_has_vector_derivative [symmetric]) with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all next case False thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto qed } note has_integral_f = this have integral_f: "integral {0..c} (f k) = F k" for k using has_integral_f[of k] by (rule integral_unique) have A: "(\x. x powr a) integrable_on {0..c} \ (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" proof (intro monotone_convergence_increasing ballI allI) fix k from has_integral_f[of k] show "f k integrable_on {0..c}" by (auto simp: integrable_on_def) next fix k :: nat and x :: real { assume x: "inverse (real (Suc k)) \ x" have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) also note x finally have "inverse (real (Suc (Suc k))) \ x" . } thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) next fix x assume x: "x \ {0..c}" show "(\k. f k x) \ x powr a" proof (cases "x = 0") case False with x have "x > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. x powr a = f k x) sequentially" by eventually_elim (insert x, simp add: f_def) moreover have "(\_. x powr a) \ x powr a" by simp ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed (simp_all add: f_def) next { fix k from a have "F k \ c powr (a + 1) / (a + 1)" by (auto simp add: F_def divide_simps) also from a have "F k \ 0" by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) hence "F k = abs (F k)" by simp finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . } thus "bounded (range(\k. integral {0..c} (f k)))" by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) qed from False c have "c > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = integral {0..c} (f k)) sequentially" by eventually_elim (simp add: integral_f F_def) moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1)" by simp ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" by (blast intro: Lim_transform_eventually) with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" by (blast intro: LIMSEQ_unique) with A show ?thesis by (simp add: has_integral_integral) qed (simp_all add: has_integral_refl) lemma integrable_on_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0..c}" using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast lemma has_integral_powr_to_inf: fixes a e :: real assumes "e < -1" "a > 0" shows "((\x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}" proof - define f :: "nat \ real \ real" where "f = (\n x. if x \ {a..n} then x powr e else 0)" define F where "F = (\x. x powr (e + 1) / (e + 1))" have has_integral_f: "(f n has_integral (F n - F a)) {a..}" if n: "n \ a" for n :: nat proof - from n assms have "((\x. x powr e) has_integral (F n - F a)) {a..n}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] F_def) hence "(f n has_integral (F n - F a)) {a..n}" by (rule has_integral_eq [rotated]) (simp add: f_def) thus "(f n has_integral (F n - F a)) {a..}" by (rule has_integral_on_superset) (auto simp: f_def) qed have integral_f: "integral {a..} (f n) = (if n \ a then F n - F a else 0)" for n :: nat proof (cases "n \ a") case True with has_integral_f[OF this] show ?thesis by (simp add: integral_unique) next case False have "(f n has_integral 0) {a}" by (rule has_integral_refl) hence "(f n has_integral 0) {a..}" by (rule has_integral_on_superset) (insert False, simp_all add: f_def) with False show ?thesis by (simp add: integral_unique) qed have *: "(\x. x powr e) integrable_on {a..} \ (\n. integral {a..} (f n)) \ integral {a..} (\x. x powr e)" proof (intro monotone_convergence_increasing allI ballI) fix n :: nat from assms have "(\x. x powr e) integrable_on {a..n}" by (auto intro!: integrable_continuous_real continuous_intros) hence "f n integrable_on {a..n}" by (rule integrable_eq) (auto simp: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat and x :: real show "f n x \ f (Suc n) x" by (auto simp: f_def) next fix x :: real assume x: "x \ {a..}" from filterlim_real_sequentially have "eventually (\n. real n \ x) at_top" by (simp add: filterlim_at_top) with x have "eventually (\n. f n x = x powr e) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ x powr e" by (simp add: tendsto_eventually) next have "norm (integral {a..} (f n)) \ -F a" for n :: nat proof (cases "n \ a") case True with assms have "a powr (e + 1) \ n powr (e + 1)" by (intro powr_mono2') simp_all with assms show ?thesis by (auto simp: divide_simps F_def integral_f) qed (insert assms, simp add: integral_f F_def field_split_simps) thus "bounded (range(\k. integral {a..} (f k)))" unfolding bounded_iff by (intro exI[of _ "-F a"]) auto qed from filterlim_real_sequentially have "eventually (\n. real n \ a) at_top" by (simp add: filterlim_at_top) hence "eventually (\n. F n - F a = integral {a..} (f n)) at_top" by eventually_elim (simp add: integral_f) moreover have "(\n. F n - F a) \ 0 / (e + 1) - F a" unfolding F_def by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr] filterlim_ident filterlim_real_sequentially | simp)+) hence "(\n. F n - F a) \ -F a" by simp ultimately have "(\n. integral {a..} (f n)) \ -F a" by (blast intro: Lim_transform_eventually) from conjunct2[OF *] and this have "integral {a..} (\x. x powr e) = -F a" by (rule LIMSEQ_unique) with conjunct1[OF *] show ?thesis by (simp add: has_integral_integral F_def) qed lemma has_integral_inverse_power_to_inf: fixes a :: real and n :: nat assumes "n > 1" "a > 0" shows "((\x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}" proof - from assms have "real_of_int (-int n) < -1" by simp note has_integral_powr_to_inf[OF this \a > 0\] also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) = 1 / (real (n - 1) * a powr (real (n - 1)))" using assms by (simp add: field_split_simps powr_add [symmetric] of_nat_diff) also from assms have "a powr (real (n - 1)) = a ^ (n - 1)" by (intro powr_realpow) finally show ?thesis by (rule has_integral_eq [rotated]) (insert assms, simp_all add: powr_minus powr_realpow field_split_simps) qed subsubsection \Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\ lemma integral_component_eq_cart[simp]: fixes f :: "'n::euclidean_space \ real^'m" assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_vec_nth,unfolded o_def] . lemma content_closed_interval: fixes a :: "'a::ordered_euclidean_space" assumes "a \ b" shows "content {a..b} = (\i\Basis. b\i - a\i)" using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a]) lemma integrable_const_ivl[intro]: fixes a::"'a::ordered_euclidean_space" shows "(\x. c) integrable_on {a..b}" unfolding cbox_interval[symmetric] by (rule integrable_const) lemma integrable_on_subinterval: fixes f :: "'n::ordered_euclidean_space \ 'a::banach" assumes "f integrable_on S" "{a..b} \ S" shows "f integrable_on {a..b}" using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval) end diff --git a/src/HOL/Analysis/Homeomorphism.thy b/src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy +++ b/src/HOL/Analysis/Homeomorphism.thy @@ -1,2252 +1,2252 @@ (* Title: HOL/Analysis/Homeomorphism.thy Author: LC Paulson (ported from HOL Light) *) section \Homeomorphism Theorems\ theory Homeomorphism imports Homotopy begin lemma homeomorphic_spheres': fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space" assumes "0 < \" and dimeq: "DIM('a) = DIM('b)" shows "(sphere a \) homeomorphic (sphere b \)" proof - obtain f :: "'a\'b" and g where "linear f" "linear g" and fg: "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" "\x. g(f x) = x" "\y. f(g y) = y" by (blast intro: isomorphisms_UNIV_UNIV [OF dimeq]) then have "continuous_on UNIV f" "continuous_on UNIV g" using linear_continuous_on linear_linear by blast+ then show ?thesis unfolding homeomorphic_minimal apply(rule_tac x="\x. b + f(x - a)" in exI) apply(rule_tac x="\x. a + g(x - b)" in exI) using assms apply (force intro: continuous_intros continuous_on_compose2 [of _ f] continuous_on_compose2 [of _ g] simp: dist_commute dist_norm fg) done qed lemma homeomorphic_spheres_gen: fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space" assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)" shows "(sphere a r homeomorphic sphere b s)" using assms homeomorphic_trans [OF homeomorphic_spheres homeomorphic_spheres'] by auto subsection \Homeomorphism of all convex compact sets with nonempty interior\ proposition fixes S :: "'a::euclidean_space set" assumes "compact S" and 0: "0 \ rel_interior S" and star: "\x. x \ S \ open_segment 0 x \ rel_interior S" shows starlike_compact_projective1_0: "S - rel_interior S homeomorphic sphere 0 1 \ affine hull S" (is "?SMINUS homeomorphic ?SPHER") and starlike_compact_projective2_0: "S homeomorphic cball 0 1 \ affine hull S" (is "S homeomorphic ?CBALL") proof - have starI: "(u *\<^sub>R x) \ rel_interior S" if "x \ S" "0 \ u" "u < 1" for x u proof (cases "x=0 \ u=0") case True with 0 show ?thesis by force next case False with that show ?thesis by (auto simp: in_segment intro: star [THEN subsetD]) qed have "0 \ S" using assms rel_interior_subset by auto define proj where "proj \ \x::'a. x /\<^sub>R norm x" have eqI: "x = y" if "proj x = proj y" "norm x = norm y" for x y using that by (force simp: proj_def) then have iff_eq: "\x y. (proj x = proj y \ norm x = norm y) \ x = y" by blast have projI: "x \ affine hull S \ proj x \ affine hull S" for x by (metis \0 \ S\ affine_hull_span_0 hull_inc span_mul proj_def) have nproj1 [simp]: "x \ 0 \ norm(proj x) = 1" for x by (simp add: proj_def) have proj0_iff [simp]: "proj x = 0 \ x = 0" for x by (simp add: proj_def) have cont_proj: "continuous_on (UNIV - {0}) proj" unfolding proj_def by (rule continuous_intros | force)+ have proj_spherI: "\x. \x \ affine hull S; x \ 0\ \ proj x \ ?SPHER" by (simp add: projI) have "bounded S" "closed S" using \compact S\ compact_eq_bounded_closed by blast+ have inj_on_proj: "inj_on proj (S - rel_interior S)" proof fix x y assume x: "x \ S - rel_interior S" and y: "y \ S - rel_interior S" and eq: "proj x = proj y" then have xynot: "x \ 0" "y \ 0" "x \ S" "y \ S" "x \ rel_interior S" "y \ rel_interior S" using 0 by auto consider "norm x = norm y" | "norm x < norm y" | "norm x > norm y" by linarith then show "x = y" proof cases assume "norm x = norm y" with iff_eq eq show "x = y" by blast next assume *: "norm x < norm y" have "x /\<^sub>R norm x = norm x *\<^sub>R (x /\<^sub>R norm x) /\<^sub>R norm (norm x *\<^sub>R (x /\<^sub>R norm x))" by force then have "proj ((norm x / norm y) *\<^sub>R y) = proj x" by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) then have [simp]: "(norm x / norm y) *\<^sub>R y = x" by (rule eqI) (simp add: \y \ 0\) have no: "0 \ norm x / norm y" "norm x / norm y < 1" using * by (auto simp: field_split_simps) then show "x = y" using starI [OF \y \ S\ no] xynot by auto next assume *: "norm x > norm y" have "y /\<^sub>R norm y = norm y *\<^sub>R (y /\<^sub>R norm y) /\<^sub>R norm (norm y *\<^sub>R (y /\<^sub>R norm y))" by force then have "proj ((norm y / norm x) *\<^sub>R x) = proj y" by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) then have [simp]: "(norm y / norm x) *\<^sub>R x = y" by (rule eqI) (simp add: \x \ 0\) have no: "0 \ norm y / norm x" "norm y / norm x < 1" using * by (auto simp: field_split_simps) then show "x = y" using starI [OF \x \ S\ no] xynot by auto qed qed have "\surf. homeomorphism (S - rel_interior S) ?SPHER proj surf" proof (rule homeomorphism_compact) show "compact (S - rel_interior S)" using \compact S\ compact_rel_boundary by blast show "continuous_on (S - rel_interior S) proj" using 0 by (blast intro: continuous_on_subset [OF cont_proj]) show "proj ` (S - rel_interior S) = ?SPHER" proof show "proj ` (S - rel_interior S) \ ?SPHER" using 0 by (force simp: hull_inc projI intro: nproj1) show "?SPHER \ proj ` (S - rel_interior S)" proof (clarsimp simp: proj_def) fix x assume "x \ affine hull S" and nox: "norm x = 1" then have "x \ 0" by auto obtain d where "0 < d" and dx: "(d *\<^sub>R x) \ rel_frontier S" and ri: "\e. \0 \ e; e < d\ \ (e *\<^sub>R x) \ rel_interior S" using ray_to_rel_frontier [OF \bounded S\ 0] \x \ affine hull S\ \x \ 0\ by auto show "x \ (\x. x /\<^sub>R norm x) ` (S - rel_interior S)" proof show "x = d *\<^sub>R x /\<^sub>R norm (d *\<^sub>R x)" using \0 < d\ by (auto simp: nox) show "d *\<^sub>R x \ S - rel_interior S" using dx \closed S\ by (auto simp: rel_frontier_def) qed qed qed qed (rule inj_on_proj) then obtain surf where surf: "homeomorphism (S - rel_interior S) ?SPHER proj surf" by blast then have cont_surf: "continuous_on (proj ` (S - rel_interior S)) surf" by (auto simp: homeomorphism_def) have surf_nz: "\x. x \ ?SPHER \ surf x \ 0" by (metis "0" DiffE homeomorphism_def imageI surf) have cont_nosp: "continuous_on (?SPHER) (\x. norm x *\<^sub>R ((surf o proj) x))" proof (intro continuous_intros) show "continuous_on (sphere 0 1 \ affine hull S) proj" by (rule continuous_on_subset [OF cont_proj], force) show "continuous_on (proj ` (sphere 0 1 \ affine hull S)) surf" by (intro continuous_on_subset [OF cont_surf]) (force simp: homeomorphism_image1 [OF surf] dest: proj_spherI) qed have surfpS: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ S" by (metis (full_types) DiffE \0 \ S\ homeomorphism_def image_eqI norm_zero proj_spherI real_vector.scale_zero_left scaleR_one surf) have *: "\y. norm y = 1 \ y \ affine hull S \ x = surf (proj y)" if "x \ S" "x \ rel_interior S" for x proof - have "proj x \ ?SPHER" by (metis (full_types) "0" hull_inc proj_spherI that) moreover have "surf (proj x) = x" by (metis Diff_iff homeomorphism_def surf that) ultimately show ?thesis by (metis \\x. x \ ?SPHER \ surf x \ 0\ hull_inc inverse_1 local.proj_def norm_sgn projI scaleR_one sgn_div_norm that(1)) qed have surfp_notin: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ rel_interior S" by (metis (full_types) DiffE one_neq_zero homeomorphism_def image_eqI norm_zero proj_spherI surf) have no_sp_im: "(\x. norm x *\<^sub>R surf (proj x)) ` (?SPHER) = S - rel_interior S" by (auto simp: surfpS image_def Bex_def surfp_notin *) have inj_spher: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?SPHER" proof fix x y assume xy: "x \ ?SPHER" "y \ ?SPHER" and eq: " norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" then have "norm x = 1" "norm y = 1" "x \ affine hull S" "y \ affine hull S" using 0 by auto with eq show "x = y" by (simp add: proj_def) (metis surf xy homeomorphism_def) qed have co01: "compact ?SPHER" by (simp add: compact_Int_closed) show "?SMINUS homeomorphic ?SPHER" using homeomorphic_def surf by blast have proj_scaleR: "\a x. 0 < a \ proj (a *\<^sub>R x) = proj x" by (simp add: proj_def) have cont_sp0: "continuous_on (affine hull S - {0}) (surf o proj)" proof (rule continuous_on_compose [OF continuous_on_subset [OF cont_proj]]) show "continuous_on (proj ` (affine hull S - {0})) surf" using homeomorphism_image1 proj_spherI surf by (intro continuous_on_subset [OF cont_surf]) fastforce qed auto obtain B where "B>0" and B: "\x. x \ S \ norm x \ B" by (metis compact_imp_bounded \compact S\ bounded_pos_less less_eq_real_def) have cont_nosp: "continuous (at x within ?CBALL) (\x. norm x *\<^sub>R surf (proj x))" if "norm x \ 1" "x \ affine hull S" for x proof (cases "x=0") case True have "(norm \ 0) (at 0 within cball 0 1 \ affine hull S)" by (simp add: tendsto_norm_zero eventually_at) with True show ?thesis apply (simp add: continuous_within) apply (rule lim_null_scaleR_bounded [where B=B]) using B \0 < B\ local.proj_def projI surfpS by (auto simp: eventually_at) next case False then have "\\<^sub>F x in at x. (x \ affine hull S - {0}) = (x \ affine hull S)" by (force simp: False eventually_at) moreover have "continuous (at x within affine hull S - {0}) (\x. surf (proj x))" using cont_sp0 False that by (auto simp add: continuous_on_eq_continuous_within) ultimately have *: "continuous (at x within affine hull S) (\x. surf (proj x))" by (simp add: continuous_within Lim_transform_within_set continuous_on_eq_continuous_within) show ?thesis by (intro continuous_within_subset [where s = "affine hull S", OF _ Int_lower2] continuous_intros *) qed have cont_nosp2: "continuous_on ?CBALL (\x. norm x *\<^sub>R ((surf o proj) x))" by (simp add: continuous_on_eq_continuous_within cont_nosp) have "norm y *\<^sub>R surf (proj y) \ S" if "y \ cball 0 1" and yaff: "y \ affine hull S" for y proof (cases "y=0") case True then show ?thesis by (simp add: \0 \ S\) next case False then have "norm y *\<^sub>R surf (proj y) = norm y *\<^sub>R surf (proj (y /\<^sub>R norm y))" by (simp add: proj_def) have "norm y \ 1" using that by simp have "surf (proj (y /\<^sub>R norm y)) \ S" using False local.proj_def nproj1 projI surfpS yaff by blast then have "surf (proj y) \ S" by (simp add: False proj_def) then show "norm y *\<^sub>R surf (proj y) \ S" by (metis dual_order.antisym le_less_linear norm_ge_zero rel_interior_subset scaleR_one starI subset_eq \norm y \ 1\) qed moreover have "x \ (\x. norm x *\<^sub>R surf (proj x)) ` (?CBALL)" if "x \ S" for x proof (cases "x=0") case True with that hull_inc show ?thesis by fastforce next case False then have psp: "proj (surf (proj x)) = proj x" by (metis homeomorphism_def hull_inc proj_spherI surf that) have nxx: "norm x *\<^sub>R proj x = x" by (simp add: False local.proj_def) have affineI: "(1 / norm (surf (proj x))) *\<^sub>R x \ affine hull S" by (metis \0 \ S\ affine_hull_span_0 hull_inc span_clauses(4) that) have sproj_nz: "surf (proj x) \ 0" by (metis False proj0_iff psp) then have "proj x = proj (proj x)" by (metis False nxx proj_scaleR zero_less_norm_iff) moreover have scaleproj: "\a r. r *\<^sub>R proj a = (r / norm a) *\<^sub>R a" by (simp add: divide_inverse local.proj_def) ultimately have "(norm (surf (proj x)) / norm x) *\<^sub>R x \ rel_interior S" by (metis (no_types) sproj_nz divide_self_if hull_inc norm_eq_zero nproj1 projI psp scaleR_one surfp_notin that) then have "(norm (surf (proj x)) / norm x) \ 1" using starI [OF that] by (meson starI [OF that] le_less_linear norm_ge_zero zero_le_divide_iff) then have nole: "norm x \ norm (surf (proj x))" by (simp add: le_divide_eq_1) let ?inx = "x /\<^sub>R norm (surf (proj x))" show ?thesis proof show "x = norm ?inx *\<^sub>R surf (proj ?inx)" by (simp add: field_simps) (metis inverse_eq_divide nxx positive_imp_inverse_positive proj_scaleR psp scaleproj sproj_nz zero_less_norm_iff) qed (auto simp: field_split_simps nole affineI) qed ultimately have im_cball: "(\x. norm x *\<^sub>R surf (proj x)) ` ?CBALL = S" by blast have inj_cball: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?CBALL" proof fix x y assume "x \ ?CBALL" "y \ ?CBALL" and eq: "norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" then have x: "x \ affine hull S" and y: "y \ affine hull S" using 0 by auto show "x = y" proof (cases "x=0 \ y=0") case True then show "x = y" using eq proj_spherI surf_nz x y by force next case False with x y have speq: "surf (proj x) = surf (proj y)" by (metis eq homeomorphism_apply2 proj_scaleR proj_spherI surf zero_less_norm_iff) then have "norm x = norm y" by (metis \x \ affine hull S\ \y \ affine hull S\ eq proj_spherI real_vector.scale_cancel_right surf_nz) moreover have "proj x = proj y" by (metis (no_types) False speq homeomorphism_apply2 proj_spherI surf x y) ultimately show "x = y" using eq eqI by blast qed qed have co01: "compact ?CBALL" by (simp add: compact_Int_closed) show "S homeomorphic ?CBALL" using homeomorphic_compact [OF co01 cont_nosp2 [unfolded o_def] im_cball inj_cball] homeomorphic_sym by blast qed corollary fixes S :: "'a::euclidean_space set" assumes "compact S" and a: "a \ rel_interior S" and star: "\x. x \ S \ open_segment a x \ rel_interior S" shows starlike_compact_projective1: "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" and starlike_compact_projective2: "S homeomorphic cball a 1 \ affine hull S" proof - have 1: "compact ((+) (-a) ` S)" by (meson assms compact_translation) have 2: "0 \ rel_interior ((+) (-a) ` S)" using a rel_interior_translation [of "- a" S] by (simp cong: image_cong_simp) have 3: "open_segment 0 x \ rel_interior ((+) (-a) ` S)" if "x \ ((+) (-a) ` S)" for x proof - have "x+a \ S" using that by auto then have "open_segment a (x+a) \ rel_interior S" by (metis star) then show ?thesis using open_segment_translation [of a 0 x] using rel_interior_translation [of "- a" S] by (fastforce simp add: ac_simps image_iff cong: image_cong_simp) qed have "S - rel_interior S homeomorphic ((+) (-a) ` S) - rel_interior ((+) (-a) ` S)" by (metis rel_interior_translation translation_diff homeomorphic_translation) also have "... homeomorphic sphere 0 1 \ affine hull ((+) (-a) ` S)" by (rule starlike_compact_projective1_0 [OF 1 2 3]) also have "... = (+) (-a) ` (sphere a 1 \ affine hull S)" by (metis affine_hull_translation left_minus sphere_translation translation_Int) also have "... homeomorphic sphere a 1 \ affine hull S" using homeomorphic_translation homeomorphic_sym by blast finally show "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" . have "S homeomorphic ((+) (-a) ` S)" by (metis homeomorphic_translation) also have "... homeomorphic cball 0 1 \ affine hull ((+) (-a) ` S)" by (rule starlike_compact_projective2_0 [OF 1 2 3]) also have "... = (+) (-a) ` (cball a 1 \ affine hull S)" by (metis affine_hull_translation left_minus cball_translation translation_Int) also have "... homeomorphic cball a 1 \ affine hull S" using homeomorphic_translation homeomorphic_sym by blast finally show "S homeomorphic cball a 1 \ affine hull S" . qed corollary starlike_compact_projective_special: assumes "compact S" and cb01: "cball (0::'a::euclidean_space) 1 \ S" and scale: "\x u. \x \ S; 0 \ u; u < 1\ \ u *\<^sub>R x \ S - frontier S" shows "S homeomorphic (cball (0::'a::euclidean_space) 1)" proof - have "ball 0 1 \ interior S" using cb01 interior_cball interior_mono by blast then have 0: "0 \ rel_interior S" by (meson centre_in_ball subsetD interior_subset_rel_interior le_numeral_extra(2) not_le) have [simp]: "affine hull S = UNIV" using \ball 0 1 \ interior S\ by (auto intro!: affine_hull_nonempty_interior) have star: "open_segment 0 x \ rel_interior S" if "x \ S" for x proof fix p assume "p \ open_segment 0 x" then obtain u where "x \ 0" and u: "0 \ u" "u < 1" and p: "u *\<^sub>R x = p" by (auto simp: in_segment) then show "p \ rel_interior S" using scale [OF that u] closure_subset frontier_def interior_subset_rel_interior by fastforce qed show ?thesis using starlike_compact_projective2_0 [OF \compact S\ 0 star] by simp qed lemma homeomorphic_convex_lemma: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "compact S" "convex T" "compact T" and affeq: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \ S homeomorphic T" proof (cases "rel_interior S = {} \ rel_interior T = {}") case True then show ?thesis by (metis Diff_empty affeq \convex S\ \convex T\ aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty) next case False then obtain a b where a: "a \ rel_interior S" and b: "b \ rel_interior T" by auto have starS: "\x. x \ S \ open_segment a x \ rel_interior S" using rel_interior_closure_convex_segment a \convex S\ closure_subset subsetCE by blast have starT: "\x. x \ T \ open_segment b x \ rel_interior T" using rel_interior_closure_convex_segment b \convex T\ closure_subset subsetCE by blast let ?aS = "(+) (-a) ` S" and ?bT = "(+) (-b) ` T" have 0: "0 \ affine hull ?aS" "0 \ affine hull ?bT" by (metis a b subsetD hull_inc image_eqI left_minus rel_interior_subset)+ have subs: "subspace (span ?aS)" "subspace (span ?bT)" by (rule subspace_span)+ moreover have "dim (span ((+) (- a) ` S)) = dim (span ((+) (- b) ` T))" by (metis 0 aff_dim_translation_eq aff_dim_zero affeq dim_span nat_int) ultimately obtain f g where "linear f" "linear g" and fim: "f ` span ?aS = span ?bT" and gim: "g ` span ?bT = span ?aS" and fno: "\x. x \ span ?aS \ norm(f x) = norm x" and gno: "\x. x \ span ?bT \ norm(g x) = norm x" and gf: "\x. x \ span ?aS \ g(f x) = x" and fg: "\x. x \ span ?bT \ f(g x) = x" by (rule isometries_subspaces) blast have [simp]: "continuous_on A f" for A using \linear f\ linear_conv_bounded_linear linear_continuous_on by blast have [simp]: "continuous_on B g" for B using \linear g\ linear_conv_bounded_linear linear_continuous_on by blast have eqspanS: "affine hull ?aS = span ?aS" by (metis a affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) have eqspanT: "affine hull ?bT = span ?bT" by (metis b affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) have "S homeomorphic cball a 1 \ affine hull S" by (rule starlike_compact_projective2 [OF \compact S\ a starS]) also have "... homeomorphic (+) (-a) ` (cball a 1 \ affine hull S)" by (metis homeomorphic_translation) also have "... = cball 0 1 \ (+) (-a) ` (affine hull S)" by (auto simp: dist_norm) also have "... = cball 0 1 \ span ?aS" using eqspanS affine_hull_translation by blast also have "... homeomorphic cball 0 1 \ span ?bT" proof (rule homeomorphicI) show fim1: "f ` (cball 0 1 \ span ?aS) = cball 0 1 \ span ?bT" proof show "f ` (cball 0 1 \ span ?aS) \ cball 0 1 \ span ?bT" using fim fno by auto show "cball 0 1 \ span ?bT \ f ` (cball 0 1 \ span ?aS)" by clarify (metis IntI fg gim gno image_eqI mem_cball_0) qed show "g ` (cball 0 1 \ span ?bT) = cball 0 1 \ span ?aS" proof show "g ` (cball 0 1 \ span ?bT) \ cball 0 1 \ span ?aS" using gim gno by auto show "cball 0 1 \ span ?aS \ g ` (cball 0 1 \ span ?bT)" by clarify (metis IntI fim1 gf image_eqI) qed qed (auto simp: fg gf) also have "... = cball 0 1 \ (+) (-b) ` (affine hull T)" using eqspanT affine_hull_translation by blast also have "... = (+) (-b) ` (cball b 1 \ affine hull T)" by (auto simp: dist_norm) also have "... homeomorphic (cball b 1 \ affine hull T)" by (metis homeomorphic_translation homeomorphic_sym) also have "... homeomorphic T" by (metis starlike_compact_projective2 [OF \compact T\ b starT] homeomorphic_sym) finally have 1: "S homeomorphic T" . have "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" by (rule starlike_compact_projective1 [OF \compact S\ a starS]) also have "... homeomorphic (+) (-a) ` (sphere a 1 \ affine hull S)" by (metis homeomorphic_translation) also have "... = sphere 0 1 \ (+) (-a) ` (affine hull S)" by (auto simp: dist_norm) also have "... = sphere 0 1 \ span ?aS" using eqspanS affine_hull_translation by blast also have "... homeomorphic sphere 0 1 \ span ?bT" proof (rule homeomorphicI) show fim1: "f ` (sphere 0 1 \ span ?aS) = sphere 0 1 \ span ?bT" proof show "f ` (sphere 0 1 \ span ?aS) \ sphere 0 1 \ span ?bT" using fim fno by auto show "sphere 0 1 \ span ?bT \ f ` (sphere 0 1 \ span ?aS)" by clarify (metis IntI fg gim gno image_eqI mem_sphere_0) qed show "g ` (sphere 0 1 \ span ?bT) = sphere 0 1 \ span ?aS" proof show "g ` (sphere 0 1 \ span ?bT) \ sphere 0 1 \ span ?aS" using gim gno by auto show "sphere 0 1 \ span ?aS \ g ` (sphere 0 1 \ span ?bT)" by clarify (metis IntI fim1 gf image_eqI) qed qed (auto simp: fg gf) also have "... = sphere 0 1 \ (+) (-b) ` (affine hull T)" using eqspanT affine_hull_translation by blast also have "... = (+) (-b) ` (sphere b 1 \ affine hull T)" by (auto simp: dist_norm) also have "... homeomorphic (sphere b 1 \ affine hull T)" by (metis homeomorphic_translation homeomorphic_sym) also have "... homeomorphic T - rel_interior T" by (metis starlike_compact_projective1 [OF \compact T\ b starT] homeomorphic_sym) finally have 2: "S - rel_interior S homeomorphic T - rel_interior T" . show ?thesis using 1 2 by blast qed lemma homeomorphic_convex_compact_sets: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "compact S" "convex T" "compact T" and affeq: "aff_dim S = aff_dim T" shows "S homeomorphic T" using homeomorphic_convex_lemma [OF assms] assms by (auto simp: rel_frontier_def) lemma homeomorphic_rel_frontiers_convex_bounded_sets: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "bounded S" "convex T" "bounded T" and affeq: "aff_dim S = aff_dim T" shows "rel_frontier S homeomorphic rel_frontier T" using assms homeomorphic_convex_lemma [of "closure S" "closure T"] by (simp add: rel_frontier_def convex_rel_interior_closure) subsection\Homeomorphisms between punctured spheres and affine sets\ text\Including the famous stereoscopic projection of the 3-D sphere to the complex plane\ text\The special case with centre 0 and radius 1\ lemma homeomorphic_punctured_affine_sphere_affine_01: assumes "b \ sphere 0 1" "affine T" "0 \ T" "b \ T" "affine p" and affT: "aff_dim T = aff_dim p + 1" shows "(sphere 0 1 \ T) - {b} homeomorphic p" proof - have [simp]: "norm b = 1" "b\b = 1" using assms by (auto simp: norm_eq_1) have [simp]: "T \ {v. b\v = 0} \ {}" using \0 \ T\ by auto have [simp]: "\ T \ {v. b\v = 0}" using \norm b = 1\ \b \ T\ by auto define f where "f \ \x. 2 *\<^sub>R b + (2 / (1 - b\x)) *\<^sub>R (x - b)" define g where "g \ \y. b + (4 / (norm y ^ 2 + 4)) *\<^sub>R (y - 2 *\<^sub>R b)" have fg[simp]: "\x. \x \ T; b\x = 0\ \ f (g x) = x" unfolding f_def g_def by (simp add: algebra_simps field_split_simps add_nonneg_eq_0_iff) have no: "(norm (f x))\<^sup>2 = 4 * (1 + b \ x) / (1 - b \ x)" if "norm x = 1" and "b \ x \ 1" for x using that sum_sqs_eq [of 1 "b \ x"] apply (simp flip: dot_square_norm add: norm_eq_1 nonzero_eq_divide_eq) apply (simp add: f_def vector_add_divide_simps inner_simps) apply (auto simp add: field_split_simps inner_commute) done have [simp]: "\u::real. 8 + u * (u * 8) = u * 16 \ u=1" by algebra have gf[simp]: "\x. \norm x = 1; b \ x \ 1\ \ g (f x) = x" unfolding g_def no by (auto simp: f_def field_split_simps) have g1: "norm (g x) = 1" if "x \ T" and "b \ x = 0" for x using that apply (simp only: g_def) apply (rule power2_eq_imp_eq) apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps) apply (simp add: algebra_simps inner_commute) done have ne1: "b \ g x \ 1" if "x \ T" and "b \ x = 0" for x using that unfolding g_def apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps add_nonneg_eq_0_iff) apply (auto simp: algebra_simps) done have "subspace T" by (simp add: assms subspace_affine) have gT: "\x. \x \ T; b \ x = 0\ \ g x \ T" unfolding g_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) have "f ` {x. norm x = 1 \ b\x \ 1} \ {x. b\x = 0}" unfolding f_def using \norm b = 1\ norm_eq_1 by (force simp: field_simps inner_add_right inner_diff_right) moreover have "f ` T \ T" unfolding f_def using assms \subspace T\ by (auto simp add: inner_add_right inner_diff_right mem_affine_3_minus subspace_mul) moreover have "{x. b\x = 0} \ T \ f ` ({x. norm x = 1 \ b\x \ 1} \ T)" by clarify (metis (mono_tags) IntI ne1 fg gT g1 imageI mem_Collect_eq) ultimately have imf: "f ` ({x. norm x = 1 \ b\x \ 1} \ T) = {x. b\x = 0} \ T" by blast have no4: "\y. b\y = 0 \ norm ((y\y + 4) *\<^sub>R b + 4 *\<^sub>R (y - 2 *\<^sub>R b)) = y\y + 4" apply (rule power2_eq_imp_eq) apply (simp_all flip: dot_square_norm) apply (auto simp: power2_eq_square algebra_simps inner_commute) done have [simp]: "\x. \norm x = 1; b \ x \ 1\ \ b \ f x = 0" by (simp add: f_def algebra_simps field_split_simps) have [simp]: "\x. \x \ T; norm x = 1; b \ x \ 1\ \ f x \ T" unfolding f_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) have "g ` {x. b\x = 0} \ {x. norm x = 1 \ b\x \ 1}" unfolding g_def apply (clarsimp simp: no4 vector_add_divide_simps divide_simps add_nonneg_eq_0_iff dot_square_norm [symmetric]) apply (auto simp: algebra_simps) done moreover have "g ` T \ T" unfolding g_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) moreover have "{x. norm x = 1 \ b\x \ 1} \ T \ g ` ({x. b\x = 0} \ T)" by clarify (metis (mono_tags, lifting) IntI gf image_iff imf mem_Collect_eq) ultimately have img: "g ` ({x. b\x = 0} \ T) = {x. norm x = 1 \ b\x \ 1} \ T" by blast have aff: "affine ({x. b\x = 0} \ T)" by (blast intro: affine_hyperplane assms) have contf: "continuous_on ({x. norm x = 1 \ b\x \ 1} \ T) f" unfolding f_def by (rule continuous_intros | force)+ have contg: "continuous_on ({x. b\x = 0} \ T) g" unfolding g_def by (rule continuous_intros | force simp: add_nonneg_eq_0_iff)+ have "(sphere 0 1 \ T) - {b} = {x. norm x = 1 \ (b\x \ 1)} \ T" using \norm b = 1\ by (auto simp: norm_eq_1) (metis vector_eq \b\b = 1\) also have "... homeomorphic {x. b\x = 0} \ T" by (rule homeomorphicI [OF imf img contf contg]) auto also have "... homeomorphic p" proof (rule homeomorphic_affine_sets [OF aff \affine p\]) show "aff_dim ({x. b \ x = 0} \ T) = aff_dim p" by (simp add: Int_commute aff_dim_affine_Int_hyperplane [OF \affine T\] affT) qed finally show ?thesis . qed theorem homeomorphic_punctured_affine_sphere_affine: fixes a :: "'a :: euclidean_space" assumes "0 < r" "b \ sphere a r" "affine T" "a \ T" "b \ T" "affine p" and aff: "aff_dim T = aff_dim p + 1" shows "(sphere a r \ T) - {b} homeomorphic p" proof - have "a \ b" using assms by auto then have inj: "inj (\x::'a. x /\<^sub>R norm (a - b))" by (simp add: inj_on_def) have "((sphere a r \ T) - {b}) homeomorphic (+) (-a) ` ((sphere a r \ T) - {b})" by (rule homeomorphic_translation) also have "... homeomorphic (*\<^sub>R) (inverse r) ` (+) (- a) ` (sphere a r \ T - {b})" by (metis \0 < r\ homeomorphic_scaling inverse_inverse_eq inverse_zero less_irrefl) also have "... = sphere 0 1 \ ((*\<^sub>R) (inverse r) ` (+) (- a) ` T) - {(b - a) /\<^sub>R r}" using assms by (auto simp: dist_norm norm_minus_commute divide_simps) also have "... homeomorphic p" using assms affine_translation [symmetric, of "- a"] aff_dim_translation_eq [of "- a"] by (intro homeomorphic_punctured_affine_sphere_affine_01) (auto simp: dist_norm norm_minus_commute affine_scaling inj) finally show ?thesis . qed corollary homeomorphic_punctured_sphere_affine: fixes a :: "'a :: euclidean_space" assumes "0 < r" and b: "b \ sphere a r" and "affine T" and affS: "aff_dim T + 1 = DIM('a)" shows "(sphere a r - {b}) homeomorphic T" using homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by auto corollary homeomorphic_punctured_sphere_hyperplane: fixes a :: "'a :: euclidean_space" assumes "0 < r" and b: "b \ sphere a r" and "c \ 0" shows "(sphere a r - {b}) homeomorphic {x::'a. c \ x = d}" using assms by (intro homeomorphic_punctured_sphere_affine) (auto simp: affine_hyperplane of_nat_diff) proposition homeomorphic_punctured_sphere_affine_gen: fixes a :: "'a :: euclidean_space" assumes "convex S" "bounded S" and a: "a \ rel_frontier S" and "affine T" and affS: "aff_dim S = aff_dim T + 1" shows "rel_frontier S - {a} homeomorphic T" proof - obtain U :: "'a set" where "affine U" "convex U" and affdS: "aff_dim U = aff_dim S" using choose_affine_subset [OF affine_UNIV aff_dim_geq] by (meson aff_dim_affine_hull affine_affine_hull affine_imp_convex) have "S \ {}" using assms by auto then obtain z where "z \ U" by (metis aff_dim_negative_iff equals0I affdS) then have bne: "ball z 1 \ U \ {}" by force then have [simp]: "aff_dim(ball z 1 \ U) = aff_dim U" using aff_dim_convex_Int_open [OF \convex U\ open_ball] by (fastforce simp add: Int_commute) have "rel_frontier S homeomorphic rel_frontier (ball z 1 \ U)" by (rule homeomorphic_rel_frontiers_convex_bounded_sets) (auto simp: \affine U\ affine_imp_convex convex_Int affdS assms) also have "... = sphere z 1 \ U" using convex_affine_rel_frontier_Int [of "ball z 1" U] by (simp add: \affine U\ bne) finally have "rel_frontier S homeomorphic sphere z 1 \ U" . then obtain h k where him: "h ` rel_frontier S = sphere z 1 \ U" and kim: "k ` (sphere z 1 \ U) = rel_frontier S" and hcon: "continuous_on (rel_frontier S) h" and kcon: "continuous_on (sphere z 1 \ U) k" and kh: "\x. x \ rel_frontier S \ k(h(x)) = x" and hk: "\y. y \ sphere z 1 \ U \ h(k(y)) = y" unfolding homeomorphic_def homeomorphism_def by auto have "rel_frontier S - {a} homeomorphic (sphere z 1 \ U) - {h a}" proof (rule homeomorphicI) show h: "h ` (rel_frontier S - {a}) = sphere z 1 \ U - {h a}" using him a kh by auto metis show "k ` (sphere z 1 \ U - {h a}) = rel_frontier S - {a}" by (force simp: h [symmetric] image_comp o_def kh) qed (auto intro: continuous_on_subset hcon kcon simp: kh hk) also have "... homeomorphic T" by (rule homeomorphic_punctured_affine_sphere_affine) (use a him in \auto simp: affS affdS \affine T\ \affine U\ \z \ U\\) finally show ?thesis . qed text\ When dealing with AR, ANR and ANR later, it's useful to know that every set is homeomorphic to a closed subset of a convex set, and if the set is locally compact we can take the convex set to be the universe.\ proposition homeomorphic_closedin_convex: fixes S :: "'m::euclidean_space set" assumes "aff_dim S < DIM('n)" obtains U and T :: "'n::euclidean_space set" where "convex U" "U \ {}" "closedin (top_of_set U) T" "S homeomorphic T" proof (cases "S = {}") case True then show ?thesis by (rule_tac U=UNIV and T="{}" in that) auto next case False then obtain a where "a \ S" by auto obtain i::'n where i: "i \ Basis" "i \ 0" using SOME_Basis Basis_zero by force have "0 \ affine hull ((+) (- a) ` S)" by (simp add: \a \ S\ hull_inc) then have "dim ((+) (- a) ` S) = aff_dim ((+) (- a) ` S)" by (simp add: aff_dim_zero) also have "... < DIM('n)" by (simp add: aff_dim_translation_eq_subtract assms cong: image_cong_simp) finally have dd: "dim ((+) (- a) ` S) < DIM('n)" by linarith have span: "span {x. i \ x = 0} = {x. i \ x = 0}" using span_eq_iff [symmetric, of "{x. i \ x = 0}"] subspace_hyperplane [of i] by simp have "dim ((+) (- a) ` S) \ dim {x. i \ x = 0}" using dd by (simp add: dim_hyperplane [OF \i \ 0\]) then obtain T where "subspace T" and Tsub: "T \ {x. i \ x = 0}" and dimT: "dim T = dim ((+) (- a) ` S)" by (rule choose_subspace_of_subspace) (simp add: span) have "subspace (span ((+) (- a) ` S))" using subspace_span by blast then obtain h k where "linear h" "linear k" and heq: "h ` span ((+) (- a) ` S) = T" and keq:"k ` T = span ((+) (- a) ` S)" and hinv [simp]: "\x. x \ span ((+) (- a) ` S) \ k(h x) = x" and kinv [simp]: "\x. x \ T \ h(k x) = x" by (auto simp: dimT intro: isometries_subspaces [OF _ \subspace T\] dimT) have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B using \linear h\ \linear k\ linear_continuous_on linear_conv_bounded_linear by blast+ have ihhhh[simp]: "\x. x \ S \ i \ h (x - a) = 0" using Tsub [THEN subsetD] heq span_superset by fastforce have "sphere 0 1 - {i} homeomorphic {x. i \ x = 0}" proof (rule homeomorphic_punctured_sphere_affine) show "affine {x. i \ x = 0}" by (auto simp: affine_hyperplane) show "aff_dim {x. i \ x = 0} + 1 = int DIM('n)" using i by clarsimp (metis DIM_positive Suc_pred add.commute of_nat_Suc) qed (use i in auto) then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \ x = 0} f g" by (force simp: homeomorphic_def) show ?thesis proof have "h ` (+) (- a) ` S \ T" using heq span_superset span_linear_image by blast then have "g ` h ` (+) (- a) ` S \ g ` {x. i \ x = 0}" using Tsub by (simp add: image_mono) also have "... \ sphere 0 1 - {i}" by (simp add: fg [unfolded homeomorphism_def]) finally have gh_sub_sph: "(g \ h) ` (+) (- a) ` S \ sphere 0 1 - {i}" by (metis image_comp) then have gh_sub_cb: "(g \ h) ` (+) (- a) ` S \ cball 0 1" by (metis Diff_subset order_trans sphere_cball) have [simp]: "\u. u \ S \ norm (g (h (u - a))) = 1" using gh_sub_sph [THEN subsetD] by (auto simp: o_def) show "convex (ball 0 1 \ (g \ h) ` (+) (- a) ` S)" by (meson ball_subset_cball convex_intermediate_ball gh_sub_cb sup.bounded_iff sup.cobounded1) show "closedin (top_of_set (ball 0 1 \ (g \ h) ` (+) (- a) ` S)) ((g \ h) ` (+) (- a) ` S)" unfolding closedin_closed by (rule_tac x="sphere 0 1" in exI) auto have ghcont: "continuous_on ((\x. x - a) ` S) (\x. g (h x))" by (rule continuous_on_compose2 [OF homeomorphism_cont2 [OF fg] hcont], force) have kfcont: "continuous_on ((\x. g (h (x - a))) ` S) (\x. k (f x))" proof (rule continuous_on_compose2 [OF kcont]) show "continuous_on ((\x. g (h (x - a))) ` S) f" using homeomorphism_cont1 [OF fg] gh_sub_sph by (fastforce intro: continuous_on_subset) qed auto have "S homeomorphic (+) (- a) ` S" by (fact homeomorphic_translation) also have "\ homeomorphic (g \ h) ` (+) (- a) ` S" apply (simp add: homeomorphic_def homeomorphism_def cong: image_cong_simp) apply (rule_tac x="g \ h" in exI) apply (rule_tac x="k \ f" in exI) apply (auto simp: ghcont kfcont span_base homeomorphism_apply2 [OF fg] image_comp cong: image_cong_simp) done finally show "S homeomorphic (g \ h) ` (+) (- a) ` S" . qed auto qed subsection\Locally compact sets in an open set\ text\ Locally compact sets are closed in an open set and are homeomorphic to an absolutely closed set if we have one more dimension to play with.\ lemma locally_compact_open_Int_closure: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains T where "open T" "S = T \ closure S" proof - have "\x\S. \T v u. u = S \ T \ x \ u \ u \ v \ v \ S \ open T \ compact v" by (metis assms locally_compact openin_open) then obtain t v where tv: "\x. x \ S \ v x \ S \ open (t x) \ compact (v x) \ (\u. x \ u \ u \ v x \ u = S \ t x)" by metis then have o: "open (\(t ` S))" by blast have "S = \ (v ` S)" using tv by blast also have "... = \(t ` S) \ closure S" proof show "\(v ` S) \ \(t ` S) \ closure S" by clarify (meson IntD2 IntI UN_I closure_subset subsetD tv) have "t x \ closure S \ v x" if "x \ S" for x proof - have "t x \ closure S \ closure (t x \ S)" by (simp add: open_Int_closure_subset that tv) also have "... \ v x" by (metis Int_commute closure_minimal compact_imp_closed that tv) finally show ?thesis . qed then show "\(t ` S) \ closure S \ \(v ` S)" by blast qed finally have e: "S = \(t ` S) \ closure S" . show ?thesis by (rule that [OF o e]) qed lemma locally_compact_closedin_open: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains T where "open T" "closedin (top_of_set T) S" by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int) lemma locally_compact_homeomorphism_projection_closed: assumes "locally compact S" obtains T and f :: "'a \ 'a :: euclidean_space \ 'b :: euclidean_space" where "closed T" "homeomorphism S T f fst" proof (cases "closed S") case True show ?thesis proof show "homeomorphism S (S \ {0}) (\x. (x, 0)) fst" by (auto simp: homeomorphism_def continuous_intros) qed (use True closed_Times in auto) next case False obtain U where "open U" and US: "U \ closure S = S" by (metis locally_compact_open_Int_closure [OF assms]) with False have Ucomp: "-U \ {}" using closure_eq by auto have [simp]: "closure (- U) = -U" by (simp add: \open U\ closed_Compl) define f :: "'a \ 'a \ 'b" where "f \ \x. (x, One /\<^sub>R setdist {x} (- U))" have "continuous_on U (\x. (x, One /\<^sub>R setdist {x} (- U)))" proof (intro continuous_intros continuous_on_setdist) show "\x\U. setdist {x} (- U) \ 0" by (simp add: Ucomp setdist_eq_0_sing_1) qed then have homU: "homeomorphism U (f`U) f fst" by (auto simp: f_def homeomorphism_def image_iff continuous_intros) have cloS: "closedin (top_of_set U) S" by (metis US closed_closure closedin_closed_Int) have cont: "isCont ((\x. setdist {x} (- U)) o fst) z" for z :: "'a \ 'b" by (rule continuous_at_compose continuous_intros continuous_at_setdist)+ have setdist1D: "setdist {a} (- U) *\<^sub>R b = One \ setdist {a} (- U) \ 0" for a::'a and b::'b by force have *: "r *\<^sub>R b = One \ b = (1 / r) *\<^sub>R One" for r and b::'b by (metis One_non_0 nonzero_divide_eq_eq real_vector.scale_eq_0_iff real_vector.scale_scale scaleR_one) have "\a b::'b. setdist {a} (- U) *\<^sub>R b = One \ (a,b) \ (\x. (x, (1 / setdist {x} (- U)) *\<^sub>R One)) ` U" by (metis (mono_tags, lifting) "*" ComplI image_eqI setdist1D setdist_sing_in_set) then have "f ` U = (\z. (setdist {fst z} (- U) *\<^sub>R snd z)) -` {One}" by (auto simp: f_def setdist_eq_0_sing_1 field_simps Ucomp) then have clfU: "closed (f ` U)" by (force intro: continuous_intros cont [unfolded o_def] continuous_closed_vimage) have "closed (f ` S)" by (metis closedin_closed_trans [OF _ clfU] homeomorphism_imp_closed_map [OF homU cloS]) then show ?thesis by (metis US homU homeomorphism_of_subsets inf_sup_ord(1) that) qed lemma locally_compact_closed_Int_open: fixes S :: "'a :: euclidean_space set" shows "locally compact S \ (\U V. closed U \ open V \ S = U \ V)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (metis closed_closure inf_commute locally_compact_open_Int_closure) show "?rhs \ ?lhs" by (meson closed_imp_locally_compact locally_compact_Int open_imp_locally_compact) qed lemma lowerdim_embeddings: assumes "DIM('a) < DIM('b)" obtains f :: "'a::euclidean_space*real \ 'b::euclidean_space" and g :: "'b \ 'a*real" and j :: 'b where "linear f" "linear g" "\z. g (f z) = z" "j \ Basis" "\x. f(x,0) \ j = 0" proof - let ?B = "Basis :: ('a*real) set" have b01: "(0,1) \ ?B" by (simp add: Basis_prod_def) have "DIM('a * real) \ DIM('b)" by (simp add: Suc_leI assms) then obtain basf :: "'a*real \ 'b" where sbf: "basf ` ?B \ Basis" and injbf: "inj_on basf Basis" by (metis finite_Basis card_le_inj) define basg:: "'b \ 'a * real" where "basg \ \i. if i \ basf ` Basis then inv_into Basis basf i else (0,1)" have bgf[simp]: "basg (basf i) = i" if "i \ Basis" for i using inv_into_f_f injbf that by (force simp: basg_def) have sbg: "basg ` Basis \ ?B" by (force simp: basg_def injbf b01) define f :: "'a*real \ 'b" where "f \ \u. \j\Basis. (u \ basg j) *\<^sub>R j" define g :: "'b \ 'a*real" where "g \ \z. (\i\Basis. (z \ basf i) *\<^sub>R i)" show ?thesis proof show "linear f" unfolding f_def by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) show "linear g" unfolding g_def by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) have *: "(\a \ Basis. a \ basf b * (x \ basg a)) = x \ b" if "b \ Basis" for x b using sbf that by auto show gf: "g (f x) = x" for x proof (rule euclidean_eqI) show "\b. b \ Basis \ g (f x) \ b = x \ b" using f_def g_def sbf by auto qed show "basf(0,1) \ Basis" using b01 sbf by auto then show "f(x,0) \ basf(0,1) = 0" for x unfolding f_def inner_sum_left using b01 inner_not_same_Basis by (fastforce intro: comm_monoid_add_class.sum.neutral) qed qed proposition locally_compact_homeomorphic_closed: fixes S :: "'a::euclidean_space set" assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)" obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T" proof - obtain U:: "('a*real)set" and h where "closed U" and homU: "homeomorphism S U h fst" using locally_compact_homeomorphism_projection_closed assms by metis obtain f :: "'a*real \ 'b" and g :: "'b \ 'a*real" where "linear f" "linear g" and gf [simp]: "\z. g (f z) = z" using lowerdim_embeddings [OF dimlt] by metis then have "inj f" by (metis injI) have gfU: "g ` f ` U = U" by (simp add: image_comp o_def) have "S homeomorphic U" using homU homeomorphic_def by blast also have "... homeomorphic f ` U" proof (rule homeomorphicI [OF refl gfU]) show "continuous_on U f" by (meson \inj f\ \linear f\ homeomorphism_cont2 linear_homeomorphism_image) show "continuous_on (f ` U) g" using \linear g\ linear_continuous_on linear_conv_bounded_linear by blast qed (auto simp: o_def) finally show ?thesis using \closed U\ \inj f\ \linear f\ closed_injective_linear_image that by blast qed lemma homeomorphic_convex_compact_lemma: fixes S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "cball 0 1 \ S" shows "S homeomorphic (cball (0::'a) 1)" proof (rule starlike_compact_projective_special[OF assms(2-3)]) fix x u assume "x \ S" and "0 \ u" and "u < (1::real)" have "open (ball (u *\<^sub>R x) (1 - u))" by (rule open_ball) moreover have "u *\<^sub>R x \ ball (u *\<^sub>R x) (1 - u)" unfolding centre_in_ball using \u < 1\ by simp moreover have "ball (u *\<^sub>R x) (1 - u) \ S" proof fix y assume "y \ ball (u *\<^sub>R x) (1 - u)" then have "dist (u *\<^sub>R x) y < 1 - u" unfolding mem_ball . with \u < 1\ have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ cball 0 1" by (simp add: dist_norm inverse_eq_divide norm_minus_commute) with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ S" .. with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \ S" using \x \ S\ \0 \ u\ \u < 1\ [THEN less_imp_le] by (rule convexD_alt) then show "y \ S" using \u < 1\ by simp qed ultimately have "u *\<^sub>R x \ interior S" .. then show "u *\<^sub>R x \ S - frontier S" using frontier_def and interior_subset by auto qed proposition homeomorphic_convex_compact_cball: fixes e :: real and S :: "'a::euclidean_space set" assumes S: "convex S" "compact S" "interior S \ {}" and "e > 0" shows "S homeomorphic (cball (b::'a) e)" proof (rule homeomorphic_trans[OF _ homeomorphic_balls(2)]) obtain a where "a \ interior S" using assms by auto then show "S homeomorphic cball (0::'a) 1" by (metis (no_types) aff_dim_cball S compact_cball convex_cball homeomorphic_convex_lemma interior_rel_interior_gen zero_less_one) qed (use \e>0\ in auto) corollary homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "convex S" "compact S" "interior S \ {}" and "convex T" "compact T" "interior T \ {}" shows "S homeomorphic T" using assms by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) lemma homeomorphic_closed_intervals: fixes a :: "'a::euclidean_space" and b and c :: "'a::euclidean_space" and d assumes "box a b \ {}" and "box c d \ {}" shows "(cbox a b) homeomorphic (cbox c d)" by (simp add: assms homeomorphic_convex_compact) lemma homeomorphic_closed_intervals_real: fixes a::real and b and c::real and d assumes "aCovering spaces and lifting results for them\ definition\<^marker>\tag important\ covering_space :: "'a::topological_space set \ ('a \ 'b) \ 'b::topological_space set \ bool" where "covering_space c p S \ continuous_on c p \ p ` c = S \ (\x \ S. \T. x \ T \ openin (top_of_set S) T \ (\v. \v = c \ p -` T \ (\u \ v. openin (top_of_set c) u) \ pairwise disjnt v \ (\u \ v. \q. homeomorphism u T p q)))" lemma covering_space_imp_continuous: "covering_space c p S \ continuous_on c p" by (simp add: covering_space_def) lemma covering_space_imp_surjective: "covering_space c p S \ p ` c = S" by (simp add: covering_space_def) lemma homeomorphism_imp_covering_space: "homeomorphism S T f g \ covering_space S f T" apply (clarsimp simp add: homeomorphism_def covering_space_def) apply (rule_tac x=T in exI, simp) apply (rule_tac x="{S}" in exI, auto) done lemma covering_space_local_homeomorphism: assumes "covering_space c p S" "x \ c" obtains T u q where "x \ T" "openin (top_of_set c) T" "p x \ u" "openin (top_of_set S) u" "homeomorphism T u p q" using assms by (clarsimp simp add: covering_space_def) (metis IntI UnionE vimage_eq) lemma covering_space_local_homeomorphism_alt: assumes p: "covering_space c p S" and "y \ S" obtains x T U q where "p x = y" "x \ T" "openin (top_of_set c) T" "y \ U" "openin (top_of_set S) U" "homeomorphism T U p q" proof - obtain x where "p x = y" "x \ c" using assms covering_space_imp_surjective by blast show ?thesis using that \p x = y\ by (auto intro: covering_space_local_homeomorphism [OF p \x \ c\]) qed proposition covering_space_open_map: fixes S :: "'a :: metric_space set" and T :: "'b :: metric_space set" assumes p: "covering_space c p S" and T: "openin (top_of_set c) T" shows "openin (top_of_set S) (p ` T)" proof - have pce: "p ` c = S" and covs: "\x. x \ S \ \X VS. x \ X \ openin (top_of_set S) X \ \VS = c \ p -` X \ (\u \ VS. openin (top_of_set c) u) \ pairwise disjnt VS \ (\u \ VS. \q. homeomorphism u X p q)" using p by (auto simp: covering_space_def) have "T \ c" by (metis openin_euclidean_subtopology_iff T) have "\X. openin (top_of_set S) X \ y \ X \ X \ p ` T" if "y \ p ` T" for y proof - have "y \ S" using \T \ c\ pce that by blast obtain U VS where "y \ U" and U: "openin (top_of_set S) U" and VS: "\VS = c \ p -` U" and openVS: "\V \ VS. openin (top_of_set c) V" and homVS: "\V. V \ VS \ \q. homeomorphism V U p q" using covs [OF \y \ S\] by auto obtain x where "x \ c" "p x \ U" "x \ T" "p x = y" using T [unfolded openin_euclidean_subtopology_iff] \y \ U\ \y \ p ` T\ by blast with VS obtain V where "x \ V" "V \ VS" by auto then obtain q where q: "homeomorphism V U p q" using homVS by blast then have ptV: "p ` (T \ V) = U \ q -` (T \ V)" using VS \V \ VS\ by (auto simp: homeomorphism_def) have ocv: "openin (top_of_set c) V" by (simp add: \V \ VS\ openVS) have "openin (top_of_set (q ` U)) (T \ V)" using q unfolding homeomorphism_def by (metis T inf.absorb_iff2 ocv openin_imp_subset openin_subtopology_Int subtopology_subtopology) then have "openin (top_of_set U) (U \ q -` (T \ V))" using continuous_on_open homeomorphism_def q by blast then have os: "openin (top_of_set S) (U \ q -` (T \ V))" using openin_trans [of U] by (simp add: Collect_conj_eq U) show ?thesis proof (intro exI conjI) show "openin (top_of_set S) (p ` (T \ V))" by (simp only: ptV os) qed (use \p x = y\ \x \ V\ \x \ T\ in auto) qed with openin_subopen show ?thesis by blast qed lemma covering_space_lift_unique_gen: fixes f :: "'a::topological_space \ 'b::topological_space" fixes g1 :: "'a \ 'c::real_normed_vector" assumes cov: "covering_space c p S" and eq: "g1 a = g2 a" and f: "continuous_on T f" "f ` T \ S" and g1: "continuous_on T g1" "g1 ` T \ c" and fg1: "\x. x \ T \ f x = p(g1 x)" and g2: "continuous_on T g2" "g2 ` T \ c" and fg2: "\x. x \ T \ f x = p(g2 x)" and u_compt: "U \ components T" and "a \ U" "x \ U" shows "g1 x = g2 x" proof - have "U \ T" by (rule in_components_subset [OF u_compt]) define G12 where "G12 \ {x \ U. g1 x - g2 x = 0}" have "connected U" by (rule in_components_connected [OF u_compt]) have contu: "continuous_on U g1" "continuous_on U g2" using \U \ T\ continuous_on_subset g1 g2 by blast+ have o12: "openin (top_of_set U) G12" unfolding G12_def proof (subst openin_subopen, clarify) fix z assume z: "z \ U" "g1 z - g2 z = 0" obtain v w q where "g1 z \ v" and ocv: "openin (top_of_set c) v" and "p (g1 z) \ w" and osw: "openin (top_of_set S) w" and hom: "homeomorphism v w p q" proof (rule covering_space_local_homeomorphism [OF cov]) show "g1 z \ c" using \U \ T\ \z \ U\ g1(2) by blast qed auto have "g2 z \ v" using \g1 z \ v\ z by auto have gg: "U \ g -` v = U \ g -` (v \ g ` U)" for g by auto have "openin (top_of_set (g1 ` U)) (v \ g1 ` U)" using ocv \U \ T\ g1 by (fastforce simp add: openin_open) then have 1: "openin (top_of_set U) (U \ g1 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) have "openin (top_of_set (g2 ` U)) (v \ g2 ` U)" using ocv \U \ T\ g2 by (fastforce simp add: openin_open) then have 2: "openin (top_of_set U) (U \ g2 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) let ?T = "(U \ g1 -` v) \ (U \ g2 -` v)" show "\T. openin (top_of_set U) T \ z \ T \ T \ {z \ U. g1 z - g2 z = 0}" proof (intro exI conjI) show "openin (top_of_set U) ?T" using "1" "2" by blast show "z \ ?T" using z by (simp add: \g1 z \ v\ \g2 z \ v\) show "?T \ {z \ U. g1 z - g2 z = 0}" using hom by (clarsimp simp: homeomorphism_def) (metis \U \ T\ fg1 fg2 subsetD) qed qed have c12: "closedin (top_of_set U) G12" unfolding G12_def by (intro continuous_intros continuous_closedin_preimage_constant contu) have "G12 = {} \ G12 = U" by (intro connected_clopen [THEN iffD1, rule_format] \connected U\ conjI o12 c12) with eq \a \ U\ have "\x. x \ U \ g1 x - g2 x = 0" by (auto simp: G12_def) then show ?thesis using \x \ U\ by force qed proposition covering_space_lift_unique: fixes f :: "'a::topological_space \ 'b::topological_space" fixes g1 :: "'a \ 'c::real_normed_vector" assumes "covering_space c p S" "g1 a = g2 a" "continuous_on T f" "f ` T \ S" "continuous_on T g1" "g1 ` T \ c" "\x. x \ T \ f x = p(g1 x)" "continuous_on T g2" "g2 ` T \ c" "\x. x \ T \ f x = p(g2 x)" "connected T" "a \ T" "x \ T" shows "g1 x = g2 x" using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv by blast lemma covering_space_locally: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes loc: "locally \ C" and cov: "covering_space C p S" and pim: "\T. \T \ C; \ T\ \ \(p ` T)" shows "locally \ S" proof - have "locally \ (p ` C)" proof (rule locally_open_map_image [OF loc]) show "continuous_on C p" using cov covering_space_imp_continuous by blast show "\T. openin (top_of_set C) T \ openin (top_of_set (p ` C)) (p ` T)" using cov covering_space_imp_surjective covering_space_open_map by blast qed (simp add: pim) then show ?thesis using covering_space_imp_surjective [OF cov] by metis qed proposition covering_space_locally_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and pim: "\T. \T \ C; \ T\ \ \(p ` T)" and qim: "\q U. \U \ S; continuous_on U q; \ U\ \ \(q ` U)" shows "locally \ S \ locally \ C" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (rule locallyI) fix V x assume V: "openin (top_of_set C) V" and "x \ V" have "p x \ p ` C" by (metis IntE V \x \ V\ imageI openin_open) then obtain T \ where "p x \ T" and opeT: "openin (top_of_set S) T" and veq: "\\ = C \ p -` T" and ope: "\U\\. openin (top_of_set C) U" and hom: "\U\\. \q. homeomorphism U T p q" using cov unfolding covering_space_def by (blast intro: that) have "x \ \\" using V veq \p x \ T\ \x \ V\ openin_imp_subset by fastforce then obtain U where "x \ U" "U \ \" by blast then obtain q where opeU: "openin (top_of_set C) U" and q: "homeomorphism U T p q" using ope hom by blast with V have "openin (top_of_set C) (U \ V)" by blast then have UV: "openin (top_of_set S) (p ` (U \ V))" using cov covering_space_open_map by blast obtain W W' where opeW: "openin (top_of_set S) W" and "\ W'" "p x \ W" "W \ W'" and W'sub: "W' \ p ` (U \ V)" using locallyE [OF L UV] \x \ U\ \x \ V\ by blast then have "W \ T" by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans) show "\U Z. openin (top_of_set C) U \ \ Z \ x \ U \ U \ Z \ Z \ V" proof (intro exI conjI) have "openin (top_of_set T) W" by (meson opeW opeT openin_imp_subset openin_subset_trans \W \ T\) then have "openin (top_of_set U) (q ` W)" by (meson homeomorphism_imp_open_map homeomorphism_symD q) then show "openin (top_of_set C) (q ` W)" using opeU openin_trans by blast show "\ (q ` W')" by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \\ W'\ continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim) show "x \ q ` W" by (metis \p x \ W\ \x \ U\ homeomorphism_def imageI q) show "q ` W \ q ` W'" using \W \ W'\ by blast have "W' \ p ` V" using W'sub by blast then show "q ` W' \ V" using W'sub homeomorphism_apply1 [OF q] by auto qed qed next assume ?rhs then show ?lhs using cov covering_space_locally pim by blast qed lemma covering_space_locally_compact_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally compact S \ locally compact C" proof (rule covering_space_locally_eq [OF assms]) show "\T. \T \ C; compact T\ \ compact (p ` T)" by (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous) qed (use compact_continuous_image in blast) lemma covering_space_locally_connected_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally connected S \ locally connected C" proof (rule covering_space_locally_eq [OF assms]) show "\T. \T \ C; connected T\ \ connected (p ` T)" by (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) qed (use connected_continuous_image in blast) lemma covering_space_locally_path_connected_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally path_connected S \ locally path_connected C" proof (rule covering_space_locally_eq [OF assms]) show "\T. \T \ C; path_connected T\ \ path_connected (p ` T)" by (meson path_connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) qed (use path_connected_continuous_image in blast) lemma covering_space_locally_compact: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally compact C" "covering_space C p S" shows "locally compact S" using assms covering_space_locally_compact_eq by blast lemma covering_space_locally_connected: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally connected C" "covering_space C p S" shows "locally connected S" using assms covering_space_locally_connected_eq by blast lemma covering_space_locally_path_connected: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally path_connected C" "covering_space C p S" shows "locally path_connected S" using assms covering_space_locally_path_connected_eq by blast proposition covering_space_lift_homotopy: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and h :: "real \ 'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and conth: "continuous_on ({0..1} \ U) h" and him: "h ` ({0..1} \ U) \ S" and heq: "\y. y \ U \ h (0,y) = p(f y)" and contf: "continuous_on U f" and fim: "f ` U \ C" obtains k where "continuous_on ({0..1} \ U) k" "k ` ({0..1} \ U) \ C" "\y. y \ U \ k(0, y) = f y" "\z. z \ {0..1} \ U \ h z = p(k z)" proof - have "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..1} \ V) k \ k ` ({0..1} \ V) \ C \ (\z \ V. k(0, z) = f z) \ (\z \ {0..1} \ V. h z = p(k z))" if "y \ U" for y proof - obtain UU where UU: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s) \ (\\. \\ = C \ p -` UU s \ (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U (UU s) p q))" using cov unfolding covering_space_def by (metis (mono_tags)) then have ope: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s)" by blast have "\k n i. open k \ open n \ t \ k \ y \ n \ i \ S \ h ` (({0..1} \ k) \ (U \ n)) \ UU i" if "t \ {0..1}" for t proof - have hinS: "h (t, y) \ S" using \y \ U\ him that by blast then have "(t,y) \ ({0..1} \ U) \ h -` UU(h(t, y))" using \y \ U\ \t \ {0..1}\ by (auto simp: ope) moreover have ope_01U: "openin (top_of_set ({0..1} \ U)) (({0..1} \ U) \ h -` UU(h(t, y)))" using hinS ope continuous_on_open_gen [OF him] conth by blast ultimately obtain V W where opeV: "open V" and "t \ {0..1} \ V" "t \ {0..1} \ V" and opeW: "open W" and "y \ U" "y \ W" and VW: "({0..1} \ V) \ (U \ W) \ (({0..1} \ U) \ h -` UU(h(t, y)))" by (rule Times_in_interior_subtopology) (auto simp: openin_open) then show ?thesis using hinS by blast qed then obtain K NN X where K: "\t. t \ {0..1} \ open (K t)" and NN: "\t. t \ {0..1} \ open (NN t)" and inUS: "\t. t \ {0..1} \ t \ K t \ y \ NN t \ X t \ S" and him: "\t. t \ {0..1} \ h ` (({0..1} \ K t) \ (U \ NN t)) \ UU (X t)" by (metis (mono_tags)) obtain \ where "\ \ ((\i. K i \ NN i)) ` {0..1}" "finite \" "{0::real..1} \ {y} \ \\" proof (rule compactE) show "compact ({0::real..1} \ {y})" by (simp add: compact_Times) show "{0..1} \ {y} \ (\i\{0..1}. K i \ NN i)" using K inUS by auto show "\B. B \ (\i. K i \ NN i) ` {0..1} \ open B" using K NN by (auto simp: open_Times) qed blast then obtain tk where "tk \ {0..1}" "finite tk" and tk: "{0::real..1} \ {y} \ (\i \ tk. K i \ NN i)" by (metis (no_types, lifting) finite_subset_image) then have "tk \ {}" by auto define n where "n = \(NN ` tk)" have "y \ n" "open n" using inUS NN \tk \ {0..1}\ \finite tk\ by (auto simp: n_def open_INT subset_iff) obtain \ where "0 < \" and \: "\T. \T \ {0..1}; diameter T < \\ \ \B\K ` tk. T \ B" proof (rule Lebesgue_number_lemma [of "{0..1}" "K ` tk"]) show "K ` tk \ {}" using \tk \ {}\ by auto show "{0..1} \ \(K ` tk)" using tk by auto show "\B. B \ K ` tk \ open B" using \tk \ {0..1}\ K by auto qed auto obtain N::nat where N: "N > 1 / \" using reals_Archimedean2 by blast then have "N > 0" using \0 < \\ order.asym by force have *: "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..of_nat n / N} \ V) k \ k ` ({0..of_nat n / N} \ V) \ C \ (\z\V. k (0, z) = f z) \ (\z\{0..of_nat n / N} \ V. h z = p (k z))" if "n \ N" for n using that proof (induction n) case 0 show ?case apply (rule_tac x=U in exI) apply (rule_tac x="f \ snd" in exI) apply (intro conjI \y \ U\ continuous_intros continuous_on_subset [OF contf]) using fim apply (auto simp: heq) done next case (Suc n) then obtain V k where opeUV: "openin (top_of_set U) V" and "y \ V" and contk: "continuous_on ({0..n/N} \ V) k" and kim: "k ` ({0..n/N} \ V) \ C" and keq: "\z. z \ V \ k (0, z) = f z" and heq: "\z. z \ {0..n/N} \ V \ h z = p (k z)" using Suc_leD by auto have "n \ N" using Suc.prems by auto obtain t where "t \ tk" and t: "{n/N .. (1 + real n) / N} \ K t" proof (rule bexE [OF \]) show "{n/N .. (1 + real n) / N} \ {0..1}" using Suc.prems by (auto simp: field_split_simps) show diameter_less: "diameter {n/N .. (1 + real n) / N} < \" using \0 < \\ N by (auto simp: field_split_simps) qed blast have t01: "t \ {0..1}" using \t \ tk\ \tk \ {0..1}\ by blast obtain \ where \: "\\ = C \ p -` UU (X t)" and opeC: "\U. U \ \ \ openin (top_of_set C) U" and "pairwise disjnt \" and homuu: "\U. U \ \ \ \q. homeomorphism U (UU (X t)) p q" using inUS [OF t01] UU by meson have n_div_N_in: "n/N \ {n/N .. (1 + real n) / N}" using N by (auto simp: field_split_simps) with t have nN_in_kkt: "n/N \ K t" by blast have "k (n/N, y) \ C \ p -` UU (X t)" proof (simp, rule conjI) show "k (n/N, y) \ C" using \y \ V\ kim keq by force have "p (k (n/N, y)) = h (n/N, y)" by (simp add: \y \ V\ heq) also have "... \ h ` (({0..1} \ K t) \ (U \ NN t))" using \y \ V\ t01 \n \ N\ by (simp add: nN_in_kkt \y \ U\ inUS field_split_simps) also have "... \ UU (X t)" using him t01 by blast finally show "p (k (n/N, y)) \ UU (X t)" . qed with \ have "k (n/N, y) \ \\" by blast then obtain W where W: "k (n/N, y) \ W" and "W \ \" by blast then obtain p' where opeC': "openin (top_of_set C) W" and hom': "homeomorphism W (UU (X t)) p p'" using homuu opeC by blast then have "W \ C" using openin_imp_subset by blast define W' where "W' = UU(X t)" have opeVW: "openin (top_of_set V) (V \ (k \ Pair (n / N)) -` W)" proof (rule continuous_openin_preimage [OF _ _ opeC']) show "continuous_on V (k \ Pair (n/N))" by (intro continuous_intros continuous_on_subset [OF contk], auto) show "(k \ Pair (n/N)) ` V \ C" using kim by (auto simp: \y \ V\ W) qed obtain N' where opeUN': "openin (top_of_set U) N'" and "y \ N'" and kimw: "k ` ({(n/N)} \ N') \ W" proof show "openin (top_of_set U) (V \ (k \ Pair (n/N)) -` W)" using opeUV opeVW openin_trans by blast qed (use \y \ V\ W in \force+\) obtain Q Q' where opeUQ: "openin (top_of_set U) Q" and cloUQ': "closedin (top_of_set U) Q'" and "y \ Q" "Q \ Q'" and Q': "Q' \ (U \ NN(t)) \ N' \ V" proof - obtain VO VX where "open VO" "open VX" and VO: "V = U \ VO" and VX: "N' = U \ VX" using opeUV opeUN' by (auto simp: openin_open) then have "open (NN(t) \ VO \ VX)" using NN t01 by blast then obtain e where "e > 0" and e: "cball y e \ NN(t) \ VO \ VX" by (metis Int_iff \N' = U \ VX\ \V = U \ VO\ \y \ N'\ \y \ V\ inUS open_contains_cball t01) show ?thesis proof show "openin (top_of_set U) (U \ ball y e)" by blast show "closedin (top_of_set U) (U \ cball y e)" using e by (auto simp: closedin_closed) qed (use \y \ U\ \e > 0\ VO VX e in auto) qed then have "y \ Q'" "Q \ (U \ NN(t)) \ N' \ V" by blast+ have neq: "{0..n/N} \ {n/N..(1 + real n) / N} = {0..(1 + real n) / N}" apply (auto simp: field_split_simps) - by (metis mult_zero_left of_nat_0_le_iff of_nat_0_less_iff order_trans real_mult_le_cancel_iff1) + by (metis not_less of_nat_0_le_iff of_nat_0_less_iff order_trans zero_le_mult_iff) then have neqQ': "{0..n/N} \ Q' \ {n/N..(1 + real n) / N} \ Q' = {0..(1 + real n) / N} \ Q'" by blast have cont: "continuous_on ({0..(1 + real n) / N} \ Q') (\x. if x \ {0..n/N} \ Q' then k x else (p' \ h) x)" unfolding neqQ' [symmetric] proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply) have "\T. closed T \ {0..n/N} \ Q' = {0..(1+n)/N} \ Q' \ T" using n_div_N_in by (rule_tac x="{0 .. n/N} \ UNIV" in exI) (auto simp: closed_Times) then show "closedin (top_of_set ({0..(1 + real n) / N} \ Q')) ({0..n/N} \ Q')" by (simp add: closedin_closed) have "\T. closed T \ {n/N..(1+n)/N} \ Q' = {0..(1+n)/N} \ Q' \ T" by (rule_tac x="{n/N..(1+n)/N} \ UNIV" in exI) (auto simp: closed_Times order_trans [rotated]) then show "closedin (top_of_set ({0..(1 + real n) / N} \ Q')) ({n/N..(1 + real n) / N} \ Q')" by (simp add: closedin_closed) show "continuous_on ({0..n/N} \ Q') k" using Q' by (auto intro: continuous_on_subset [OF contk]) have "continuous_on ({n/N..(1 + real n) / N} \ Q') h" proof (rule continuous_on_subset [OF conth]) show "{n/N..(1 + real n) / N} \ Q' \ {0..1} \ U" proof (clarsimp, intro conjI) fix a b assume "b \ Q'" and a: "n/N \ a" "a \ (1 + real n) / N" have "0 \ n/N" "(1 + real n) / N \ 1" using a Suc.prems by (auto simp: divide_simps) with a show "0 \ a" "a \ 1" by linarith+ show "b \ U" using \b \ Q'\ cloUQ' closedin_imp_subset by blast qed qed moreover have "continuous_on (h ` ({n/N..(1 + real n) / N} \ Q')) p'" proof (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom']]) have "h ` ({n/N..(1 + real n) / N} \ Q') \ h ` (({0..1} \ K t) \ (U \ NN t))" proof (rule image_mono) show "{n/N..(1 + real n) / N} \ Q' \ ({0..1} \ K t) \ (U \ NN t)" proof (clarsimp, intro conjI) fix a::real and b assume "b \ Q'" "n/N \ a" "a \ (1 + real n) / N" show "0 \ a" by (meson \n/N \ a\ divide_nonneg_nonneg of_nat_0_le_iff order_trans) show "a \ 1" using Suc.prems \a \ (1 + real n) / N\ order_trans by force show "a \ K t" using \a \ (1 + real n) / N\ \n/N \ a\ t by auto show "b \ U" using \b \ Q'\ cloUQ' closedin_imp_subset by blast show "b \ NN t" using Q' \b \ Q'\ by auto qed qed with him show "h ` ({n/N..(1 + real n) / N} \ Q') \ UU (X t)" using t01 by blast qed ultimately show "continuous_on ({n/N..(1 + real n) / N} \ Q') (p' \ h)" by (rule continuous_on_compose) have "k (n/N, b) = p' (h (n/N, b))" if "b \ Q'" for b proof - have "k (n/N, b) \ W" using that Q' kimw by force then have "k (n/N, b) = p' (p (k (n/N, b)))" by (simp add: homeomorphism_apply1 [OF hom']) then show ?thesis using Q' that by (force simp: heq) qed then show "\x. x \ {n/N..(1 + real n) / N} \ Q' \ x \ {0..n/N} \ Q' \ k x = (p' \ h) x" by auto qed have h_in_UU: "h (x, y) \ UU (X t)" if "y \ Q" "\ x \ n/N" "0 \ x" "x \ (1 + real n) / N" for x y proof - have "x \ 1" using Suc.prems that order_trans by force moreover have "x \ K t" by (meson atLeastAtMost_iff le_less not_le subset_eq t that) moreover have "y \ U" using \y \ Q\ opeUQ openin_imp_subset by blast moreover have "y \ NN t" using Q' \Q \ Q'\ \y \ Q\ by auto ultimately have "(x, y) \ (({0..1} \ K t) \ (U \ NN t))" using that by auto then have "h (x, y) \ h ` (({0..1} \ K t) \ (U \ NN t))" by blast also have "... \ UU (X t)" by (metis him t01) finally show ?thesis . qed let ?k = "(\x. if x \ {0..n/N} \ Q' then k x else (p' \ h) x)" show ?case proof (intro exI conjI) show "continuous_on ({0..real (Suc n) / N} \ Q) ?k" using \Q \ Q'\ by (auto intro: continuous_on_subset [OF cont]) have "\x y. \x \ n/N; y \ Q'; 0 \ x\ \ k (x, y) \ C" using kim Q' by force moreover have "p' (h (x, y)) \ C" if "y \ Q" "\ x \ n/N" "0 \ x" "x \ (1 + real n) / N" for x y proof (rule \W \ C\ [THEN subsetD]) show "p' (h (x, y)) \ W" using homeomorphism_image2 [OF hom', symmetric] h_in_UU Q' \Q \ Q'\ \W \ C\ that by auto qed ultimately show "?k ` ({0..real (Suc n) / N} \ Q) \ C" using Q' \Q \ Q'\ by force show "\z\Q. ?k (0, z) = f z" using Q' keq \Q \ Q'\ by auto show "\z \ {0..real (Suc n) / N} \ Q. h z = p(?k z)" using \Q \ U \ NN t \ N' \ V\ heq Q' \Q \ Q'\ by (auto simp: homeomorphism_apply2 [OF hom'] dest: h_in_UU) qed (auto simp: \y \ Q\ opeUQ) qed show ?thesis using *[OF order_refl] N \0 < \\ by (simp add: split: if_split_asm) qed then obtain V fs where opeV: "\y. y \ U \ openin (top_of_set U) (V y)" and V: "\y. y \ U \ y \ V y" and contfs: "\y. y \ U \ continuous_on ({0..1} \ V y) (fs y)" and *: "\y. y \ U \ (fs y) ` ({0..1} \ V y) \ C \ (\z \ V y. fs y (0, z) = f z) \ (\z \ {0..1} \ V y. h z = p(fs y z))" by (metis (mono_tags)) then have VU: "\y. y \ U \ V y \ U" by (meson openin_imp_subset) obtain k where contk: "continuous_on ({0..1} \ U) k" and k: "\x i. \i \ U; x \ {0..1} \ U \ {0..1} \ V i\ \ k x = fs i x" proof (rule pasting_lemma_exists) let ?X = "top_of_set ({0..1::real} \ U)" show "topspace ?X \ (\i\U. {0..1} \ V i)" using V by force show "\i. i \ U \ openin (top_of_set ({0..1} \ U)) ({0..1} \ V i)" by (simp add: Abstract_Topology.openin_Times opeV) show "\i. i \ U \ continuous_map (subtopology (top_of_set ({0..1} \ U)) ({0..1} \ V i)) euclidean (fs i)" by (metis contfs subtopology_subtopology continuous_map_iff_continuous Times_Int_Times VU inf.absorb_iff2 inf.idem) show "fs i x = fs j x" if "i \ U" "j \ U" and x: "x \ topspace ?X \ {0..1} \ V i \ {0..1} \ V j" for i j x proof - obtain u y where "x = (u, y)" "y \ V i" "y \ V j" "0 \ u" "u \ 1" using x by auto show ?thesis proof (rule covering_space_lift_unique [OF cov, of _ "(0,y)" _ "{0..1} \ {y}" h]) show "fs i (0, y) = fs j (0, y)" using*V by (simp add: \y \ V i\ \y \ V j\ that) show conth_y: "continuous_on ({0..1} \ {y}) h" using VU \y \ V j\ that by (auto intro: continuous_on_subset [OF conth]) show "h ` ({0..1} \ {y}) \ S" using \y \ V i\ assms(3) VU that by fastforce show "continuous_on ({0..1} \ {y}) (fs i)" using continuous_on_subset [OF contfs] \i \ U\ by (simp add: \y \ V i\ subset_iff) show "fs i ` ({0..1} \ {y}) \ C" using "*" \y \ V i\ \i \ U\ by fastforce show "\x. x \ {0..1} \ {y} \ h x = p (fs i x)" using "*" \y \ V i\ \i \ U\ by blast show "continuous_on ({0..1} \ {y}) (fs j)" using continuous_on_subset [OF contfs] \j \ U\ by (simp add: \y \ V j\ subset_iff) show "fs j ` ({0..1} \ {y}) \ C" using "*" \y \ V j\ \j \ U\ by fastforce show "\x. x \ {0..1} \ {y} \ h x = p (fs j x)" using "*" \y \ V j\ \j \ U\ by blast show "connected ({0..1::real} \ {y})" using connected_Icc connected_Times connected_sing by blast show "(0, y) \ {0..1::real} \ {y}" by force show "x \ {0..1} \ {y}" using \x = (u, y)\ x by blast qed qed qed force show ?thesis proof show "k ` ({0..1} \ U) \ C" using V*k VU by fastforce show "\y. y \ U \ k (0, y) = f y" by (simp add: V*k) show "\z. z \ {0..1} \ U \ h z = p (k z)" using V*k by auto qed (auto simp: contk) qed corollary covering_space_lift_homotopy_alt: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and h :: "'c::real_normed_vector \ real \ 'b" assumes cov: "covering_space C p S" and conth: "continuous_on (U \ {0..1}) h" and him: "h ` (U \ {0..1}) \ S" and heq: "\y. y \ U \ h (y,0) = p(f y)" and contf: "continuous_on U f" and fim: "f ` U \ C" obtains k where "continuous_on (U \ {0..1}) k" "k ` (U \ {0..1}) \ C" "\y. y \ U \ k(y, 0) = f y" "\z. z \ U \ {0..1} \ h z = p(k z)" proof - have "continuous_on ({0..1} \ U) (h \ (\z. (snd z, fst z)))" by (intro continuous_intros continuous_on_subset [OF conth]) auto then obtain k where contk: "continuous_on ({0..1} \ U) k" and kim: "k ` ({0..1} \ U) \ C" and k0: "\y. y \ U \ k(0, y) = f y" and heqp: "\z. z \ {0..1} \ U \ (h \ (\z. Pair (snd z) (fst z))) z = p(k z)" apply (rule covering_space_lift_homotopy [OF cov _ _ _ contf fim]) using him by (auto simp: contf heq) show ?thesis proof show "continuous_on (U \ {0..1}) (k \ (\z. (snd z, fst z)))" by (intro continuous_intros continuous_on_subset [OF contk]) auto qed (use kim heqp in \auto simp: k0\) qed corollary covering_space_lift_homotopic_function: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and g:: "'c::real_normed_vector \ 'a" assumes cov: "covering_space C p S" and contg: "continuous_on U g" and gim: "g ` U \ C" and pgeq: "\y. y \ U \ p(g y) = f y" and hom: "homotopic_with_canon (\x. True) U S f f'" obtains g' where "continuous_on U g'" "image g' U \ C" "\y. y \ U \ p(g' y) = f' y" proof - obtain h where conth: "continuous_on ({0..1::real} \ U) h" and him: "h ` ({0..1} \ U) \ S" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = f' x" using hom by (auto simp: homotopic_with_def) have "\y. y \ U \ h (0, y) = p (g y)" by (simp add: h0 pgeq) then obtain k where contk: "continuous_on ({0..1} \ U) k" and kim: "k ` ({0..1} \ U) \ C" and k0: "\y. y \ U \ k(0, y) = g y" and heq: "\z. z \ {0..1} \ U \ h z = p(k z)" using covering_space_lift_homotopy [OF cov conth him _ contg gim] by metis show ?thesis proof show "continuous_on U (k \ Pair 1)" by (meson contk atLeastAtMost_iff continuous_on_o_Pair order_refl zero_le_one) show "(k \ Pair 1) ` U \ C" using kim by auto show "\y. y \ U \ p ((k \ Pair 1) y) = f' y" by (auto simp: h1 heq [symmetric]) qed qed corollary covering_space_lift_inessential_function: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and U :: "'c::real_normed_vector set" assumes cov: "covering_space C p S" and hom: "homotopic_with_canon (\x. True) U S f (\x. a)" obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" proof (cases "U = {}") case True then show ?thesis using that continuous_on_empty by blast next case False then obtain b where b: "b \ C" "p b = a" using covering_space_imp_surjective [OF cov] homotopic_with_imp_subset2 [OF hom] by auto then have gim: "(\y. b) ` U \ C" by blast show ?thesis proof (rule covering_space_lift_homotopic_function [OF cov continuous_on_const gim]) show "\y. y \ U \ p b = a" using b by auto qed (use that homotopic_with_symD [OF hom] in auto) qed subsection\ Lifting of general functions to covering space\ proposition covering_space_lift_path_strong: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and "a \ C" and "path g" and pag: "path_image g \ S" and pas: "pathstart g = p a" obtains h where "path h" "path_image h \ C" "pathstart h = a" and "\t. t \ {0..1} \ p(h t) = g t" proof - obtain k:: "real \ 'c \ 'a" where contk: "continuous_on ({0..1} \ {undefined}) k" and kim: "k ` ({0..1} \ {undefined}) \ C" and k0: "k (0, undefined) = a" and pk: "\z. z \ {0..1} \ {undefined} \ p(k z) = (g \ fst) z" proof (rule covering_space_lift_homotopy [OF cov, of "{undefined}" "g \ fst"]) show "continuous_on ({0..1::real} \ {undefined::'c}) (g \ fst)" using \path g\ by (intro continuous_intros) (simp add: path_def) show "(g \ fst) ` ({0..1} \ {undefined}) \ S" using pag by (auto simp: path_image_def) show "(g \ fst) (0, y) = p a" if "y \ {undefined}" for y::'c by (metis comp_def fst_conv pas pathstart_def) qed (use assms in auto) show ?thesis proof show "path (k \ (\t. Pair t undefined))" unfolding path_def by (intro continuous_on_compose continuous_intros continuous_on_subset [OF contk]) auto show "path_image (k \ (\t. (t, undefined))) \ C" using kim by (auto simp: path_image_def) show "pathstart (k \ (\t. (t, undefined))) = a" by (auto simp: pathstart_def k0) show "\t. t \ {0..1} \ p ((k \ (\t. (t, undefined))) t) = g t" by (auto simp: pk) qed qed corollary covering_space_lift_path: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \ S" obtains h where "path h" "path_image h \ C" "\t. t \ {0..1} \ p(h t) = g t" proof - obtain a where "a \ C" "pathstart g = p a" by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE) show ?thesis using covering_space_lift_path_strong [OF cov \a \ C\ \path g\ pig] by (metis \pathstart g = p a\ that) qed proposition covering_space_lift_homotopic_paths: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g1" and pig1: "path_image g1 \ S" and "path g2" and pig2: "path_image g2 \ S" and hom: "homotopic_paths S g1 g2" and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" and h1h2: "pathstart h1 = pathstart h2" shows "homotopic_paths C h1 h2" proof - obtain h :: "real \ real \ 'b" where conth: "continuous_on ({0..1} \ {0..1}) h" and him: "h ` ({0..1} \ {0..1}) \ S" and h0: "\x. h (0, x) = g1 x" and h1: "\x. h (1, x) = g2 x" and heq0: "\t. t \ {0..1} \ h (t, 0) = g1 0" and heq1: "\t. t \ {0..1} \ h (t, 1) = g1 1" using hom by (auto simp: homotopic_paths_def homotopic_with_def pathstart_def pathfinish_def) obtain k where contk: "continuous_on ({0..1} \ {0..1}) k" and kim: "k ` ({0..1} \ {0..1}) \ C" and kh2: "\y. y \ {0..1} \ k (y, 0) = h2 0" and hpk: "\z. z \ {0..1} \ {0..1} \ h z = p (k z)" proof (rule covering_space_lift_homotopy_alt [OF cov conth him]) show "\y. y \ {0..1} \ h (y, 0) = p (h2 0)" by (metis atLeastAtMost_iff h1h2 heq0 order_refl pathstart_def ph1 zero_le_one) qed (use path_image_def pih2 in \fastforce+\) have contg1: "continuous_on {0..1} g1" and contg2: "continuous_on {0..1} g2" using \path g1\ \path g2\ path_def by blast+ have g1im: "g1 ` {0..1} \ S" and g2im: "g2 ` {0..1} \ S" using path_image_def pig1 pig2 by auto have conth1: "continuous_on {0..1} h1" and conth2: "continuous_on {0..1} h2" using \path h1\ \path h2\ path_def by blast+ have h1im: "h1 ` {0..1} \ C" and h2im: "h2 ` {0..1} \ C" using path_image_def pih1 pih2 by auto show ?thesis unfolding homotopic_paths pathstart_def pathfinish_def proof (intro exI conjI ballI) show keqh1: "k(0, x) = h1 x" if "x \ {0..1}" for x proof (rule covering_space_lift_unique [OF cov _ contg1 g1im]) show "k (0,0) = h1 0" by (metis atLeastAtMost_iff h1h2 kh2 order_refl pathstart_def zero_le_one) show "continuous_on {0..1} (\a. k (0, a))" by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g1 x = p (k (0, x))" by (metis atLeastAtMost_iff h0 hpk zero_le_one mem_Sigma_iff order_refl) qed (use conth1 h1im kim that in \auto simp: ph1\) show "k(1, x) = h2 x" if "x \ {0..1}" for x proof (rule covering_space_lift_unique [OF cov _ contg2 g2im]) show "k (1,0) = h2 0" by (metis atLeastAtMost_iff kh2 order_refl zero_le_one) show "continuous_on {0..1} (\a. k (1, a))" by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g2 x = p (k (1, x))" by (metis atLeastAtMost_iff h1 hpk mem_Sigma_iff order_refl zero_le_one) qed (use conth2 h2im kim that in \auto simp: ph2\) show "\t. t \ {0..1} \ (k \ Pair t) 0 = h1 0" by (metis comp_apply h1h2 kh2 pathstart_def) show "(k \ Pair t) 1 = h1 1" if "t \ {0..1}" for t proof (rule covering_space_lift_unique [OF cov, of "\a. (k \ Pair a) 1" 0 "\a. h1 1" "{0..1}" "\x. g1 1"]) show "(k \ Pair 0) 1 = h1 1" using keqh1 by auto show "continuous_on {0..1} (\a. (k \ Pair a) 1)" by (auto intro!: continuous_intros continuous_on_compose2 [OF contk]) show "\x. x \ {0..1} \ g1 1 = p ((k \ Pair x) 1)" using heq1 hpk by auto qed (use contk kim g1im h1im that in \auto simp: ph1\) qed (use contk kim in auto) qed corollary covering_space_monodromy: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g1" and pig1: "path_image g1 \ S" and "path g2" and pig2: "path_image g2 \ S" and hom: "homotopic_paths S g1 g2" and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" and h1h2: "pathstart h1 = pathstart h2" shows "pathfinish h1 = pathfinish h2" using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish by blast corollary covering_space_lift_homotopic_path: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and hom: "homotopic_paths S f f'" and "path g" and pig: "path_image g \ C" and a: "pathstart g = a" and b: "pathfinish g = b" and pgeq: "\t. t \ {0..1} \ p(g t) = f t" obtains g' where "path g'" "path_image g' \ C" "pathstart g' = a" "pathfinish g' = b" "\t. t \ {0..1} \ p(g' t) = f' t" proof (rule covering_space_lift_path_strong [OF cov, of a f']) show "a \ C" using a pig by auto show "path f'" "path_image f' \ S" using hom homotopic_paths_imp_path homotopic_paths_imp_subset by blast+ show "pathstart f' = p a" by (metis a atLeastAtMost_iff hom homotopic_paths_imp_pathstart order_refl pathstart_def pgeq zero_le_one) qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig) proposition covering_space_lift_general: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and "a \ C" "z \ U" and U: "path_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ \ \q. path q \ path_image q \ C \ pathstart q = a \ pathfinish q = a \ homotopic_paths S (f \ r) (p \ q)" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof - have *: "\g h. path g \ path_image g \ U \ pathstart g = z \ pathfinish g = y \ path h \ path_image h \ C \ pathstart h = a \ (\t \ {0..1}. p(h t) = f(g t))" if "y \ U" for y proof - obtain g where "path g" "path_image g \ U" and pastg: "pathstart g = z" and pafig: "pathfinish g = y" using U \z \ U\ \y \ U\ by (force simp: path_connected_def) obtain h where "path h" "path_image h \ C" "pathstart h = a" and "\t. t \ {0..1} \ p(h t) = (f \ g) t" proof (rule covering_space_lift_path_strong [OF cov \a \ C\]) show "path (f \ g)" using \path g\ \path_image g \ U\ contf continuous_on_subset path_continuous_image by blast show "path_image (f \ g) \ S" by (metis \path_image g \ U\ fim image_mono path_image_compose subset_trans) show "pathstart (f \ g) = p a" by (simp add: feq pastg pathstart_compose) qed auto then show ?thesis by (metis \path g\ \path_image g \ U\ comp_apply pafig pastg) qed have "\l. \g h. path g \ path_image g \ U \ pathstart g = z \ pathfinish g = y \ path h \ path_image h \ C \ pathstart h = a \ (\t \ {0..1}. p(h t) = f(g t)) \ pathfinish h = l" for y proof - have "pathfinish h = pathfinish h'" if g: "path g" "path_image g \ U" "pathstart g = z" "pathfinish g = y" and h: "path h" "path_image h \ C" "pathstart h = a" and phg: "\t. t \ {0..1} \ p(h t) = f(g t)" and g': "path g'" "path_image g' \ U" "pathstart g' = z" "pathfinish g' = y" and h': "path h'" "path_image h' \ C" "pathstart h' = a" and phg': "\t. t \ {0..1} \ p(h' t) = f(g' t)" for g h g' h' proof - obtain q where "path q" and piq: "path_image q \ C" and pastq: "pathstart q = a" and pafiq: "pathfinish q = a" and homS: "homotopic_paths S (f \ g +++ reversepath g') (p \ q)" using g g' hom [of "g +++ reversepath g'"] by (auto simp: subset_path_image_join) have papq: "path (p \ q)" using homS homotopic_paths_imp_path by blast have pipq: "path_image (p \ q) \ S" using homS homotopic_paths_imp_subset by blast obtain q' where "path q'" "path_image q' \ C" and "pathstart q' = pathstart q" "pathfinish q' = pathfinish q" and pq'_eq: "\t. t \ {0..1} \ p (q' t) = (f \ g +++ reversepath g') t" using covering_space_lift_homotopic_path [OF cov homotopic_paths_sym [OF homS] \path q\ piq refl refl] by auto have "q' t = (h \ (*\<^sub>R) 2) t" if "0 \ t" "t \ 1/2" for t proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \ (*\<^sub>R) 2" "{0..1/2}" "f \ g \ (*\<^sub>R) 2" t]) show "q' 0 = (h \ (*\<^sub>R) 2) 0" by (metis \pathstart q' = pathstart q\ comp_def h(3) pastq pathstart_def pth_4(2)) show "continuous_on {0..1/2} (f \ g \ (*\<^sub>R) 2)" proof (intro continuous_intros continuous_on_path [OF \path g\] continuous_on_subset [OF contf]) show "g ` (*\<^sub>R) 2 ` {0..1/2} \ U" using g path_image_def by fastforce qed auto show "(f \ g \ (*\<^sub>R) 2) ` {0..1/2} \ S" using g(2) path_image_def fim by fastforce show "(h \ (*\<^sub>R) 2) ` {0..1/2} \ C" using h path_image_def by fastforce show "q' ` {0..1/2} \ C" using \path_image q' \ C\ path_image_def by fastforce show "\x. x \ {0..1/2} \ (f \ g \ (*\<^sub>R) 2) x = p (q' x)" by (auto simp: joinpaths_def pq'_eq) show "\x. x \ {0..1/2} \ (f \ g \ (*\<^sub>R) 2) x = p ((h \ (*\<^sub>R) 2) x)" by (simp add: phg) show "continuous_on {0..1/2} q'" by (simp add: continuous_on_path \path q'\) show "continuous_on {0..1/2} (h \ (*\<^sub>R) 2)" by (intro continuous_intros continuous_on_path [OF \path h\]) auto qed (use that in auto) moreover have "q' t = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) t" if "1/2 < t" "t \ 1" for t proof (rule covering_space_lift_unique [OF cov, of q' 1 "reversepath h' \ (\t. 2 *\<^sub>R t - 1)" "{1/2<..1}" "f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)" t]) show "q' 1 = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) 1" using h' \pathfinish q' = pathfinish q\ pafiq by (simp add: pathstart_def pathfinish_def reversepath_def) show "continuous_on {1/2<..1} (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1))" proof (intro continuous_intros continuous_on_path \path g'\ continuous_on_subset [OF contf]) show "reversepath g' ` (\t. 2 *\<^sub>R t - 1) ` {1/2<..1} \ U" using g' by (auto simp: path_image_def reversepath_def) qed (use g' in auto) show "(f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ S" using g'(2) path_image_def fim by (auto simp: image_subset_iff path_image_def reversepath_def) show "q' ` {1/2<..1} \ C" using \path_image q' \ C\ path_image_def by fastforce show "(reversepath h' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ C" using h' by (simp add: path_image_def reversepath_def subset_eq) show "\x. x \ {1/2<..1} \ (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p (q' x)" by (auto simp: joinpaths_def pq'_eq) show "\x. x \ {1/2<..1} \ (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p ((reversepath h' \ (\t. 2 *\<^sub>R t - 1)) x)" by (simp add: phg' reversepath_def) show "continuous_on {1/2<..1} q'" by (auto intro: continuous_on_path [OF \path q'\]) show "continuous_on {1/2<..1} (reversepath h' \ (\t. 2 *\<^sub>R t - 1))" by (intro continuous_intros continuous_on_path \path h'\) (use h' in auto) qed (use that in auto) ultimately have "q' t = (h +++ reversepath h') t" if "0 \ t" "t \ 1" for t using that by (simp add: joinpaths_def) then have "path(h +++ reversepath h')" by (auto intro: path_eq [OF \path q'\]) then show ?thesis by (auto simp: \path h\ \path h'\) qed then show ?thesis by metis qed then obtain l :: "'c \ 'a" where l: "\y g h. \path g; path_image g \ U; pathstart g = z; pathfinish g = y; path h; path_image h \ C; pathstart h = a; \t. t \ {0..1} \ p(h t) = f(g t)\ \ pathfinish h = l y" by metis show ?thesis proof show pleq: "p (l y) = f y" if "y \ U" for y using*[OF \y \ U\] by (metis l atLeastAtMost_iff order_refl pathfinish_def zero_le_one) show "l z = a" using l [of "linepath z z" z "linepath a a"] by (auto simp: assms) show LC: "l ` U \ C" by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE) have "\T. openin (top_of_set U) T \ y \ T \ T \ U \ l -` X" if X: "openin (top_of_set C) X" and "y \ U" "l y \ X" for X y proof - have "X \ C" using X openin_euclidean_subtopology_iff by blast have "f y \ S" using fim \y \ U\ by blast then obtain W \ where WV: "f y \ W \ openin (top_of_set S) W \ (\\ = C \ p -` W \ (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U W p q))" using cov by (force simp: covering_space_def) then have "l y \ \\" using \X \ C\ pleq that by auto then obtain W' where "l y \ W'" and "W' \ \" by blast with WV obtain p' where opeCW': "openin (top_of_set C) W'" and homUW': "homeomorphism W' W p p'" by blast then have contp': "continuous_on W p'" and p'im: "p' ` W \ W'" using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+ obtain V where "y \ V" "y \ U" and fimW: "f ` V \ W" "V \ U" and "path_connected V" and opeUV: "openin (top_of_set U) V" proof - have "openin (top_of_set U) (U \ f -` W)" using WV contf continuous_on_open_gen fim by auto then obtain UO where "openin (top_of_set U) UO \ path_connected UO \ y \ UO \ UO \ U \ f -` W" using U WV \y \ U\ unfolding locally_path_connected by (meson IntI vimage_eq) then show ?thesis by (meson \y \ U\ image_subset_iff_subset_vimage le_inf_iff that) qed have "W' \ C" "W \ S" using opeCW' WV openin_imp_subset by auto have p'im: "p' ` W \ W'" using homUW' homeomorphism_image2 by fastforce show ?thesis proof (intro exI conjI) have "openin (top_of_set S) (W \ p' -` (W' \ X))" proof (rule openin_trans) show "openin (top_of_set W) (W \ p' -` (W' \ X))" using X \W' \ C\ by (intro continuous_openin_preimage [OF contp' p'im]) (auto simp: openin_open) show "openin (top_of_set S) W" using WV by blast qed then show "openin (top_of_set U) (V \ (U \ (f -` (W \ (p' -` (W' \ X))))))" by (blast intro: opeUV openin_subtopology_self continuous_openin_preimage [OF contf fim]) have "p' (f y) \ X" using \l y \ W'\ homeomorphism_apply1 [OF homUW'] pleq \y \ U\ \l y \ X\ by fastforce then show "y \ V \ (U \ f -` (W \ p' -` (W' \ X)))" using \y \ U\ \y \ V\ WV p'im by auto show "V \ (U \ f -` (W \ p' -` (W' \ X))) \ U \ l -` X" proof (intro subsetI IntI; clarify) fix y' assume y': "y' \ V" "y' \ U" "f y' \ W" "p' (f y') \ W'" "p' (f y') \ X" then obtain \ where "path \" "path_image \ \ V" "pathstart \ = y" "pathfinish \ = y'" by (meson \path_connected V\ \y \ V\ path_connected_def) obtain pp qq where pp: "path pp" "path_image pp \ U" "pathstart pp = z" "pathfinish pp = y" and qq: "path qq" "path_image qq \ C" "pathstart qq = a" and pqqeq: "\t. t \ {0..1} \ p(qq t) = f(pp t)" using*[OF \y \ U\] by blast have finW: "\x. \0 \ x; x \ 1\ \ f (\ x) \ W" using \path_image \ \ V\ by (auto simp: image_subset_iff path_image_def fimW [THEN subsetD]) have "pathfinish (qq +++ (p' \ f \ \)) = l y'" proof (rule l [of "pp +++ \" y' "qq +++ (p' \ f \ \)"]) show "path (pp +++ \)" by (simp add: \path \\ \path pp\ \pathfinish pp = y\ \pathstart \ = y\) show "path_image (pp +++ \) \ U" using \V \ U\ \path_image \ \ V\ \path_image pp \ U\ not_in_path_image_join by blast show "pathstart (pp +++ \) = z" by (simp add: \pathstart pp = z\) show "pathfinish (pp +++ \) = y'" by (simp add: \pathfinish \ = y'\) have "pathfinish qq = l y" using \path pp\ \path qq\ \path_image pp \ U\ \path_image qq \ C\ \pathfinish pp = y\ \pathstart pp = z\ \pathstart qq = a\ l pqqeq by blast also have "... = p' (f y)" using \l y \ W'\ homUW' homeomorphism_apply1 pleq that(2) by fastforce finally have "pathfinish qq = p' (f y)" . then have paqq: "pathfinish qq = pathstart (p' \ f \ \)" by (simp add: \pathstart \ = y\ pathstart_compose) have "continuous_on (path_image \) (p' \ f)" proof (rule continuous_on_compose) show "continuous_on (path_image \) f" using \path_image \ \ V\ \V \ U\ contf continuous_on_subset by blast show "continuous_on (f ` path_image \) p'" proof (rule continuous_on_subset [OF contp']) show "f ` path_image \ \ W" by (auto simp: path_image_def pathfinish_def pathstart_def finW) qed qed then show "path (qq +++ (p' \ f \ \))" using \path \\ \path qq\ paqq path_continuous_image path_join_imp by blast show "path_image (qq +++ (p' \ f \ \)) \ C" proof (rule subset_path_image_join) show "path_image qq \ C" by (simp add: \path_image qq \ C\) show "path_image (p' \ f \ \) \ C" by (metis \W' \ C\ \path_image \ \ V\ dual_order.trans fimW(1) image_comp image_mono p'im path_image_compose) qed show "pathstart (qq +++ (p' \ f \ \)) = a" by (simp add: \pathstart qq = a\) show "p ((qq +++ (p' \ f \ \)) \) = f ((pp +++ \) \)" if \: "\ \ {0..1}" for \ proof (simp add: joinpaths_def, safe) show "p (qq (2*\)) = f (pp (2*\))" if "\*2 \ 1" using \\ \ {0..1}\ pqqeq that by auto show "p (p' (f (\ (2*\ - 1)))) = f (\ (2*\ - 1))" if "\ \*2 \ 1" using that \ by (auto intro: homeomorphism_apply2 [OF homUW' finW]) qed qed with \pathfinish \ = y'\ \p' (f y') \ X\ show "y' \ l -` X" unfolding pathfinish_join by (simp add: pathfinish_def) qed qed qed then show "continuous_on U l" by (metis IntD1 IntD2 vimage_eq openin_subopen continuous_on_open_gen [OF LC]) qed qed corollary covering_space_lift_stronger: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" "a \ C" "z \ U" and U: "path_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ \ \b. homotopic_paths S (f \ r) (linepath b b)" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof (rule covering_space_lift_general [OF cov U contf fim feq]) fix r assume "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" then obtain b where b: "homotopic_paths S (f \ r) (linepath b b)" using hom by blast then have "f (pathstart r) = b" by (metis homotopic_paths_imp_pathstart pathstart_compose pathstart_linepath) then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" by (simp add: b \pathstart r = z\) then have "homotopic_paths S (f \ r) (p \ linepath a a)" by (simp add: o_def feq linepath_def) then show "\q. path q \ path_image q \ C \ pathstart q = a \ pathfinish q = a \ homotopic_paths S (f \ r) (p \ q)" by (force simp: \a \ C\) qed auto corollary covering_space_lift_strong: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" "a \ C" "z \ U" and scU: "simply_connected U" and lpcU: "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq]) show "path_connected U" using scU simply_connected_eq_contractible_loop_some by blast fix r assume r: "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" have "linepath (f z) (f z) = f \ linepath z z" by (simp add: o_def linepath_def) then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" by (metis r contf fim homotopic_paths_continuous_image scU simply_connected_eq_contractible_path) then show "\b. homotopic_paths S (f \ r) (linepath b b)" by blast qed blast corollary covering_space_lift: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and U: "simply_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" proof (cases "U = {}") case True with that show ?thesis by auto next case False then obtain z where "z \ U" by blast then obtain a where "a \ C" "f z = p a" by (metis cov covering_space_imp_surjective fim image_iff image_subset_iff) then show ?thesis by (metis that covering_space_lift_strong [OF cov _ \z \ U\ U contf fim]) qed subsection\<^marker>\tag unimportant\ \Homeomorphisms of arc images\ lemma homeomorphism_arc: fixes g :: "real \ 'a::t2_space" assumes "arc g" obtains h where "homeomorphism {0..1} (path_image g) g h" using assms by (force simp: arc_def homeomorphism_compact path_def path_image_def) lemma homeomorphic_arc_image_interval: fixes g :: "real \ 'a::t2_space" and a::real assumes "arc g" "a < b" shows "(path_image g) homeomorphic {a..b}" proof - have "(path_image g) homeomorphic {0..1::real}" by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc) also have "\ homeomorphic {a..b}" using assms by (force intro: homeomorphic_closed_intervals_real) finally show ?thesis . qed lemma homeomorphic_arc_images: fixes g :: "real \ 'a::t2_space" and h :: "real \ 'b::t2_space" assumes "arc g" "arc h" shows "(path_image g) homeomorphic (path_image h)" proof - have "(path_image g) homeomorphic {0..1::real}" by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc) also have "\ homeomorphic (path_image h)" by (meson assms homeomorphic_def homeomorphism_arc) finally show ?thesis . qed end diff --git a/src/HOL/Analysis/Line_Segment.thy b/src/HOL/Analysis/Line_Segment.thy --- a/src/HOL/Analysis/Line_Segment.thy +++ b/src/HOL/Analysis/Line_Segment.thy @@ -1,1195 +1,1195 @@ (* Title: HOL/Analysis/Line_Segment.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) section \Line Segment\ theory Line_Segment imports Convex Topology_Euclidean_Space begin subsection\<^marker>\tag unimportant\ \Topological Properties of Convex Sets, Metric Spaces and Functions\ lemma convex_supp_sum: assumes "convex S" and 1: "supp_sum u I = 1" and "\i. i \ I \ 0 \ u i \ (u i = 0 \ f i \ S)" shows "supp_sum (\i. u i *\<^sub>R f i) I \ S" proof - have fin: "finite {i \ I. u i \ 0}" using 1 sum.infinite by (force simp: supp_sum_def support_on_def) then have "supp_sum (\i. u i *\<^sub>R f i) I = sum (\i. u i *\<^sub>R f i) {i \ I. u i \ 0}" by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def) also have "... \ S" using 1 assms by (force simp: supp_sum_def support_on_def intro: convex_sum [OF fin \convex S\]) finally show ?thesis . qed lemma sphere_eq_empty [simp]: fixes a :: "'a::{real_normed_vector, perfect_space}" shows "sphere a r = {} \ r < 0" by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) lemma cone_closure: fixes S :: "'a::real_normed_vector set" assumes "cone S" shows "cone (closure S)" proof (cases "S = {}") case True then show ?thesis by auto next case False then have "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have "0 \ closure S \ (\c. c > 0 \ (*\<^sub>R) c ` closure S = closure S)" using closure_subset by (auto simp: closure_scaleR) then show ?thesis using False cone_iff[of "closure S"] by auto qed corollary component_complement_connected: fixes S :: "'a::real_normed_vector set" assumes "connected S" "C \ components (-S)" shows "connected(-C)" using component_diff_connected [of S UNIV] assms by (auto simp: Compl_eq_Diff_UNIV) proposition clopen: fixes S :: "'a :: real_normed_vector set" shows "closed S \ open S \ S = {} \ S = UNIV" by (force intro!: connected_UNIV [unfolded connected_clopen, rule_format]) corollary compact_open: fixes S :: "'a :: euclidean_space set" shows "compact S \ open S \ S = {}" by (auto simp: compact_eq_bounded_closed clopen) corollary finite_imp_not_open: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "\finite S; open S\ \ S={}" using clopen [of S] finite_imp_closed not_bounded_UNIV by blast corollary empty_interior_finite: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "finite S \ interior S = {}" by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open) text \Balls, being convex, are connected.\ lemma convex_local_global_minimum: fixes s :: "'a::real_normed_vector set" assumes "e > 0" and "convex_on s f" and "ball x e \ s" and "\y\ball x e. f x \ f y" shows "\y\s. f x \ f y" proof (rule ccontr) have "x \ s" using assms(1,3) by auto assume "\ ?thesis" then obtain y where "y\s" and y: "f x > f y" by auto then have xy: "0 < dist x y" by auto then obtain u where "0 < u" "u \ 1" and u: "u < e / dist x y" using field_lbound_gt_zero[of 1 "e / dist x y"] xy \e>0\ by auto then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \ (1-u) * f x + u * f y" using \x\s\ \y\s\ using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] by auto moreover have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" by (simp add: algebra_simps) have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ ball x e" unfolding mem_ball dist_norm unfolding * and norm_scaleR and abs_of_pos[OF \0] unfolding dist_norm[symmetric] using u unfolding pos_less_divide_eq[OF xy] by auto then have "f x \ f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" using assms(4) by auto ultimately show False using mult_strict_left_mono[OF y \u>0\] unfolding left_diff_distrib by auto qed lemma convex_ball [iff]: fixes x :: "'a::real_normed_vector" shows "convex (ball x e)" proof (auto simp: convex_def) fix y z assume yz: "dist x y < e" "dist x z < e" fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ u * dist x y + v * dist x z" using uv yz using convex_on_dist [of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" using convex_bound_lt[OF yz uv] by auto qed lemma convex_cball [iff]: fixes x :: "'a::real_normed_vector" shows "convex (cball x e)" proof - { fix y z assume yz: "dist x y \ e" "dist x z \ e" fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ u * dist x y + v * dist x z" using uv yz using convex_on_dist [of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ e" using convex_bound_le[OF yz uv] by auto } then show ?thesis by (auto simp: convex_def Ball_def) qed lemma connected_ball [iff]: fixes x :: "'a::real_normed_vector" shows "connected (ball x e)" using convex_connected convex_ball by auto lemma connected_cball [iff]: fixes x :: "'a::real_normed_vector" shows "connected (cball x e)" using convex_connected convex_cball by auto lemma bounded_convex_hull: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "bounded (convex hull s)" proof - from assms obtain B where B: "\x\s. norm x \ B" unfolding bounded_iff by auto show ?thesis by (simp add: bounded_subset[OF bounded_cball, of _ 0 B] B subsetI subset_hull) qed lemma finite_imp_bounded_convex_hull: fixes s :: "'a::real_normed_vector set" shows "finite s \ bounded (convex hull s)" using bounded_convex_hull finite_imp_bounded by auto subsection \Midpoint\ definition\<^marker>\tag important\ midpoint :: "'a::real_vector \ 'a \ 'a" where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" lemma midpoint_idem [simp]: "midpoint x x = x" unfolding midpoint_def by simp lemma midpoint_sym: "midpoint a b = midpoint b a" unfolding midpoint_def by (auto simp add: scaleR_right_distrib) lemma midpoint_eq_iff: "midpoint a b = c \ a + b = c + c" proof - have "midpoint a b = c \ scaleR 2 (midpoint a b) = scaleR 2 c" by simp then show ?thesis unfolding midpoint_def scaleR_2 [symmetric] by simp qed lemma fixes a::real assumes "a \ b" shows ge_midpoint_1: "a \ midpoint a b" and le_midpoint_1: "midpoint a b \ b" by (simp_all add: midpoint_def assms) lemma dist_midpoint: fixes a b :: "'a::real_normed_vector" shows "dist a (midpoint a b) = (dist a b) / 2" (is ?t1) "dist b (midpoint a b) = (dist a b) / 2" (is ?t2) "dist (midpoint a b) a = (dist a b) / 2" (is ?t3) "dist (midpoint a b) b = (dist a b) / 2" (is ?t4) proof - have *: "\x y::'a. 2 *\<^sub>R x = - y \ norm x = (norm y) / 2" unfolding equation_minus_iff by auto have **: "\x y::'a. 2 *\<^sub>R x = y \ norm x = (norm y) / 2" by auto note scaleR_right_distrib [simp] show ?t1 unfolding midpoint_def dist_norm apply (rule **) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done show ?t2 unfolding midpoint_def dist_norm apply (rule *) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done show ?t3 unfolding midpoint_def dist_norm apply (rule *) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done show ?t4 unfolding midpoint_def dist_norm apply (rule **) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done qed lemma midpoint_eq_endpoint [simp]: "midpoint a b = a \ a = b" "midpoint a b = b \ a = b" unfolding midpoint_eq_iff by auto lemma midpoint_plus_self [simp]: "midpoint a b + midpoint a b = a + b" using midpoint_eq_iff by metis lemma midpoint_linear_image: "linear f \ midpoint(f a)(f b) = f(midpoint a b)" by (simp add: linear_iff midpoint_def) subsection \Open and closed segments\ definition\<^marker>\tag important\ closed_segment :: "'a::real_vector \ 'a \ 'a set" where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \ u \ u \ 1}" definition\<^marker>\tag important\ open_segment :: "'a::real_vector \ 'a \ 'a set" where "open_segment a b \ closed_segment a b - {a,b}" lemmas segment = open_segment_def closed_segment_def lemma in_segment: "x \ closed_segment a b \ (\u. 0 \ u \ u \ 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" "x \ open_segment a b \ a \ b \ (\u. 0 < u \ u < 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" using less_eq_real_def by (auto simp: segment algebra_simps) lemma closed_segment_linear_image: "closed_segment (f a) (f b) = f ` (closed_segment a b)" if "linear f" proof - interpret linear f by fact show ?thesis by (force simp add: in_segment add scale) qed lemma open_segment_linear_image: "\linear f; inj f\ \ open_segment (f a) (f b) = f ` (open_segment a b)" by (force simp: open_segment_def closed_segment_linear_image inj_on_def) lemma closed_segment_translation: "closed_segment (c + a) (c + b) = image (\x. c + x) (closed_segment a b)" apply safe apply (rule_tac x="x-c" in image_eqI) apply (auto simp: in_segment algebra_simps) done lemma open_segment_translation: "open_segment (c + a) (c + b) = image (\x. c + x) (open_segment a b)" by (simp add: open_segment_def closed_segment_translation translation_diff) lemma closed_segment_of_real: "closed_segment (of_real x) (of_real y) = of_real ` closed_segment x y" apply (auto simp: image_iff in_segment scaleR_conv_of_real) apply (rule_tac x="(1-u)*x + u*y" in bexI) apply (auto simp: in_segment) done lemma open_segment_of_real: "open_segment (of_real x) (of_real y) = of_real ` open_segment x y" apply (auto simp: image_iff in_segment scaleR_conv_of_real) apply (rule_tac x="(1-u)*x + u*y" in bexI) apply (auto simp: in_segment) done lemma closed_segment_Reals: "\x \ Reals; y \ Reals\ \ closed_segment x y = of_real ` closed_segment (Re x) (Re y)" by (metis closed_segment_of_real of_real_Re) lemma open_segment_Reals: "\x \ Reals; y \ Reals\ \ open_segment x y = of_real ` open_segment (Re x) (Re y)" by (metis open_segment_of_real of_real_Re) lemma open_segment_PairD: "(x, x') \ open_segment (a, a') (b, b') \ (x \ open_segment a b \ a = b) \ (x' \ open_segment a' b' \ a' = b')" by (auto simp: in_segment) lemma closed_segment_PairD: "(x, x') \ closed_segment (a, a') (b, b') \ x \ closed_segment a b \ x' \ closed_segment a' b'" by (auto simp: closed_segment_def) lemma closed_segment_translation_eq [simp]: "d + x \ closed_segment (d + a) (d + b) \ x \ closed_segment a b" proof - have *: "\d x a b. x \ closed_segment a b \ d + x \ closed_segment (d + a) (d + b)" apply (simp add: closed_segment_def) apply (erule ex_forward) apply (simp add: algebra_simps) done show ?thesis using * [where d = "-d"] * by (fastforce simp add:) qed lemma open_segment_translation_eq [simp]: "d + x \ open_segment (d + a) (d + b) \ x \ open_segment a b" by (simp add: open_segment_def) lemma of_real_closed_segment [simp]: "of_real x \ closed_segment (of_real a) (of_real b) \ x \ closed_segment a b" apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward) using of_real_eq_iff by fastforce lemma of_real_open_segment [simp]: "of_real x \ open_segment (of_real a) (of_real b) \ x \ open_segment a b" apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward del: exE) using of_real_eq_iff by fastforce lemma convex_contains_segment: "convex S \ (\a\S. \b\S. closed_segment a b \ S)" unfolding convex_alt closed_segment_def by auto lemma closed_segment_in_Reals: "\x \ closed_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" by (meson subsetD convex_Reals convex_contains_segment) lemma open_segment_in_Reals: "\x \ open_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" by (metis Diff_iff closed_segment_in_Reals open_segment_def) lemma closed_segment_subset: "\x \ S; y \ S; convex S\ \ closed_segment x y \ S" by (simp add: convex_contains_segment) lemma closed_segment_subset_convex_hull: "\x \ convex hull S; y \ convex hull S\ \ closed_segment x y \ convex hull S" using convex_contains_segment by blast lemma segment_convex_hull: "closed_segment a b = convex hull {a,b}" proof - have *: "\x. {x} \ {}" by auto show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton by (safe; rule_tac x="1 - u" in exI; force) qed lemma open_closed_segment: "u \ open_segment w z \ u \ closed_segment w z" by (auto simp add: closed_segment_def open_segment_def) lemma segment_open_subset_closed: "open_segment a b \ closed_segment a b" by (auto simp: closed_segment_def open_segment_def) lemma bounded_closed_segment: fixes a :: "'a::real_normed_vector" shows "bounded (closed_segment a b)" by (rule boundedI[where B="max (norm a) (norm b)"]) (auto simp: closed_segment_def max_def convex_bound_le intro!: norm_triangle_le) lemma bounded_open_segment: fixes a :: "'a::real_normed_vector" shows "bounded (open_segment a b)" by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed]) lemmas bounded_segment = bounded_closed_segment open_closed_segment lemma ends_in_segment [iff]: "a \ closed_segment a b" "b \ closed_segment a b" unfolding segment_convex_hull by (auto intro!: hull_subset[unfolded subset_eq, rule_format]) lemma eventually_closed_segment: fixes x0::"'a::real_normed_vector" assumes "open X0" "x0 \ X0" shows "\\<^sub>F x in at x0 within U. closed_segment x0 x \ X0" proof - from openE[OF assms] obtain e where e: "0 < e" "ball x0 e \ X0" . then have "\\<^sub>F x in at x0 within U. x \ ball x0 e" by (auto simp: dist_commute eventually_at) then show ?thesis proof eventually_elim case (elim x) have "x0 \ ball x0 e" using \e > 0\ by simp from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim] have "closed_segment x0 x \ ball x0 e" . also note \\ \ X0\ finally show ?case . qed qed lemma closed_segment_commute: "closed_segment a b = closed_segment b a" proof - have "{a, b} = {b, a}" by auto thus ?thesis by (simp add: segment_convex_hull) qed lemma segment_bound1: assumes "x \ closed_segment a b" shows "norm (x - a) \ norm (b - a)" proof - obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" using assms by (auto simp add: closed_segment_def) then show "norm (x - a) \ norm (b - a)" apply clarify apply (auto simp: algebra_simps) apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le) done qed lemma segment_bound: assumes "x \ closed_segment a b" shows "norm (x - a) \ norm (b - a)" "norm (x - b) \ norm (b - a)" by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1)+ lemma open_segment_commute: "open_segment a b = open_segment b a" proof - have "{a, b} = {b, a}" by auto thus ?thesis by (simp add: closed_segment_commute open_segment_def) qed lemma closed_segment_idem [simp]: "closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps) lemma open_segment_idem [simp]: "open_segment a a = {}" by (simp add: open_segment_def) lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \ {a,b}" using open_segment_def by auto lemma convex_contains_open_segment: "convex s \ (\a\s. \b\s. open_segment a b \ s)" by (simp add: convex_contains_segment closed_segment_eq_open) lemma closed_segment_eq_real_ivl1: fixes a b::real assumes "a \ b" shows "closed_segment a b = {a .. b}" proof safe fix x assume "x \ closed_segment a b" then obtain u where u: "0 \ u" "u \ 1" and x_def: "x = (1 - u) * a + u * b" by (auto simp: closed_segment_def) have "u * a \ u * b" "(1 - u) * a \ (1 - u) * b" by (auto intro!: mult_left_mono u assms) then show "x \ {a .. b}" unfolding x_def by (auto simp: algebra_simps) next show "\x. x \ {a..b} \ x \ closed_segment a b" by (force simp: closed_segment_def divide_simps algebra_simps intro: exI[where x="(x - a) / (b - a)" for x]) qed lemma closed_segment_eq_real_ivl: fixes a b::real shows "closed_segment a b = (if a \ b then {a .. b} else {b .. a})" using closed_segment_eq_real_ivl1[of a b] closed_segment_eq_real_ivl1[of b a] by (auto simp: closed_segment_commute) lemma open_segment_eq_real_ivl: fixes a b::real shows "open_segment a b = (if a \ b then {a<..x. (v - u) * x + u) ` {0..1}" by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl) lemma closed_segment_same_Re: assumes "Re a = Re b" shows "closed_segment a b = {z. Re z = Re a \ Im z \ closed_segment (Im a) (Im b)}" proof safe fix z assume "z \ closed_segment a b" then obtain u where u: "u \ {0..1}" "z = a + of_real u * (b - a)" by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) from assms show "Re z = Re a" by (auto simp: u) from u(1) show "Im z \ closed_segment (Im a) (Im b)" by (force simp: u closed_segment_def algebra_simps) next fix z assume [simp]: "Re z = Re a" and "Im z \ closed_segment (Im a) (Im b)" then obtain u where u: "u \ {0..1}" "Im z = Im a + of_real u * (Im b - Im a)" by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) from u(1) show "z \ closed_segment a b" using assms by (force simp: u closed_segment_def algebra_simps scaleR_conv_of_real complex_eq_iff) qed lemma closed_segment_same_Im: assumes "Im a = Im b" shows "closed_segment a b = {z. Im z = Im a \ Re z \ closed_segment (Re a) (Re b)}" proof safe fix z assume "z \ closed_segment a b" then obtain u where u: "u \ {0..1}" "z = a + of_real u * (b - a)" by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) from assms show "Im z = Im a" by (auto simp: u) from u(1) show "Re z \ closed_segment (Re a) (Re b)" by (force simp: u closed_segment_def algebra_simps) next fix z assume [simp]: "Im z = Im a" and "Re z \ closed_segment (Re a) (Re b)" then obtain u where u: "u \ {0..1}" "Re z = Re a + of_real u * (Re b - Re a)" by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) from u(1) show "z \ closed_segment a b" using assms by (force simp: u closed_segment_def algebra_simps scaleR_conv_of_real complex_eq_iff) qed lemma dist_in_closed_segment: fixes a :: "'a :: euclidean_space" assumes "x \ closed_segment a b" shows "dist x a \ dist a b \ dist x b \ dist a b" proof (intro conjI) obtain u where u: "0 \ u" "u \ 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" using assms by (force simp: in_segment algebra_simps) have "dist x a = u * dist a b" apply (simp add: dist_norm algebra_simps x) by (metis \0 \ u\ abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib) also have "... \ dist a b" by (simp add: mult_left_le_one_le u) finally show "dist x a \ dist a b" . have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" by (simp add: dist_norm algebra_simps x) also have "... = (1-u) * dist a b" proof - have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" using \u \ 1\ by force then show ?thesis by (simp add: dist_norm real_vector.scale_right_diff_distrib) qed also have "... \ dist a b" by (simp add: mult_left_le_one_le u) finally show "dist x b \ dist a b" . qed lemma dist_in_open_segment: fixes a :: "'a :: euclidean_space" assumes "x \ open_segment a b" shows "dist x a < dist a b \ dist x b < dist a b" proof (intro conjI) obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" using assms by (force simp: in_segment algebra_simps) have "dist x a = u * dist a b" apply (simp add: dist_norm algebra_simps x) by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \0 < u\) also have *: "... < dist a b" - by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \u < 1\) + using assms mult_less_cancel_right2 u(2) by fastforce finally show "dist x a < dist a b" . have ab_ne0: "dist a b \ 0" using * by fastforce have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" by (simp add: dist_norm algebra_simps x) also have "... = (1-u) * dist a b" proof - have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" using \u < 1\ by force then show ?thesis by (simp add: dist_norm real_vector.scale_right_diff_distrib) qed also have "... < dist a b" using ab_ne0 \0 < u\ by simp finally show "dist x b < dist a b" . qed lemma dist_decreases_open_segment_0: fixes x :: "'a :: euclidean_space" assumes "x \ open_segment 0 b" shows "dist c x < dist c 0 \ dist c x < dist c b" proof (rule ccontr, clarsimp simp: not_less) obtain u where u: "0 \ b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b" using assms by (auto simp: in_segment) have xb: "x \ b < b \ b" using u x by auto assume "norm c \ dist c x" then have "c \ c \ (c - x) \ (c - x)" by (simp add: dist_norm norm_le) moreover have "0 < x \ b" using u x by auto ultimately have less: "c \ b < x \ b" by (simp add: x algebra_simps inner_commute u) assume "dist c b \ dist c x" then have "(c - b) \ (c - b) \ (c - x) \ (c - x)" by (simp add: dist_norm norm_le) then have "(b \ b) * (1 - u*u) \ 2 * (b \ c) * (1-u)" by (simp add: x algebra_simps inner_commute) then have "(1+u) * (b \ b) * (1-u) \ 2 * (b \ c) * (1-u)" by (simp add: algebra_simps) then have "(1+u) * (b \ b) \ 2 * (b \ c)" using \u < 1\ by auto with xb have "c \ b \ x \ b" by (auto simp: x algebra_simps inner_commute) with less show False by auto qed proposition dist_decreases_open_segment: fixes a :: "'a :: euclidean_space" assumes "x \ open_segment a b" shows "dist c x < dist c a \ dist c x < dist c b" proof - have *: "x - a \ open_segment 0 (b - a)" using assms by (metis diff_self open_segment_translation_eq uminus_add_conv_diff) show ?thesis using dist_decreases_open_segment_0 [OF *, of "c-a"] assms by (simp add: dist_norm) qed corollary open_segment_furthest_le: fixes a b x y :: "'a::euclidean_space" assumes "x \ open_segment a b" shows "norm (y - x) < norm (y - a) \ norm (y - x) < norm (y - b)" by (metis assms dist_decreases_open_segment dist_norm) corollary dist_decreases_closed_segment: fixes a :: "'a :: euclidean_space" assumes "x \ closed_segment a b" shows "dist c x \ dist c a \ dist c x \ dist c b" apply (cases "x \ open_segment a b") using dist_decreases_open_segment less_eq_real_def apply blast by (metis DiffI assms empty_iff insertE open_segment_def order_refl) corollary segment_furthest_le: fixes a b x y :: "'a::euclidean_space" assumes "x \ closed_segment a b" shows "norm (y - x) \ norm (y - a) \ norm (y - x) \ norm (y - b)" by (metis assms dist_decreases_closed_segment dist_norm) lemma convex_intermediate_ball: fixes a :: "'a :: euclidean_space" shows "\ball a r \ T; T \ cball a r\ \ convex T" apply (simp add: convex_contains_open_segment, clarify) by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment) lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \ closed_segment a b" apply (clarsimp simp: midpoint_def in_segment) apply (rule_tac x="(1 + u) / 2" in exI) apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib) by (metis field_sum_of_halves scaleR_left.add) lemma notin_segment_midpoint: fixes a :: "'a :: euclidean_space" shows "a \ b \ a \ closed_segment (midpoint a b) b" by (auto simp: dist_midpoint dest!: dist_in_closed_segment) subsubsection\More lemmas, especially for working with the underlying formula\ lemma segment_eq_compose: fixes a :: "'a :: real_vector" shows "(\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\x. a + x) o (\u. u *\<^sub>R (b - a))" by (simp add: o_def algebra_simps) lemma segment_degen_1: fixes a :: "'a :: real_vector" shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \ a=b \ u=1" proof - { assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b" then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b" by (simp add: algebra_simps) then have "a=b \ u=1" by simp } then show ?thesis by (auto simp: algebra_simps) qed lemma segment_degen_0: fixes a :: "'a :: real_vector" shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \ a=b \ u=0" using segment_degen_1 [of "1-u" b a] by (auto simp: algebra_simps) lemma add_scaleR_degen: fixes a b ::"'a::real_vector" assumes "(u *\<^sub>R b + v *\<^sub>R a) = (u *\<^sub>R a + v *\<^sub>R b)" "u \ v" shows "a=b" by (metis (no_types, hide_lams) add.commute add_diff_eq diff_add_cancel real_vector.scale_cancel_left real_vector.scale_left_diff_distrib assms) lemma closed_segment_image_interval: "closed_segment a b = (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}" by (auto simp: set_eq_iff image_iff closed_segment_def) lemma open_segment_image_interval: "open_segment a b = (if a=b then {} else (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})" by (auto simp: open_segment_def closed_segment_def segment_degen_0 segment_degen_1) lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval lemma closed_segment_neq_empty [simp]: "closed_segment a b \ {}" by auto lemma open_segment_eq_empty [simp]: "open_segment a b = {} \ a = b" proof - { assume a1: "open_segment a b = {}" have "{} \ {0::real<..<1}" by simp then have "a = b" using a1 open_segment_image_interval by fastforce } then show ?thesis by auto qed lemma open_segment_eq_empty' [simp]: "{} = open_segment a b \ a = b" using open_segment_eq_empty by blast lemmas segment_eq_empty = closed_segment_neq_empty open_segment_eq_empty lemma inj_segment: fixes a :: "'a :: real_vector" assumes "a \ b" shows "inj_on (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) I" proof fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" then have "x *\<^sub>R (b - a) = y *\<^sub>R (b - a)" by (simp add: algebra_simps) with assms show "x = y" by (simp add: real_vector.scale_right_imp_eq) qed lemma finite_closed_segment [simp]: "finite(closed_segment a b) \ a = b" apply auto apply (rule ccontr) apply (simp add: segment_image_interval) using infinite_Icc [OF zero_less_one] finite_imageD [OF _ inj_segment] apply blast done lemma finite_open_segment [simp]: "finite(open_segment a b) \ a = b" by (auto simp: open_segment_def) lemmas finite_segment = finite_closed_segment finite_open_segment lemma closed_segment_eq_sing: "closed_segment a b = {c} \ a = c \ b = c" by auto lemma open_segment_eq_sing: "open_segment a b \ {c}" by (metis finite_insert finite_open_segment insert_not_empty open_segment_image_interval) lemmas segment_eq_sing = closed_segment_eq_sing open_segment_eq_sing lemma open_segment_bound1: assumes "x \ open_segment a b" shows "norm (x - a) < norm (b - a)" proof - obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \ b" using assms by (auto simp add: open_segment_image_interval split: if_split_asm) then show "norm (x - a) < norm (b - a)" apply clarify apply (auto simp: algebra_simps) apply (simp add: scaleR_diff_right [symmetric]) done qed lemma compact_segment [simp]: fixes a :: "'a::real_normed_vector" shows "compact (closed_segment a b)" by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros) lemma closed_segment [simp]: fixes a :: "'a::real_normed_vector" shows "closed (closed_segment a b)" by (simp add: compact_imp_closed) lemma closure_closed_segment [simp]: fixes a :: "'a::real_normed_vector" shows "closure(closed_segment a b) = closed_segment a b" by simp lemma open_segment_bound: assumes "x \ open_segment a b" shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)" apply (simp add: assms open_segment_bound1) by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute) lemma closure_open_segment [simp]: "closure (open_segment a b) = (if a = b then {} else closed_segment a b)" for a :: "'a::euclidean_space" proof (cases "a = b") case True then show ?thesis by simp next case False have "closure ((\u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\u. u *\<^sub>R (b - a)) ` closure {0<..<1}" apply (rule closure_injective_linear_image [symmetric]) apply (use False in \auto intro!: injI\) done then have "closure ((\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1}) = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b) ` closure {0<..<1}" using closure_translation [of a "((\x. x *\<^sub>R b - x *\<^sub>R a) ` {0<..<1})"] by (simp add: segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right image_image) then show ?thesis by (simp add: segment_image_interval closure_greaterThanLessThan [symmetric] del: closure_greaterThanLessThan) qed lemma closed_open_segment_iff [simp]: fixes a :: "'a::euclidean_space" shows "closed(open_segment a b) \ a = b" by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2)) lemma compact_open_segment_iff [simp]: fixes a :: "'a::euclidean_space" shows "compact(open_segment a b) \ a = b" by (simp add: bounded_open_segment compact_eq_bounded_closed) lemma convex_closed_segment [iff]: "convex (closed_segment a b)" unfolding segment_convex_hull by(rule convex_convex_hull) lemma convex_open_segment [iff]: "convex (open_segment a b)" proof - have "convex ((\u. u *\<^sub>R (b - a)) ` {0<..<1})" by (rule convex_linear_image) auto then have "convex ((+) a ` (\u. u *\<^sub>R (b - a)) ` {0<..<1})" by (rule convex_translation) then show ?thesis by (simp add: image_image open_segment_image_interval segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right) qed lemmas convex_segment = convex_closed_segment convex_open_segment lemma subset_closed_segment: "closed_segment a b \ closed_segment c d \ a \ closed_segment c d \ b \ closed_segment c d" by auto (meson contra_subsetD convex_closed_segment convex_contains_segment) lemma subset_co_segment: "closed_segment a b \ open_segment c d \ a \ open_segment c d \ b \ open_segment c d" using closed_segment_subset by blast lemma subset_open_segment: fixes a :: "'a::euclidean_space" shows "open_segment a b \ open_segment c d \ a = b \ a \ closed_segment c d \ b \ closed_segment c d" (is "?lhs = ?rhs") proof (cases "a = b") case True then show ?thesis by simp next case False show ?thesis proof assume rhs: ?rhs with \a \ b\ have "c \ d" using closed_segment_idem singleton_iff by auto have "\uc. (1 - u) *\<^sub>R ((1 - ua) *\<^sub>R c + ua *\<^sub>R d) + u *\<^sub>R ((1 - ub) *\<^sub>R c + ub *\<^sub>R d) = (1 - uc) *\<^sub>R c + uc *\<^sub>R d \ 0 < uc \ uc < 1" if neq: "(1 - ua) *\<^sub>R c + ua *\<^sub>R d \ (1 - ub) *\<^sub>R c + ub *\<^sub>R d" "c \ d" and "a = (1 - ua) *\<^sub>R c + ua *\<^sub>R d" "b = (1 - ub) *\<^sub>R c + ub *\<^sub>R d" and u: "0 < u" "u < 1" and uab: "0 \ ua" "ua \ 1" "0 \ ub" "ub \ 1" for u ua ub proof - have "ua \ ub" using neq by auto moreover have "(u - 1) * ua \ 0" using u uab by (simp add: mult_nonpos_nonneg) ultimately have lt: "(u - 1) * ua < u * ub" using u uab by (metis antisym_conv diff_ge_0_iff_ge le_less_trans mult_eq_0_iff mult_le_0_iff not_less) have "p * ua + q * ub < p+q" if p: "0 < p" and q: "0 < q" for p q proof - have "\ p \ 0" "\ q \ 0" using p q not_less by blast+ then show ?thesis by (metis \ua \ ub\ add_less_cancel_left add_less_cancel_right add_mono_thms_linordered_field(5) less_eq_real_def mult_cancel_left1 mult_less_cancel_left2 uab(2) uab(4)) qed then have "(1 - u) * ua + u * ub < 1" using u \ua \ ub\ by (metis diff_add_cancel diff_gt_0_iff_gt) with lt show ?thesis by (rule_tac x="ua + u*(ub-ua)" in exI) (simp add: algebra_simps) qed with rhs \a \ b\ \c \ d\ show ?lhs unfolding open_segment_image_interval closed_segment_def by (fastforce simp add:) next assume lhs: ?lhs with \a \ b\ have "c \ d" by (meson finite_open_segment rev_finite_subset) have "closure (open_segment a b) \ closure (open_segment c d)" using lhs closure_mono by blast then have "closed_segment a b \ closed_segment c d" by (simp add: \a \ b\ \c \ d\) then show ?rhs by (force simp: \a \ b\) qed qed lemma subset_oc_segment: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d \ a = b \ a \ closed_segment c d \ b \ closed_segment c d" apply (simp add: subset_open_segment [symmetric]) apply (rule iffI) apply (metis closure_closed_segment closure_mono closure_open_segment subset_closed_segment subset_open_segment) apply (meson dual_order.trans segment_open_subset_closed) done lemmas subset_segment = subset_closed_segment subset_co_segment subset_oc_segment subset_open_segment lemma dist_half_times2: fixes a :: "'a :: real_normed_vector" shows "dist ((1 / 2) *\<^sub>R (a + b)) x * 2 = dist (a+b) (2 *\<^sub>R x)" proof - have "norm ((1 / 2) *\<^sub>R (a + b) - x) * 2 = norm (2 *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x))" by simp also have "... = norm ((a + b) - 2 *\<^sub>R x)" by (simp add: real_vector.scale_right_diff_distrib) finally show ?thesis by (simp only: dist_norm) qed lemma closed_segment_as_ball: "closed_segment a b = affine hull {a,b} \ cball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)" proof (cases "b = a") case True then show ?thesis by (auto simp: hull_inc) next case False then have *: "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a)) = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1)" for x proof - have "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a)) = ((\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a))" unfolding eq_diff_eq [symmetric] by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R x)) \ norm (b - a))" by (simp add: dist_half_times2) (simp add: dist_norm) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) \ norm (b - a))" by auto also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((1 - u * 2) *\<^sub>R (b - a)) \ norm (b - a))" by (simp add: algebra_simps scaleR_2) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ * norm (b - a) \ norm (b - a))" by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ \ 1)" by (simp add: mult_le_cancel_right2 False) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1)" by auto finally show ?thesis . qed show ?thesis by (simp add: affine_hull_2 Set.set_eq_iff closed_segment_def *) qed lemma open_segment_as_ball: "open_segment a b = affine hull {a,b} \ ball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)" proof (cases "b = a") case True then show ?thesis by (auto simp: hull_inc) next case False then have *: "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 < u \ u < 1)" for x proof - have "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = ((\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a))" unfolding eq_diff_eq [symmetric] by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R x)) < norm (b - a))" by (simp add: dist_half_times2) (simp add: dist_norm) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) < norm (b - a))" by auto also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((1 - u * 2) *\<^sub>R (b - a)) < norm (b - a))" by (simp add: algebra_simps scaleR_2) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ * norm (b - a) < norm (b - a))" by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ < 1)" by (simp add: mult_le_cancel_right2 False) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 < u \ u < 1)" by auto finally show ?thesis . qed show ?thesis using False by (force simp: affine_hull_2 Set.set_eq_iff open_segment_image_interval *) qed lemmas segment_as_ball = closed_segment_as_ball open_segment_as_ball lemma connected_segment [iff]: fixes x :: "'a :: real_normed_vector" shows "connected (closed_segment x y)" by (simp add: convex_connected) lemma is_interval_closed_segment_1[intro, simp]: "is_interval (closed_segment a b)" for a b::real unfolding closed_segment_eq_real_ivl by (auto simp: is_interval_def) lemma IVT'_closed_segment_real: fixes f :: "real \ real" assumes "y \ closed_segment (f a) (f b)" assumes "continuous_on (closed_segment a b) f" shows "\x \ closed_segment a b. f x = y" using IVT'[of f a y b] IVT'[of "-f" a "-y" b] IVT'[of f b y a] IVT'[of "-f" b "-y" a] assms by (cases "a \ b"; cases "f b \ f a") (auto simp: closed_segment_eq_real_ivl continuous_on_minus) subsection \Betweenness\ definition\<^marker>\tag important\ "between = (\(a,b) x. x \ closed_segment a b)" lemma betweenI: assumes "0 \ u" "u \ 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" shows "between (a, b) x" using assms unfolding between_def closed_segment_def by auto lemma betweenE: assumes "between (a, b) x" obtains u where "0 \ u" "u \ 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" using assms unfolding between_def closed_segment_def by auto lemma between_implies_scaled_diff: assumes "between (S, T) X" "between (S, T) Y" "S \ Y" obtains c where "(X - Y) = c *\<^sub>R (S - Y)" proof - from \between (S, T) X\ obtain u\<^sub>X where X: "X = u\<^sub>X *\<^sub>R S + (1 - u\<^sub>X) *\<^sub>R T" by (metis add.commute betweenE eq_diff_eq) from \between (S, T) Y\ obtain u\<^sub>Y where Y: "Y = u\<^sub>Y *\<^sub>R S + (1 - u\<^sub>Y) *\<^sub>R T" by (metis add.commute betweenE eq_diff_eq) have "X - Y = (u\<^sub>X - u\<^sub>Y) *\<^sub>R (S - T)" proof - from X Y have "X - Y = u\<^sub>X *\<^sub>R S - u\<^sub>Y *\<^sub>R S + ((1 - u\<^sub>X) *\<^sub>R T - (1 - u\<^sub>Y) *\<^sub>R T)" by simp also have "\ = (u\<^sub>X - u\<^sub>Y) *\<^sub>R S - (u\<^sub>X - u\<^sub>Y) *\<^sub>R T" by (simp add: scaleR_left.diff) finally show ?thesis by (simp add: real_vector.scale_right_diff_distrib) qed moreover from Y have "S - Y = (1 - u\<^sub>Y) *\<^sub>R (S - T)" by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) moreover note \S \ Y\ ultimately have "(X - Y) = ((u\<^sub>X - u\<^sub>Y) / (1 - u\<^sub>Y)) *\<^sub>R (S - Y)" by auto from this that show thesis by blast qed lemma between_mem_segment: "between (a,b) x \ x \ closed_segment a b" unfolding between_def by auto lemma between: "between (a, b) (x::'a::euclidean_space) \ dist a b = (dist a x) + (dist x b)" proof (cases "a = b") case True then show ?thesis by (auto simp add: between_def dist_commute) next case False then have Fal: "norm (a - b) \ 0" and Fal2: "norm (a - b) > 0" by auto have *: "\u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps) have "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" if "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" for u proof - have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" unfolding that(1) by (auto simp add:algebra_simps) show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" unfolding norm_minus_commute[of x a] * using \0 \ u\ \u \ 1\ by simp qed moreover have "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" if "dist a b = dist a x + dist x b" proof - let ?\ = "norm (a - x) / norm (a - b)" show "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" proof (intro exI conjI) show "?\ \ 1" using Fal2 unfolding that[unfolded dist_norm] norm_ge_zero by auto show "x = (1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b" proof (subst euclidean_eq_iff; intro ballI) fix i :: 'a assume i: "i \ Basis" have "((1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b) \ i = ((norm (a - b) - norm (a - x)) * (a \ i) + norm (a - x) * (b \ i)) / norm (a - b)" using Fal by (auto simp add: field_simps inner_simps) also have "\ = x\i" apply (rule divide_eq_imp[OF Fal]) unfolding that[unfolded dist_norm] using that[unfolded dist_triangle_eq] i apply (subst (asm) euclidean_eq_iff) apply (auto simp add: field_simps inner_simps) done finally show "x \ i = ((1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b) \ i" by auto qed qed (use Fal2 in auto) qed ultimately show ?thesis by (force simp add: between_def closed_segment_def dist_triangle_eq) qed lemma between_midpoint: fixes a :: "'a::euclidean_space" shows "between (a,b) (midpoint a b)" (is ?t1) and "between (b,a) (midpoint a b)" (is ?t2) proof - have *: "\x y z. x = (1/2::real) *\<^sub>R z \ y = (1/2) *\<^sub>R z \ norm z = norm x + norm y" by auto show ?t1 ?t2 unfolding between midpoint_def dist_norm by (auto simp add: field_simps inner_simps euclidean_eq_iff[where 'a='a] intro!: *) qed lemma between_mem_convex_hull: "between (a,b) x \ x \ convex hull {a,b}" unfolding between_mem_segment segment_convex_hull .. lemma between_triv_iff [simp]: "between (a,a) b \ a=b" by (auto simp: between_def) lemma between_triv1 [simp]: "between (a,b) a" by (auto simp: between_def) lemma between_triv2 [simp]: "between (a,b) b" by (auto simp: between_def) lemma between_commute: "between (a,b) = between (b,a)" by (auto simp: between_def closed_segment_commute) lemma between_antisym: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,c) b\ \ a = b" by (auto simp: between dist_commute) lemma between_trans: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,c) d\ \ between (b,c) d" using dist_triangle2 [of b c d] dist_triangle3 [of b d a] by (auto simp: between dist_commute) lemma between_norm: fixes a :: "'a :: euclidean_space" shows "between (a,b) x \ norm(x - a) *\<^sub>R (b - x) = norm(b - x) *\<^sub>R (x - a)" by (auto simp: between dist_triangle_eq norm_minus_commute algebra_simps) lemma between_swap: fixes A B X Y :: "'a::euclidean_space" assumes "between (A, B) X" assumes "between (A, B) Y" shows "between (X, B) Y \ between (A, Y) X" using assms by (auto simp add: between) lemma between_translation [simp]: "between (a + y,a + z) (a + x) \ between (y,z) x" by (auto simp: between_def) lemma between_trans_2: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,b) d\ \ between (c,d) a" by (metis between_commute between_swap between_trans) lemma between_scaleR_lift [simp]: fixes v :: "'a::euclidean_space" shows "between (a *\<^sub>R v, b *\<^sub>R v) (c *\<^sub>R v) \ v = 0 \ between (a, b) c" by (simp add: between dist_norm scaleR_left_diff_distrib [symmetric] distrib_right [symmetric]) lemma between_1: fixes x::real shows "between (a,b) x \ (a \ x \ x \ b) \ (b \ x \ x \ a)" by (auto simp: between_mem_segment closed_segment_eq_real_ivl) end \ No newline at end of file diff --git a/src/HOL/Analysis/Poly_Roots.thy b/src/HOL/Analysis/Poly_Roots.thy --- a/src/HOL/Analysis/Poly_Roots.thy +++ b/src/HOL/Analysis/Poly_Roots.thy @@ -1,234 +1,235 @@ (* Author: John Harrison and Valentina Bruno Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson *) section \Polynomial Functions: Extremal Behaviour and Root Counts\ theory Poly_Roots imports Complex_Main begin subsection\Basics about polynomial functions: extremal behaviour and root counts\ lemma sub_polyfun: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\jk= Suc j..n. a k * y^(k - Suc j) * x^j)" proof - have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (\i\n. a i * (x^i - y^i))" by (simp add: algebra_simps sum_subtractf [symmetric]) also have "... = (\i\n. a i * (x - y) * (\ji\n. (\jji=Suc j..n. a i * y^(i - Suc j) * x^j))" by (simp add: sum.nested_swap') finally show ?thesis . qed lemma sub_polyfun_alt: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\jkk = Suc j..n. a k * y^(k - Suc j) * x^j) = (\k b. \z. (\i\n. c i * z^i) = (z-a) * (\ii\n. c i * a^i)" proof - { fix z have "(\i\n. c i * z^i) - (\i\n. c i * a^i) = (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j)" by (simp add: sub_polyfun sum_distrib_right) then have "(\i\n. c i * z^i) = (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j) + (\i\n. c i * a^i)" by (simp add: algebra_simps) } then show ?thesis by (intro exI allI) qed lemma polyfun_linear_factor_root: fixes a :: "'a::{comm_ring,monoid_mult}" assumes "(\i\n. c i * a^i) = 0" shows "\b. \z. (\i\n. c i * z^i) = (z-a) * (\i b ==> norm(x) \ a ==> norm(x + y) \ b" by (metis norm_triangle_mono order.trans order_refl) proposition polyfun_extremal_lemma: fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "e > 0" shows "\M. \z. M \ norm z \ norm(\i\n. c i * z^i) \ e * norm(z) ^ Suc n" proof (induction n) case 0 show ?case by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) next case (Suc n) then obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" .. show ?case proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify) fix z::'a assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \ norm z" then have norm1: "0 < norm z" "M \ norm z" "(e + norm (c (Suc n))) / e \ norm z" by auto then have norm2: "(e + norm (c (Suc n))) \ e * norm z" "(norm z * norm z ^ n) > 0" apply (metis assms less_divide_eq mult.commute not_le) using norm1 apply (metis mult_pos_pos zero_less_power) done have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) = (e + norm (c (Suc n))) * (norm z * norm z ^ n)" by (simp add: norm_mult norm_power algebra_simps) also have "... \ (e * norm z) * (norm z * norm z ^ n)" - using norm2 by (metis real_mult_le_cancel_iff1) + using norm2 + using assms mult_mono by fastforce also have "... = e * (norm z * (norm z * norm z ^ n))" by (simp add: algebra_simps) finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) \ e * (norm z * (norm z * norm z ^ n))" . then show "norm (\i\Suc n. c i * z^i) \ e * norm z ^ Suc (Suc n)" using M norm1 by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle) qed qed lemma norm_lemma_xy: assumes "\b\ + 1 \ norm(y) - a" "norm(x) \ a" shows "b \ norm(x + y)" proof - have "b \ norm y - norm x" using assms by linarith then show ?thesis by (metis (no_types) add.commute norm_diff_ineq order_trans) qed proposition polyfun_extremal: fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "\k. k \ 0 \ k \ n \ c k \ 0" shows "eventually (\z. norm(\i\n. c i * z^i) \ B) at_infinity" using assms proof (induction n) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases "c (Suc n) = 0") case True with Suc show ?thesis by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq) next case False with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n] obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ norm (c (Suc n)) / 2 * norm z ^ Suc n" by auto show ?thesis unfolding eventually_at_infinity proof (rule exI [where x="max M (max 1 ((\B\ + 1) / (norm (c (Suc n)) / 2)))"], clarsimp) fix z::'a assume les: "M \ norm z" "1 \ norm z" "(\B\ * 2 + 2) / norm (c (Suc n)) \ norm z" then have "\B\ * 2 + 2 \ norm z * norm (c (Suc n))" by (metis False pos_divide_le_eq zero_less_norm_iff) then have "\B\ * 2 + 2 \ norm z ^ (Suc n) * norm (c (Suc n))" by (metis \1 \ norm z\ order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc) then show "B \ norm ((\i\n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les apply auto apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"]) apply (simp_all add: norm_mult norm_power) done qed qed qed proposition polyfun_rootbound: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" assumes "\k. k \ n \ c k \ 0" shows "finite {z. (\i\n. c i * z^i) = 0} \ card {z. (\i\n. c i * z^i) = 0} \ n" using assms proof (induction n arbitrary: c) case (Suc n) show ?case proof (cases "{z. (\i\Suc n. c i * z^i) = 0} = {}") case False then obtain a where a: "(\i\Suc n. c i * a^i) = 0" by auto from polyfun_linear_factor_root [OF this] obtain b where "\z. (\i\Suc n. c i * z^i) = (z - a) * (\i< Suc n. b i * z^i)" by auto then have b: "\z. (\i\Suc n. c i * z^i) = (z - a) * (\i\n. b i * z^i)" by (metis lessThan_Suc_atMost) then have ins_ab: "{z. (\i\Suc n. c i * z^i) = 0} = insert a {z. (\i\n. b i * z^i) = 0}" by auto have c0: "c 0 = - (a * b 0)" using b [of 0] by simp then have extr_prem: "\ (\k\n. b k \ 0) \ \k. k \ 0 \ k \ Suc n \ c k \ 0" by (metis Suc.prems le0 minus_zero mult_zero_right) have "\k\n. b k \ 0" apply (rule ccontr) using polyfun_extremal [OF extr_prem, of 1] apply (auto simp: eventually_at_infinity b simp del: sum.atMost_Suc) apply (drule_tac x="of_real ba" in spec, simp) done then show ?thesis using Suc.IH [of b] ins_ab by (auto simp: card_insert_if) qed simp qed simp corollary fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" assumes "\k. k \ n \ c k \ 0" shows polyfun_rootbound_finite: "finite {z. (\i\n. c i * z^i) = 0}" and polyfun_rootbound_card: "card {z. (\i\n. c i * z^i) = 0} \ n" using polyfun_rootbound [OF assms] by auto proposition polyfun_finite_roots: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "finite {z. (\i\n. c i * z^i) = 0} \ (\k. k \ n \ c k \ 0)" proof (cases " \k\n. c k \ 0") case True then show ?thesis by (blast intro: polyfun_rootbound_finite) next case False then show ?thesis by (auto simp: infinite_UNIV_char_0) qed lemma polyfun_eq_0: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "(\z. (\i\n. c i * z^i) = 0) \ (\k. k \ n \ c k = 0)" proof (cases "(\z. (\i\n. c i * z^i) = 0)") case True then have "\ finite {z. (\i\n. c i * z^i) = 0}" by (simp add: infinite_UNIV_char_0) with True show ?thesis by (metis (poly_guards_query) polyfun_rootbound_finite) next case False then show ?thesis by auto qed theorem polyfun_eq_const: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "(\z. (\i\n. c i * z^i) = k) \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" proof - {fix z have "(\i\n. c i * z^i) = (\i\n. (if i = 0 then c 0 - k else c i) * z^i) + k" by (induct n) auto } then have "(\z. (\i\n. c i * z^i) = k) \ (\z. (\i\n. (if i = 0 then c 0 - k else c i) * z^i) = 0)" by auto also have "... \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" by (auto simp: polyfun_eq_0) finally show ?thesis . qed end diff --git a/src/HOL/Analysis/Polytope.thy b/src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy +++ b/src/HOL/Analysis/Polytope.thy @@ -1,3913 +1,3913 @@ section \Faces, Extreme Points, Polytopes, Polyhedra etc\ text\Ported from HOL Light by L C Paulson\ theory Polytope imports Cartesian_Euclidean_Space Path_Connected begin subsection \Faces of a (usually convex) set\ definition\<^marker>\tag important\ face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) where "T face_of S \ T \ S \ convex T \ (\a \ S. \b \ S. \x \ T. x \ open_segment a b \ a \ T \ b \ T)" lemma face_ofD: "\T face_of S; x \ open_segment a b; a \ S; b \ S; x \ T\ \ a \ T \ b \ T" unfolding face_of_def by blast lemma face_of_translation_eq [simp]: "((+) a ` T face_of (+) a ` S) \ T face_of S" proof - have *: "\a T S. T face_of S \ ((+) a ` T face_of (+) a ` S)" by (simp add: face_of_def) show ?thesis by (force simp: image_comp o_def dest: * [where a = "-a"] intro: *) qed lemma face_of_linear_image: assumes "linear f" "inj f" shows "(f ` c face_of f ` S) \ c face_of S" by (simp add: face_of_def inj_image_subset_iff inj_image_mem_iff open_segment_linear_image assms) lemma face_of_refl: "convex S \ S face_of S" by (auto simp: face_of_def) lemma face_of_refl_eq: "S face_of S \ convex S" by (auto simp: face_of_def) lemma empty_face_of [iff]: "{} face_of S" by (simp add: face_of_def) lemma face_of_empty [simp]: "S face_of {} \ S = {}" by (meson empty_face_of face_of_def subset_empty) lemma face_of_trans [trans]: "\S face_of T; T face_of u\ \ S face_of u" unfolding face_of_def by (safe; blast) lemma face_of_face: "T face_of S \ (f face_of T \ f face_of S \ f \ T)" unfolding face_of_def by (safe; blast) lemma face_of_subset: "\F face_of S; F \ T; T \ S\ \ F face_of T" unfolding face_of_def by (safe; blast) lemma face_of_slice: "\F face_of S; convex T\ \ (F \ T) face_of (S \ T)" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Int: "\t1 face_of S; t2 face_of S\ \ (t1 \ t2) face_of S" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Inter: "\A \ {}; \T. T \ A \ T face_of S\ \ (\ A) face_of S" unfolding face_of_def by (blast intro: convex_Inter) lemma face_of_Int_Int: "\F face_of T; F' face_of t'\ \ (F \ F') face_of (T \ t')" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_imp_subset: "T face_of S \ T \ S" unfolding face_of_def by blast proposition face_of_imp_eq_affine_Int: fixes S :: "'a::euclidean_space set" assumes S: "convex S" and T: "T face_of S" shows "T = (affine hull T) \ S" proof - have "convex T" using T by (simp add: face_of_def) have *: False if x: "x \ affine hull T" and "x \ S" "x \ T" and y: "y \ rel_interior T" for x y proof - obtain e where "e>0" and e: "cball y e \ affine hull T \ T" using y by (auto simp: rel_interior_cball) have "y \ x" "y \ S" "y \ T" using face_of_imp_subset rel_interior_subset T that by blast+ then have zne: "\u. \u \ {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \ T\ \ False" using \x \ S\ \x \ T\ \T face_of S\ unfolding face_of_def by (meson greaterThanLessThan_iff in_segment(2)) have in01: "min (1/2) (e / norm (x - y)) \ {0<..<1}" using \y \ x\ \e > 0\ by simp have \
: "norm (min (1/2) (e / norm (x - y)) *\<^sub>R y - min (1/2) (e / norm (x - y)) *\<^sub>R x) \ e" using \e > 0\ by (simp add: scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right) show False apply (rule zne [OF in01 e [THEN subsetD]]) using \y \ T\ apply (simp add: hull_inc mem_affine x) by (simp add: dist_norm algebra_simps \
) qed show ?thesis proof (rule subset_antisym) show "T \ affine hull T \ S" using assms by (simp add: hull_subset face_of_imp_subset) show "affine hull T \ S \ T" using "*" \convex T\ rel_interior_eq_empty by fastforce qed qed lemma face_of_imp_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "T face_of S" shows "closed T" by (metis affine_affine_hull affine_closed closed_Int face_of_imp_eq_affine_Int assms) lemma face_of_Int_supporting_hyperplane_le_strong: assumes "convex(S \ {x. a \ x = b})" and aleb: "\x. x \ S \ a \ x \ b" shows "(S \ {x. a \ x = b}) face_of S" proof - have *: "a \ u = a \ x" if "x \ open_segment u v" "u \ S" "v \ S" and b: "b = a \ x" for u v x proof (rule antisym) show "a \ u \ a \ x" using aleb \u \ S\ \b = a \ x\ by blast next obtain \ where "b = a \ ((1 - \) *\<^sub>R u + \ *\<^sub>R v)" "0 < \" "\ < 1" using \b = a \ x\ \x \ open_segment u v\ in_segment by (auto simp: open_segment_image_interval split: if_split_asm) then have "b + \ * (a \ u) \ a \ u + \ * b" using aleb [OF \v \ S\] by (simp add: algebra_simps) then have "(1 - \) * b \ (1 - \) * (a \ u)" by (simp add: algebra_simps) then have "b \ a \ u" using \\ < 1\ by auto with b show "a \ x \ a \ u" by simp qed show ?thesis using "*" open_segment_commute by (fastforce simp add: face_of_def assms) qed lemma face_of_Int_supporting_hyperplane_ge_strong: "\convex(S \ {x. a \ x = b}); \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" using face_of_Int_supporting_hyperplane_le_strong [of S "-a" "-b"] by simp lemma face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_le_strong) lemma face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_ge_strong) lemma face_of_imp_convex: "T face_of S \ convex T" using face_of_def by blast lemma face_of_imp_compact: fixes S :: "'a::euclidean_space set" shows "\convex S; compact S; T face_of S\ \ compact T" by (meson bounded_subset compact_eq_bounded_closed face_of_imp_closed face_of_imp_subset) lemma face_of_Int_subface: "\A \ B face_of A; A \ B face_of B; C face_of A; D face_of B\ \ (C \ D) face_of C \ (C \ D) face_of D" by (meson face_of_Int_Int face_of_face inf_le1 inf_le2) lemma subset_of_face_of: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "u \ S" "T \ (rel_interior u) \ {}" shows "u \ T" proof fix c assume "c \ u" obtain b where "b \ T" "b \ rel_interior u" using assms by auto then obtain e where "e>0" "b \ u" and e: "cball b e \ affine hull u \ u" by (auto simp: rel_interior_cball) show "c \ T" proof (cases "b=c") case True with \b \ T\ show ?thesis by blast next case False define d where "d = b + (e / norm(b - c)) *\<^sub>R (b - c)" have "d \ cball b e \ affine hull u" using \e > 0\ \b \ u\ \c \ u\ by (simp add: d_def dist_norm hull_inc mem_affine_3_minus False) with e have "d \ u" by blast have nbc: "norm (b - c) + e > 0" using \e > 0\ by (metis add.commute le_less_trans less_add_same_cancel2 norm_ge_zero) then have [simp]: "d \ c" using False scaleR_cancel_left [of "1 + (e / norm (b - c))" b c] by (simp add: algebra_simps d_def) (simp add: field_split_simps) have [simp]: "((e - e * e / (e + norm (b - c))) / norm (b - c)) = (e / (e + norm (b - c)))" using False nbc by (simp add: divide_simps) (simp add: algebra_simps) have "b \ open_segment d c" apply (simp add: open_segment_image_interval) apply (simp add: d_def algebra_simps image_def) apply (rule_tac x="e / (e + norm (b - c))" in bexI) using False nbc \0 < e\ by (auto simp: algebra_simps) then have "d \ T \ c \ T" by (meson \b \ T\ \c \ u\ \d \ u\ assms face_ofD subset_iff) then show ?thesis .. qed qed lemma face_of_eq: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "U face_of S" "(rel_interior T) \ (rel_interior U) \ {}" shows "T = U" using assms unfolding disjoint_iff_not_equal by (metis IntI empty_iff face_of_imp_subset mem_rel_interior_ball subset_antisym subset_of_face_of) lemma face_of_disjoint_rel_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_interior S = {}" by (meson assms subset_of_face_of face_of_imp_subset order_refl subset_antisym) lemma face_of_disjoint_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ interior S = {}" proof - have "T \ interior S \ rel_interior S" by (meson inf_sup_ord(2) interior_subset_rel_interior order.trans) thus ?thesis by (metis (no_types) Int_greatest assms face_of_disjoint_rel_interior inf_sup_ord(1) subset_empty) qed lemma face_of_subset_rel_boundary: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ (S - rel_interior S)" by (meson DiffI assms disjoint_iff_not_equal face_of_disjoint_rel_interior face_of_imp_subset rev_subsetD subsetI) lemma face_of_subset_rel_frontier: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_frontier S" using assms closure_subset face_of_disjoint_rel_interior face_of_imp_subset rel_frontier_def by fastforce lemma face_of_aff_dim_lt: fixes S :: "'a::euclidean_space set" assumes "convex S" "T face_of S" "T \ S" shows "aff_dim T < aff_dim S" proof - have "aff_dim T \ aff_dim S" by (simp add: face_of_imp_subset aff_dim_subset assms) moreover have "aff_dim T \ aff_dim S" proof (cases "T = {}") case True then show ?thesis by (metis aff_dim_empty \T \ S\) next case False then show ?thesis by (metis Set.set_insert assms convex_rel_frontier_aff_dim dual_order.irrefl face_of_imp_convex face_of_subset_rel_frontier insert_not_empty subsetI) qed ultimately show ?thesis by simp qed lemma subset_of_face_of_affine_hull: fixes S :: "'a::euclidean_space set" assumes T: "T face_of S" and "convex S" "U \ S" and dis: "\ disjnt (affine hull T) (rel_interior U)" shows "U \ T" proof (rule subset_of_face_of [OF T \U \ S\]) show "T \ rel_interior U \ {}" using face_of_imp_eq_affine_Int [OF \convex S\ T] rel_interior_subset [of U] dis \U \ S\ disjnt_def by fastforce qed lemma affine_hull_face_of_disjoint_rel_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" "F face_of S" "F \ S" shows "affine hull F \ rel_interior S = {}" by (metis assms disjnt_def face_of_imp_subset order_refl subset_antisym subset_of_face_of_affine_hull) lemma affine_diff_divide: assumes "affine S" "k \ 0" "k \ 1" and xy: "x \ S" "y /\<^sub>R (1 - k) \ S" shows "(x - y) /\<^sub>R k \ S" proof - have "inverse(k) *\<^sub>R (x - y) = (1 - inverse k) *\<^sub>R inverse(1 - k) *\<^sub>R y + inverse(k) *\<^sub>R x" using assms by (simp add: algebra_simps) (simp add: scaleR_left_distrib [symmetric] field_split_simps) then show ?thesis using \affine S\ xy by (auto simp: affine_alt) qed proposition face_of_convex_hulls: assumes S: "finite S" "T \ S" and disj: "affine hull T \ convex hull (S - T) = {}" shows "(convex hull T) face_of (convex hull S)" proof - have fin: "finite T" "finite (S - T)" using assms by (auto simp: finite_subset) have *: "x \ convex hull T" if x: "x \ convex hull S" and y: "y \ convex hull S" and w: "w \ convex hull T" "w \ open_segment x y" for x y w proof - have waff: "w \ affine hull T" using convex_hull_subset_affine_hull w by blast obtain a b where a: "\i. i \ S \ 0 \ a i" and asum: "sum a S = 1" and aeqx: "(\i\S. a i *\<^sub>R i) = x" and b: "\i. i \ S \ 0 \ b i" and bsum: "sum b S = 1" and beqy: "(\i\S. b i *\<^sub>R i) = y" using x y by (auto simp: assms convex_hull_finite) obtain u where "(1 - u) *\<^sub>R x + u *\<^sub>R y \ convex hull T" "x \ y" and weq: "w = (1 - u) *\<^sub>R x + u *\<^sub>R y" and u01: "0 < u" "u < 1" using w by (auto simp: open_segment_image_interval split: if_split_asm) define c where "c i = (1 - u) * a i + u * b i" for i have cge0: "\i. i \ S \ 0 \ c i" using a b u01 by (simp add: c_def) have sumc1: "sum c S = 1" by (simp add: c_def sum.distrib sum_distrib_left [symmetric] asum bsum) have sumci_xy: "(\i\S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y" apply (simp add: c_def sum.distrib scaleR_left_distrib) by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] aeqx beqy) show ?thesis proof (cases "sum c (S - T) = 0") case True have ci0: "\i. i \ (S - T) \ c i = 0" using True cge0 fin(2) sum_nonneg_eq_0_iff by auto have a0: "a i = 0" if "i \ (S - T)" for i using ci0 [OF that] u01 a [of i] b [of i] that by (simp add: c_def Groups.ordered_comm_monoid_add_class.add_nonneg_eq_0_iff) have [simp]: "sum a T = 1" using assms by (metis sum.mono_neutral_cong_right a0 asum) show ?thesis apply (simp add: convex_hull_finite \finite T\) apply (rule_tac x=a in exI) using a0 assms apply (auto simp: cge0 a aeqx [symmetric] sum.mono_neutral_right) done next case False define k where "k = sum c (S - T)" have "k > 0" using False unfolding k_def by (metis DiffD1 antisym_conv cge0 sum_nonneg not_less) have weq_sumsum: "w = sum (\x. c x *\<^sub>R x) T + sum (\x. c x *\<^sub>R x) (S - T)" by (metis (no_types) add.commute S(1) S(2) sum.subset_diff sumci_xy weq) show ?thesis proof (cases "k = 1") case True then have "sum c T = 0" by (simp add: S k_def sum_diff sumc1) then have [simp]: "sum c (S - T) = 1" by (simp add: S sum_diff sumc1) have ci0: "\i. i \ T \ c i = 0" by (meson \finite T\ \sum c T = 0\ \T \ S\ cge0 sum_nonneg_eq_0_iff subsetCE) then have [simp]: "(\i\S-T. c i *\<^sub>R i) = w" by (simp add: weq_sumsum) have "w \ convex hull (S - T)" apply (simp add: convex_hull_finite fin) apply (rule_tac x=c in exI) apply (auto simp: cge0 weq True k_def) done then show ?thesis using disj waff by blast next case False then have sumcf: "sum c T = 1 - k" by (simp add: S k_def sum_diff sumc1) have ge0: "\x. x \ T \ 0 \ inverse (1 - k) * c x" by (metis \T \ S\ cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg subsetD sum_nonneg sumcf) have eq1: "(\x\T. inverse (1 - k) * c x) = 1" by (metis False eq_iff_diff_eq_0 mult.commute right_inverse sum_distrib_left sumcf) have "(\i\T. c i *\<^sub>R i) /\<^sub>R (1 - k) \ convex hull T" apply (simp add: convex_hull_finite fin) apply (rule_tac x="\i. inverse (1-k) * c i" in exI) by (metis (mono_tags, lifting) eq1 ge0 scaleR_scaleR scale_sum_right sum.cong) with \0 < k\ have "inverse(k) *\<^sub>R (w - sum (\i. c i *\<^sub>R i) T) \ affine hull T" by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD]) moreover have "inverse(k) *\<^sub>R (w - sum (\x. c x *\<^sub>R x) T) \ convex hull (S - T)" apply (simp add: weq_sumsum convex_hull_finite fin) apply (rule_tac x="\i. inverse k * c i" in exI) using \k > 0\ cge0 apply (auto simp: scaleR_right.sum sum_distrib_left [symmetric] k_def [symmetric]) done ultimately show ?thesis using disj by blast qed qed qed have [simp]: "convex hull T \ convex hull S" by (simp add: \T \ S\ hull_mono) show ?thesis using open_segment_commute by (auto simp: face_of_def intro: *) qed proposition face_of_convex_hull_insert: assumes "finite S" "a \ affine hull S" and T: "T face_of convex hull S" shows "T face_of convex hull insert a S" proof - have "convex hull S face_of convex hull insert a S" by (simp add: assms face_of_convex_hulls insert_Diff_if subset_insertI) then show ?thesis using T face_of_trans by blast qed proposition face_of_affine_trivial: assumes "affine S" "T face_of S" shows "T = {} \ T = S" proof (rule ccontr, clarsimp) assume "T \ {}" "T \ S" then obtain a where "a \ T" by auto then have "a \ S" using \T face_of S\ face_of_imp_subset by blast have "S \ T" proof fix b assume "b \ S" show "b \ T" proof (cases "a = b") case True with \a \ T\ show ?thesis by auto next case False then have "a \ 2 *\<^sub>R a - b" by (simp add: scaleR_2) with False have "a \ open_segment (2 *\<^sub>R a - b) b" apply (clarsimp simp: open_segment_def closed_segment_def) apply (rule_tac x="1/2" in exI) by (simp add: algebra_simps) moreover have "2 *\<^sub>R a - b \ S" by (rule mem_affine [OF \affine S\ \a \ S\ \b \ S\, of 2 "-1", simplified]) moreover note \b \ S\ \a \ T\ ultimately show ?thesis by (rule face_ofD [OF \T face_of S\, THEN conjunct2]) qed qed then show False using \T \ S\ \T face_of S\ face_of_imp_subset by blast qed lemma face_of_affine_eq: "affine S \ (T face_of S \ T = {} \ T = S)" using affine_imp_convex face_of_affine_trivial face_of_refl by auto proposition Inter_faces_finite_altbound: fixes T :: "'a::euclidean_space set set" assumes cfaI: "\c. c \ T \ c face_of S" shows "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ \F' = \T" proof (cases "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ (\c. c \ T \ c \ (\F') \ (\F'))") case True then obtain c where c: "\F'. \finite F'; F' \ T; card F' \ DIM('a) + 2\ \ c F' \ T \ c F' \ (\F') \ (\F')" by metis define d where "d = rec_nat {c{}} (\n r. insert (c r) r)" have [simp]: "d 0 = {c {}}" by (simp add: d_def) have dSuc [simp]: "\n. d (Suc n) = insert (c (d n)) (d n)" by (simp add: d_def) have dn_notempty: "d n \ {}" for n by (induction n) auto have dn_le_Suc: "d n \ T \ finite(d n) \ card(d n) \ Suc n" if "n \ DIM('a) + 2" for n using that proof (induction n) case 0 then show ?case by (simp add: c) next case (Suc n) then show ?case by (auto simp: c card_insert_if) qed have aff_dim_le: "aff_dim(\(d n)) \ DIM('a) - int n" if "n \ DIM('a) + 2" for n using that proof (induction n) case 0 then show ?case by (simp add: aff_dim_le_DIM) next case (Suc n) have fs: "\(d (Suc n)) face_of S" by (meson Suc.prems cfaI dn_le_Suc dn_notempty face_of_Inter subsetCE) have condn: "convex (\(d n))" using Suc.prems nat_le_linear not_less_eq_eq by (blast intro: face_of_imp_convex cfaI convex_Inter dest: dn_le_Suc) have fdn: "\(d (Suc n)) face_of \(d n)" by (metis (no_types, lifting) Inter_anti_mono Suc.prems dSuc cfaI dn_le_Suc dn_notempty face_of_Inter face_of_imp_subset face_of_subset subset_iff subset_insertI) have ne: "\(d (Suc n)) \ \(d n)" by (metis (no_types, lifting) Suc.prems Suc_leD c complete_lattice_class.Inf_insert dSuc dn_le_Suc less_irrefl order.trans) have *: "\m::int. \d. \d'::int. d < d' \ d' \ m - n \ d \ m - of_nat(n+1)" by arith have "aff_dim (\(d (Suc n))) < aff_dim (\(d n))" by (rule face_of_aff_dim_lt [OF condn fdn ne]) moreover have "aff_dim (\(d n)) \ int (DIM('a)) - int n" using Suc by auto ultimately have "aff_dim (\(d (Suc n))) \ int (DIM('a)) - (n+1)" by arith then show ?case by linarith qed have "aff_dim (\(d (DIM('a) + 2))) \ -2" using aff_dim_le [OF order_refl] by simp with aff_dim_geq [of "\(d (DIM('a) + 2))"] show ?thesis using order.trans by fastforce next case False then show ?thesis apply simp apply (erule ex_forward) by blast qed lemma faces_of_translation: "{F. F face_of image (\x. a + x) S} = image (image (\x. a + x)) {F. F face_of S}" proof - have "\F. F face_of (+) a ` S \ \G. G face_of S \ F = (+) a ` G" by (metis face_of_imp_subset face_of_translation_eq subset_imageE) then show ?thesis by (auto simp: image_iff) qed proposition face_of_Times: assumes "F face_of S" and "F' face_of S'" shows "(F \ F') face_of (S \ S')" proof - have "F \ F' \ S \ S'" using assms [unfolded face_of_def] by blast moreover have "convex (F \ F')" using assms [unfolded face_of_def] by (blast intro: convex_Times) moreover have "a \ F \ a' \ F' \ b \ F \ b' \ F'" if "a \ S" "b \ S" "a' \ S'" "b' \ S'" "x \ F \ F'" "x \ open_segment (a,a') (b,b')" for a b a' b' x proof (cases "b=a \ b'=a'") case True with that show ?thesis using assms by (force simp: in_segment dest: face_ofD) next case False with assms [unfolded face_of_def] that show ?thesis by (blast dest!: open_segment_PairD) qed ultimately show ?thesis unfolding face_of_def by blast qed corollary face_of_Times_decomp: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "C face_of (S \ S') \ (\F F'. F face_of S \ F' face_of S' \ C = F \ F')" (is "?lhs = ?rhs") proof assume C: ?lhs show ?rhs proof (cases "C = {}") case True then show ?thesis by auto next case False have 1: "fst ` C \ S" "snd ` C \ S'" using C face_of_imp_subset by fastforce+ have "convex C" using C by (metis face_of_imp_convex) have conv: "convex (fst ` C)" "convex (snd ` C)" by (simp_all add: \convex C\ convex_linear_image linear_fst linear_snd) have fstab: "a \ fst ` C \ b \ fst ` C" if "a \ S" "b \ S" "x \ open_segment a b" "(x,x') \ C" for a b x x' proof - have *: "(x,x') \ open_segment (a,x') (b,x')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF C *] that face_of_imp_subset [OF C] by force qed have fst: "fst ` C face_of S" by (force simp: face_of_def 1 conv fstab) have sndab: "a' \ snd ` C \ b' \ snd ` C" if "a' \ S'" "b' \ S'" "x' \ open_segment a' b'" "(x,x') \ C" for a' b' x x' proof - have *: "(x,x') \ open_segment (x,a') (x,b')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF C *] that face_of_imp_subset [OF C] by force qed have snd: "snd ` C face_of S'" by (force simp: face_of_def 1 conv sndab) have cc: "rel_interior C \ rel_interior (fst ` C) \ rel_interior (snd ` C)" by (force simp: face_of_Times rel_interior_Times conv fst snd \convex C\ linear_fst linear_snd rel_interior_convex_linear_image [symmetric]) have "C = fst ` C \ snd ` C" proof (rule face_of_eq [OF C]) show "fst ` C \ snd ` C face_of S \ S'" by (simp add: face_of_Times rel_interior_Times conv fst snd) show "rel_interior C \ rel_interior (fst ` C \ snd ` C) \ {}" using False rel_interior_eq_empty \convex C\ cc by (auto simp: face_of_Times rel_interior_Times conv fst) qed with fst snd show ?thesis by metis qed next assume ?rhs with face_of_Times show ?lhs by auto qed lemma face_of_Times_eq: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "(F \ F') face_of (S \ S') \ F = {} \ F' = {} \ F face_of S \ F' face_of S'" by (auto simp: face_of_Times_decomp times_eq_iff) lemma hyperplane_face_of_halfspace_le: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_le [OF convex_halfspace_le [of a b], of a b] show ?thesis by auto qed lemma hyperplane_face_of_halfspace_ge: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_ge [OF convex_halfspace_ge [of b a], of b a] show ?thesis by auto qed lemma face_of_halfspace_le: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis using face_of_affine_eq affine_UNIV by auto next case False then have ine: "interior {x. a \ x \ b} \ {}" using halfspace_eq_empty_lt interior_halfspace_le by blast show ?thesis proof assume L: ?lhs have "F face_of {x. a \ x = b}" if "F \ {x. a \ x \ b}" proof - have "F face_of rel_frontier {x. a \ x \ b}" proof (rule face_of_subset [OF L]) show "F \ rel_frontier {x. a \ x \ b}" by (simp add: L face_of_subset_rel_frontier that) qed (force simp: rel_frontier_def closed_halfspace_le) then show ?thesis using False by (simp add: frontier_halfspace_le rel_frontier_nonempty_interior [OF ine]) qed with L show ?rhs using affine_hyperplane face_of_affine_eq by blast next assume ?rhs then show ?lhs by (metis convex_halfspace_le empty_face_of face_of_refl hyperplane_face_of_halfspace_le) qed qed lemma face_of_halfspace_ge: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" using face_of_halfspace_le [of F "-a" "-b"] by simp subsection\Exposed faces\ text\That is, faces that are intersection with supporting hyperplane\ definition\<^marker>\tag important\ exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(exposed'_face'_of)" 50) where "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b})" lemma empty_exposed_face_of [iff]: "{} exposed_face_of S" apply (simp add: exposed_face_of_def) apply (rule_tac x=0 in exI) apply (rule_tac x=1 in exI, force) done lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \ convex S" proof assume S: "convex S" have "S \ {x. 0 \ x \ 0} \ S = S \ {x. 0 \ x = 0}" by auto with S show "S exposed_face_of S" using exposed_face_of_def face_of_refl_eq by blast qed (simp add: exposed_face_of_def face_of_refl_eq) lemma exposed_face_of_refl: "convex S \ S exposed_face_of S" by simp lemma exposed_face_of: "T exposed_face_of S \ T face_of S \ (T = {} \ T = S \ (\a b. a \ 0 \ S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b}))" proof (cases "T = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "T = S") case True then show ?thesis by (simp add: face_of_refl_eq) next case False with \T \ {}\ show ?thesis apply (auto simp: exposed_face_of_def) apply (metis inner_zero_left) done qed qed lemma exposed_face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" by (force simp: exposed_face_of_def face_of_Int_supporting_hyperplane_le) lemma exposed_face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" using exposed_face_of_Int_supporting_hyperplane_le [of S "-a" "-b"] by simp proposition exposed_face_of_Int: assumes "T exposed_face_of S" and "u exposed_face_of S" shows "(T \ u) exposed_face_of S" proof - obtain a b where T: "S \ {x. a \ x = b} face_of S" and S: "S \ {x. a \ x \ b}" and teq: "T = S \ {x. a \ x = b}" using assms by (auto simp: exposed_face_of_def) obtain a' b' where u: "S \ {x. a' \ x = b'} face_of S" and s': "S \ {x. a' \ x \ b'}" and ueq: "u = S \ {x. a' \ x = b'}" using assms by (auto simp: exposed_face_of_def) have tu: "T \ u face_of S" using T teq u ueq by (simp add: face_of_Int) have ss: "S \ {x. (a + a') \ x \ b + b'}" using S s' by (force simp: inner_left_distrib) show ?thesis apply (simp add: exposed_face_of_def tu) apply (rule_tac x="a+a'" in exI) apply (rule_tac x="b+b'" in exI) using S s' apply (fastforce simp: ss inner_left_distrib teq ueq) done qed proposition exposed_face_of_Inter: fixes P :: "'a::euclidean_space set set" assumes "P \ {}" and "\T. T \ P \ T exposed_face_of S" shows "\P exposed_face_of S" proof - obtain Q where "finite Q" and QsubP: "Q \ P" "card Q \ DIM('a) + 2" and IntQ: "\Q = \P" using Inter_faces_finite_altbound [of P S] assms [unfolded exposed_face_of] by force show ?thesis proof (cases "Q = {}") case True then show ?thesis by (metis IntQ Inter_UNIV_conv(2) assms(1) assms(2) ex_in_conv) next case False have "Q \ {T. T exposed_face_of S}" using QsubP assms by blast moreover have "Q \ {T. T exposed_face_of S} \ \Q exposed_face_of S" using \finite Q\ False by (induction Q rule: finite_induct; use exposed_face_of_Int in fastforce) ultimately show ?thesis by (simp add: IntQ) qed qed proposition exposed_face_of_sums: assumes "convex S" and "convex T" and "F exposed_face_of {x + y | x y. x \ S \ y \ T}" (is "F exposed_face_of ?ST") obtains k l where "k exposed_face_of S" "l exposed_face_of T" "F = {x + y | x y. x \ k \ y \ l}" proof (cases "F = {}") case True then show ?thesis using that by blast next case False show ?thesis proof (cases "F = ?ST") case True then show ?thesis using assms exposed_face_of_refl_eq that by blast next case False obtain p where "p \ F" using \F \ {}\ by blast moreover obtain u z where T: "?ST \ {x. u \ x = z} face_of ?ST" and S: "?ST \ {x. u \ x \ z}" and feq: "F = ?ST \ {x. u \ x = z}" using assms by (auto simp: exposed_face_of_def) ultimately obtain a0 b0 where p: "p = a0 + b0" and "a0 \ S" "b0 \ T" and z: "u \ p = z" by auto have lez: "u \ (x + y) \ z" if "x \ S" "y \ T" for x y using S that by auto have sef: "S \ {x. u \ x = u \ a0} exposed_face_of S" proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex S\]) show "\x. x \ S \ u \ x \ u \ a0" by (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \b0 \ T\]) qed have tef: "T \ {x. u \ x = u \ b0} exposed_face_of T" proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex T\]) show "\x. x \ T \ u \ x \ u \ b0" by (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \a0 \ S\]) qed have "{x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0} \ F" by (auto simp: feq) (metis inner_right_distrib p z) moreover have "F \ {x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0}" proof - have "\x y. \z = u \ (x + y); x \ S; y \ T\ \ u \ x = u \ a0 \ u \ y = u \ b0" using z p \a0 \ S\ \b0 \ T\ apply (simp add: inner_right_distrib) apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute) done then show ?thesis using feq by blast qed ultimately have "F = {x + y |x y. x \ S \ {x. u \ x = u \ a0} \ y \ T \ {x. u \ x = u \ b0}}" by blast then show ?thesis by (rule that [OF sef tef]) qed qed proposition exposed_face_of_parallel: "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b} \ (T \ {} \ T \ S \ a \ 0) \ (T \ S \ (\w \ affine hull S. (w + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (clarsimp simp: exposed_face_of_def) fix a b assume faceS: "S \ {x. a \ x = b} face_of S" and Ssub: "S \ {x. a \ x \ b}" show "\c d. S \ {x. c \ x \ d} \ S \ {x. a \ x = b} = S \ {x. c \ x = d} \ (S \ {x. a \ x = b} \ {} \ S \ {x. a \ x = b} \ S \ c \ 0) \ (S \ {x. a \ x = b} \ S \ (\w \ affine hull S. w + c \ affine hull S))" proof (cases "affine hull S \ {x. -a \ x \ -b} = {} \ affine hull S \ {x. - a \ x \ - b}") case True then show ?thesis proof assume "affine hull S \ {x. - a \ x \ - b} = {}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="1" in exI) using hull_subset by fastforce next assume "affine hull S \ {x. - a \ x \ - b}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="0" in exI) using Ssub hull_subset by fastforce qed next case False then obtain a' b' where "a' \ 0" and le: "affine hull S \ {x. a' \ x \ b'} = affine hull S \ {x. - a \ x \ - b}" and eq: "affine hull S \ {x. a' \ x = b'} = affine hull S \ {x. - a \ x = - b}" and mem: "\w. w \ affine hull S \ w + a' \ affine hull S" using affine_parallel_slice affine_affine_hull by metis show ?thesis proof (intro conjI impI allI ballI exI) have *: "S \ - (affine hull S \ {x. P x}) \ affine hull S \ {x. Q x} \ S \ {x. \ P x \ Q x}" for P Q using hull_subset by fastforce have "S \ {x. \ (a' \ x \ b') \ a' \ x = b'}" by (rule *) (use le eq Ssub in auto) then show "S \ {x. - a' \ x \ - b'}" by auto show "S \ {x. a \ x = b} = S \ {x. - a' \ x = - b'}" using eq hull_subset [of S affine] by force show "\S \ {x. a \ x = b} \ {}; S \ {x. a \ x = b} \ S\ \ - a' \ 0" using \a' \ 0\ by auto show "w + - a' \ affine hull S" if "S \ {x. a \ x = b} \ S" "w \ affine hull S" for w proof - have "w + 1 *\<^sub>R (w - (w + a')) \ affine hull S" using affine_affine_hull mem mem_affine_3_minus that(2) by blast then show ?thesis by simp qed qed qed qed next assume ?rhs then show ?lhs unfolding exposed_face_of_def by blast qed subsection\Extreme points of a set: its singleton faces\ definition\<^marker>\tag important\ extreme_point_of :: "['a::real_vector, 'a set] \ bool" (infixr "(extreme'_point'_of)" 50) where "x extreme_point_of S \ x \ S \ (\a \ S. \b \ S. x \ open_segment a b)" lemma extreme_point_of_stillconvex: "convex S \ (x extreme_point_of S \ x \ S \ convex(S - {x}))" by (fastforce simp add: convex_contains_segment extreme_point_of_def open_segment_def) lemma face_of_singleton: "{x} face_of S \ x extreme_point_of S" by (fastforce simp add: extreme_point_of_def face_of_def) lemma extreme_point_not_in_REL_INTERIOR: fixes S :: "'a::real_normed_vector set" shows "\x extreme_point_of S; S \ {x}\ \ x \ rel_interior S" by (metis disjoint_iff face_of_disjoint_rel_interior face_of_singleton insertI1) lemma extreme_point_not_in_interior: fixes S :: "'a::{real_normed_vector, perfect_space} set" assumes "x extreme_point_of S" shows "x \ interior S" proof (cases "S = {x}") case False then show ?thesis by (meson assms subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior) qed (simp add: empty_interior_finite) lemma extreme_point_of_face: "F face_of S \ v extreme_point_of F \ v extreme_point_of S \ v \ F" by (meson empty_subsetI face_of_face face_of_singleton insert_subset) lemma extreme_point_of_convex_hull: "x extreme_point_of (convex hull S) \ x \ S" using hull_minimal [of S "(convex hull S) - {x}" convex] using hull_subset [of S convex] by (force simp add: extreme_point_of_stillconvex) proposition extreme_points_of_convex_hull: "{x. x extreme_point_of (convex hull S)} \ S" using extreme_point_of_convex_hull by auto lemma extreme_point_of_empty [simp]: "\ (x extreme_point_of {})" by (simp add: extreme_point_of_def) lemma extreme_point_of_singleton [iff]: "x extreme_point_of {a} \ x = a" using extreme_point_of_stillconvex by auto lemma extreme_point_of_translation_eq: "(a + x) extreme_point_of (image (\x. a + x) S) \ x extreme_point_of S" by (auto simp: extreme_point_of_def) lemma extreme_points_of_translation: "{x. x extreme_point_of (image (\x. a + x) S)} = (\x. a + x) ` {x. x extreme_point_of S}" using extreme_point_of_translation_eq by auto (metis (no_types, lifting) image_iff mem_Collect_eq minus_add_cancel) lemma extreme_points_of_translation_subtract: "{x. x extreme_point_of (image (\x. x - a) S)} = (\x. x - a) ` {x. x extreme_point_of S}" using extreme_points_of_translation [of "- a" S] by simp lemma extreme_point_of_Int: "\x extreme_point_of S; x extreme_point_of T\ \ x extreme_point_of (S \ T)" by (simp add: extreme_point_of_def) lemma extreme_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" by (metis convex_singleton face_of_Int_supporting_hyperplane_le_strong face_of_singleton) lemma extreme_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" using extreme_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] by simp lemma exposed_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" unfolding exposed_face_of_def by (force simp: face_of_singleton extreme_point_of_Int_supporting_hyperplane_le) lemma exposed_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] by simp lemma extreme_point_of_convex_hull_insert: assumes "finite S" "a \ convex hull S" shows "a extreme_point_of (convex hull (insert a S))" proof (cases "a \ S") case False then show ?thesis using face_of_convex_hulls [of "insert a S" "{a}"] assms by (auto simp: face_of_singleton hull_same) qed (use assms in \simp add: hull_inc\) subsection\Facets\ definition\<^marker>\tag important\ facet_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(facet'_of)" 50) where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" lemma facet_of_empty [simp]: "\ S facet_of {}" by (simp add: facet_of_def) lemma facet_of_irrefl [simp]: "\ S facet_of S " by (simp add: facet_of_def) lemma facet_of_imp_face_of: "F facet_of S \ F face_of S" by (simp add: facet_of_def) lemma facet_of_imp_subset: "F facet_of S \ F \ S" by (simp add: face_of_imp_subset facet_of_def) lemma hyperplane_facet_of_halfspace_le: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_ge hyperplane_face_of_halfspace_le Suc_leI of_nat_diff aff_dim_halfspace_le) lemma hyperplane_facet_of_halfspace_ge: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_le hyperplane_face_of_halfspace_ge Suc_leI of_nat_diff aff_dim_halfspace_ge) lemma facet_of_halfspace_le: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" (is "?lhs = ?rhs") proof assume c: ?lhs with c facet_of_irrefl show ?rhs by (force simp: aff_dim_halfspace_le facet_of_def face_of_halfspace_le cong: conj_cong split: if_split_asm) next assume ?rhs then show ?lhs by (simp add: hyperplane_facet_of_halfspace_le) qed lemma facet_of_halfspace_ge: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" using facet_of_halfspace_le [of F "-a" "-b"] by simp subsection \Edges: faces of affine dimension 1\ (*FIXME too small subsection, rearrange? *) definition\<^marker>\tag important\ edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) where "e edge_of S \ e face_of S \ aff_dim e = 1" lemma edge_of_imp_subset: "S edge_of T \ S \ T" by (simp add: edge_of_def face_of_imp_subset) subsection\Existence of extreme points\ proposition different_norm_3_collinear_points: fixes a :: "'a::euclidean_space" assumes "x \ open_segment a b" "norm(a) = norm(b)" "norm(x) = norm(b)" shows False proof - obtain u where "norm ((1 - u) *\<^sub>R a + u *\<^sub>R b) = norm b" and "a \ b" and u01: "0 < u" "u < 1" using assms by (auto simp: open_segment_image_interval if_splits) then have "(1 - u) *\<^sub>R a \ (1 - u) *\<^sub>R a + ((1 - u) * 2) *\<^sub>R a \ u *\<^sub>R b = (1 - u * u) *\<^sub>R (a \ a)" using assms by (simp add: norm_eq algebra_simps inner_commute) then have "(1 - u) *\<^sub>R ((1 - u) *\<^sub>R a \ a + (2 * u) *\<^sub>R a \ b) = (1 - u) *\<^sub>R ((1 + u) *\<^sub>R (a \ a))" by (simp add: algebra_simps) then have "(1 - u) *\<^sub>R (a \ a) + (2 * u) *\<^sub>R (a \ b) = (1 + u) *\<^sub>R (a \ a)" using u01 by auto then have "a \ b = a \ a" using u01 by (simp add: algebra_simps) then have "a = b" using \norm(a) = norm(b)\ norm_eq vector_eq by fastforce then show ?thesis using \a \ b\ by force qed proposition extreme_point_exists_convex: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" "S \ {}" obtains x where "x extreme_point_of S" proof - obtain x where "x \ S" and xsup: "\y. y \ S \ norm y \ norm x" using distance_attains_sup [of S 0] assms by auto have False if "a \ S" "b \ S" and x: "x \ open_segment a b" for a b proof - have noax: "norm a \ norm x" and nobx: "norm b \ norm x" using xsup that by auto have "a \ b" using empty_iff open_segment_idem x by auto show False by (metis dist_0_norm dist_decreases_open_segment noax nobx not_le x) qed then show ?thesis by (meson \x \ S\ extreme_point_of_def that) qed subsection\Krein-Milman, the weaker form\ proposition Krein_Milman: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = closure(convex hull {x. x extreme_point_of S})" proof (cases "S = {}") case True then show ?thesis by simp next case False have "closed S" by (simp add: \compact S\ compact_imp_closed) have "closure (convex hull {x. x extreme_point_of S}) \ S" by (simp add: \closed S\ assms closure_minimal extreme_point_of_def hull_minimal) moreover have "u \ closure (convex hull {x. x extreme_point_of S})" if "u \ S" for u proof (rule ccontr) assume unot: "u \ closure(convex hull {x. x extreme_point_of S})" then obtain a b where "a \ u < b" and ab: "\x. x \ closure(convex hull {x. x extreme_point_of S}) \ b < a \ x" using separating_hyperplane_closed_point [of "closure(convex hull {x. x extreme_point_of S})"] by blast have "continuous_on S ((\) a)" by (rule continuous_intros)+ then obtain m where "m \ S" and m: "\y. y \ S \ a \ m \ a \ y" using continuous_attains_inf [of S "\x. a \ x"] \compact S\ \u \ S\ by auto define T where "T = S \ {x. a \ x = a \ m}" have "m \ T" by (simp add: T_def \m \ S\) moreover have "compact T" by (simp add: T_def compact_Int_closed [OF \compact S\ closed_hyperplane]) moreover have "convex T" by (simp add: T_def convex_Int [OF \convex S\ convex_hyperplane]) ultimately obtain v where v: "v extreme_point_of T" using extreme_point_exists_convex [of T] by auto then have "{v} face_of T" by (simp add: face_of_singleton) also have "T face_of S" by (simp add: T_def m face_of_Int_supporting_hyperplane_ge [OF \convex S\]) finally have "v extreme_point_of S" by (simp add: face_of_singleton) then have "b < a \ v" using closure_subset by (simp add: closure_hull hull_inc ab) then show False using \a \ u < b\ \{v} face_of T\ face_of_imp_subset m T_def that by fastforce qed ultimately show ?thesis by blast qed text\Now the sharper form.\ lemma Krein_Milman_Minkowski_aux: fixes S :: "'a::euclidean_space set" assumes n: "dim S = n" and S: "compact S" "convex S" "0 \ S" shows "0 \ convex hull {x. x extreme_point_of S}" using n S proof (induction n arbitrary: S rule: less_induct) case (less n S) show ?case proof (cases "0 \ rel_interior S") case True with Krein_Milman less.prems show ?thesis by (metis subsetD convex_convex_hull convex_rel_interior_closure rel_interior_subset) next case False have "rel_interior S \ {}" by (simp add: rel_interior_convex_nonempty_aux less) then obtain c where c: "c \ rel_interior S" by blast obtain a where "a \ 0" and le_ay: "\y. y \ S \ a \ 0 \ a \ y" and less_ay: "\y. y \ rel_interior S \ a \ 0 < a \ y" by (blast intro: supporting_hyperplane_rel_boundary intro!: less False) have face: "S \ {x. a \ x = 0} face_of S" using face_of_Int_supporting_hyperplane_ge le_ay \convex S\ by auto then have co: "compact (S \ {x. a \ x = 0})" "convex (S \ {x. a \ x = 0})" using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+ have "a \ y = 0" if "y \ span (S \ {x. a \ x = 0})" for y proof - have "y \ span {x. a \ x = 0}" by (metis inf.cobounded2 span_mono subsetCE that) then show ?thesis by (blast intro: span_induct [OF _ subspace_hyperplane]) qed then have "dim (S \ {x. a \ x = 0}) < n" by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff inf_le1 \dim S = n\ not_le rel_interior_subset span_0 span_base) then have "0 \ convex hull {x. x extreme_point_of (S \ {x. a \ x = 0})}" by (rule less.IH) (auto simp: co less.prems) then show ?thesis by (metis (mono_tags, lifting) Collect_mono_iff face extreme_point_of_face hull_mono subset_iff) qed qed theorem Krein_Milman_Minkowski: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = convex hull {x. x extreme_point_of S}" proof show "S \ convex hull {x. x extreme_point_of S}" proof fix a assume [simp]: "a \ S" have 1: "compact ((+) (- a) ` S)" by (simp add: \compact S\ compact_translation_subtract cong: image_cong_simp) have 2: "convex ((+) (- a) ` S)" by (simp add: \convex S\ compact_translation_subtract) show a_invex: "a \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski_aux [OF refl 1 2] convex_hull_translation [of "-a"] by (auto simp: extreme_points_of_translation_subtract translation_assoc cong: image_cong_simp) qed next show "convex hull {x. x extreme_point_of S} \ S" proof - have "{a. a extreme_point_of S} \ S" using extreme_point_of_def by blast then show ?thesis by (simp add: \convex S\ hull_minimal) qed qed subsection\Applying it to convex hulls of explicitly indicated finite sets\ corollary Krein_Milman_polytope: fixes S :: "'a::euclidean_space set" shows "finite S \ convex hull S = convex hull {x. x extreme_point_of (convex hull S)}" by (simp add: Krein_Milman_Minkowski finite_imp_compact_convex_hull) lemma extreme_points_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ {x. x extreme_point_of (convex hull S)} = S" by (metis (full_types) Krein_Milman_Minkowski compact_convex_hull convex_convex_hull extreme_points_of_convex_hull psubsetI) lemma extreme_point_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ (x extreme_point_of (convex hull S) \ x \ S)" using extreme_points_of_convex_hull_eq by auto lemma extreme_point_of_convex_hull_convex_independent: fixes S :: "'a::euclidean_space set" assumes "compact S" and S: "\a. a \ S \ a \ convex hull (S - {a})" shows "(x extreme_point_of (convex hull S) \ x \ S)" proof - have "convex hull T \ convex hull S" if "T \ S" for T proof - obtain a where "T \ S" "a \ S" "a \ T" using \T \ S\ by blast then show ?thesis by (metis (full_types) Diff_eq_empty_iff Diff_insert0 S hull_mono hull_subset insert_Diff_single subsetCE) qed then show ?thesis by (rule extreme_point_of_convex_hull_eq [OF \compact S\]) qed lemma extreme_point_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" shows "\ affine_dependent S \ (x extreme_point_of (convex hull S) \ x \ S)" by (metis aff_independent_finite affine_dependent_def affine_hull_convex_hull extreme_point_of_convex_hull_convex_independent finite_imp_compact hull_inc) text\Elementary proofs exist, not requiring Euclidean spaces and all this development\ lemma extreme_point_of_convex_hull_2: fixes x :: "'a::euclidean_space" shows "x extreme_point_of (convex hull {a,b}) \ x = a \ x = b" proof - have "x extreme_point_of (convex hull {a,b}) \ x \ {a,b}" by (intro extreme_point_of_convex_hull_affine_independent affine_independent_2) then show ?thesis by simp qed lemma extreme_point_of_segment: fixes x :: "'a::euclidean_space" shows "x extreme_point_of closed_segment a b \ x = a \ x = b" by (simp add: extreme_point_of_convex_hull_2 segment_convex_hull) lemma face_of_convex_hull_subset: fixes S :: "'a::euclidean_space set" assumes "compact S" and T: "T face_of (convex hull S)" obtains s' where "s' \ S" "T = convex hull s'" proof show "{x. x extreme_point_of T} \ S" using T extreme_point_of_convex_hull extreme_point_of_face by blast show "T = convex hull {x. x extreme_point_of T}" proof (rule Krein_Milman_Minkowski) show "compact T" using T assms compact_convex_hull face_of_imp_compact by auto show "convex T" using T face_of_imp_convex by blast qed qed lemma face_of_convex_hull_aux: assumes eq: "x *\<^sub>R p = u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c" and x: "u + v + w = x" "x \ 0" and S: "affine S" "a \ S" "b \ S" "c \ S" shows "p \ S" proof - have "p = (u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x" by (metis \x \ 0\ eq mult.commute right_inverse scaleR_one scaleR_scaleR) moreover have "affine hull {a,b,c} \ S" by (simp add: S hull_minimal) moreover have "(u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x \ affine hull {a,b,c}" apply (simp add: affine_hull_3) apply (rule_tac x="u/x" in exI) apply (rule_tac x="v/x" in exI) apply (rule_tac x="w/x" in exI) using x apply (auto simp: field_split_simps) done ultimately show ?thesis by force qed proposition face_of_convex_hull_insert_eq: fixes a :: "'a :: euclidean_space" assumes "finite S" and a: "a \ affine hull S" shows "(F face_of (convex hull (insert a S)) \ F face_of (convex hull S) \ (\F'. F' face_of (convex hull S) \ F = convex hull (insert a F')))" (is "F face_of ?CAS \ _") proof safe assume F: "F face_of ?CAS" and *: "\F'. F' face_of convex hull S \ F = convex hull insert a F'" obtain T where T: "T \ insert a S" and FeqT: "F = convex hull T" by (metis F \finite S\ compact_insert finite_imp_compact face_of_convex_hull_subset) show "F face_of convex hull S" proof (cases "a \ T") case True have "F = convex hull insert a (convex hull T \ convex hull S)" proof have "T \ insert a (convex hull T \ convex hull S)" using T hull_subset by fastforce then show "F \ convex hull insert a (convex hull T \ convex hull S)" by (simp add: FeqT hull_mono) show "convex hull insert a (convex hull T \ convex hull S) \ F" by (simp add: FeqT True hull_inc hull_minimal) qed moreover have "convex hull T \ convex hull S face_of convex hull S" by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI) ultimately show ?thesis using * by force next case False then show ?thesis by (metis FeqT F T face_of_subset hull_mono subset_insert subset_insertI) qed next assume "F face_of convex hull S" show "F face_of ?CAS" by (simp add: \F face_of convex hull S\ a face_of_convex_hull_insert \finite S\) next fix F assume F: "F face_of convex hull S" show "convex hull insert a F face_of ?CAS" proof (cases "S = {}") case True then show ?thesis using F face_of_affine_eq by auto next case False have anotc: "a \ convex hull S" by (metis (no_types) a affine_hull_convex_hull hull_inc) show ?thesis proof (cases "F = {}") case True show ?thesis using anotc by (simp add: \F = {}\ \finite S\ extreme_point_of_convex_hull_insert face_of_singleton) next case False have "convex hull insert a F \ ?CAS" by (simp add: F a \finite S\ convex_hull_subset face_of_convex_hull_insert face_of_imp_subset hull_inc) moreover have "(\y v. (1 - ub) *\<^sub>R a + ub *\<^sub>R b = (1 - v) *\<^sub>R a + v *\<^sub>R y \ 0 \ v \ v \ 1 \ y \ F) \ (\x u. (1 - uc) *\<^sub>R a + uc *\<^sub>R c = (1 - u) *\<^sub>R a + u *\<^sub>R x \ 0 \ u \ u \ 1 \ x \ F)" if *: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x \ open_segment ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 \ ub" "ub \ 1" "0 \ uc" "uc \ 1" "0 \ ux" "ux \ 1" and b: "b \ convex hull S" and c: "c \ convex hull S" and "x \ F" for b c ub uc ux x proof - have xah: "x \ affine hull S" using F convex_hull_subset_affine_hull face_of_imp_subset \x \ F\ by blast have ah: "b \ affine hull S" "c \ affine hull S" using b c convex_hull_subset_affine_hull by blast+ obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \ (1 - uc) *\<^sub>R a + uc *\<^sub>R c" and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x = (1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 < v" "v < 1" using * by (auto simp: in_segment) then have 0: "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a + (ux *\<^sub>R x - (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)) = 0" by (auto simp: algebra_simps) then have "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a = ((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x" by (auto simp: algebra_simps) then have "a \ affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \ 0" by (rule face_of_convex_hull_aux) (use b c xah ah that in \auto simp: algebra_simps\) then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0" using a by blast with 0 have equx: "(1 - v) * ub + v * uc = ux" and uxx: "ux *\<^sub>R x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)" by auto (auto simp: algebra_simps) show ?thesis proof (cases "uc = 0") case True then show ?thesis using equx \0 \ ub\ \ub \ 1\ \v < 1\ uxx \x \ F\ by force next case False show ?thesis proof (cases "ub = 0") case True then show ?thesis using equx \0 \ uc\ \uc \ 1\ \0 < v\ uxx \x \ F\ by force next case False then have "0 < ub" "0 < uc" using \uc \ 0\ \0 \ ub\ \0 \ uc\ by auto then have "(1 - v) * ub > 0" "v * uc > 0" by (simp_all add: \0 < uc\ \0 < v\ \v < 1\) then have "ux \ 0" using equx \0 < v\ by auto have "b \ F \ c \ F" proof (cases "b = c") case True then show ?thesis by (metis \ux \ 0\ equx real_vector.scale_cancel_left scaleR_add_left uxx \x \ F\) next case False have "x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c) /\<^sub>R ux" by (metis \ux \ 0\ uxx mult.commute right_inverse scaleR_one scaleR_scaleR) also have "... = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" using \ux \ 0\ equx apply (auto simp: field_split_simps) by (metis add.commute add_diff_eq add_divide_distrib diff_add_cancel scaleR_add_left) finally have "x = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" . then have "x \ open_segment b c" apply (simp add: in_segment \b \ c\) apply (rule_tac x="(v * uc) / ux" in exI) using \0 \ ux\ \ux \ 0\ \0 < uc\ \0 < v\ \0 < ub\ \v < 1\ equx apply (force simp: field_split_simps) done then show ?thesis by (rule face_ofD [OF F _ b c \x \ F\]) qed with \0 \ ub\ \ub \ 1\ \0 \ uc\ \uc \ 1\ show ?thesis by blast qed qed qed moreover have "convex hull F = F" by (meson F convex_hull_eq face_of_imp_convex) ultimately show ?thesis unfolding face_of_def by (fastforce simp: convex_hull_insert_alt \S \ {}\ \F \ {}\) qed qed qed lemma face_of_convex_hull_insert2: fixes a :: "'a :: euclidean_space" assumes S: "finite S" and a: "a \ affine hull S" and F: "F face_of convex hull S" shows "convex hull (insert a F) face_of convex hull (insert a S)" by (metis F face_of_convex_hull_insert_eq [OF S a]) proposition face_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(T face_of (convex hull S) \ (\c. c \ S \ T = convex hull c))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson \T face_of convex hull S\ aff_independent_finite assms face_of_convex_hull_subset finite_imp_compact) next assume ?rhs then obtain c where "c \ S" and T: "T = convex hull c" by blast have "affine hull c \ affine hull (S - c) = {}" by (intro disjoint_affine_hull [OF assms \c \ S\], auto) then have "affine hull c \ convex hull (S - c) = {}" using convex_hull_subset_affine_hull by fastforce then show ?lhs by (metis face_of_convex_hulls \c \ S\ aff_independent_finite assms T) qed lemma facet_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "T facet_of (convex hull S) \ T \ {} \ (\u. u \ S \ T = convex hull (S - {u}))" (is "?lhs = ?rhs") proof assume ?lhs then have "T face_of (convex hull S)" "T \ {}" and afft: "aff_dim T = aff_dim (convex hull S) - 1" by (auto simp: facet_of_def) then obtain c where "c \ S" and c: "T = convex hull c" by (auto simp: face_of_convex_hull_affine_independent [OF assms]) then have affs: "aff_dim S = aff_dim c + 1" by (metis aff_dim_convex_hull afft eq_diff_eq) have "\ affine_dependent c" using \c \ S\ affine_dependent_subset assms by blast with affs have "card (S - c) = 1" apply (simp add: aff_dim_affine_independent [symmetric] aff_dim_convex_hull) by (metis aff_dim_affine_independent aff_independent_finite One_nat_def \c \ S\ add.commute add_diff_cancel_right' assms card_Diff_subset card_mono of_nat_1 of_nat_diff of_nat_eq_iff) then obtain u where u: "u \ S - c" by (metis DiffI \c \ S\ aff_independent_finite assms cancel_comm_monoid_add_class.diff_cancel card_Diff_subset subsetI subset_antisym zero_neq_one) then have u: "S = insert u c" by (metis Diff_subset \c \ S\ \card (S - c) = 1\ card_1_singletonE double_diff insert_Diff insert_subset singletonD) have "T = convex hull (c - {u})" by (metis Diff_empty Diff_insert0 \T facet_of convex hull S\ c facet_of_irrefl insert_absorb u) with \T \ {}\ show ?rhs using c u by auto next assume ?rhs then obtain u where "T \ {}" "u \ S" and u: "T = convex hull (S - {u})" by (force simp: facet_of_def) then have "\ S \ {u}" using \T \ {}\ u by auto have "aff_dim (S - {u}) = aff_dim S - 1" using assms \u \ S\ unfolding affine_dependent_def by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S]) then have "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1" by (simp add: aff_dim_convex_hull) then show ?lhs by (metis Diff_subset \T \ {}\ assms face_of_convex_hull_affine_independent facet_of_def u) qed lemma facet_of_convex_hull_affine_independent_alt: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(T facet_of (convex hull S) \ 2 \ card S \ (\u. u \ S \ T = convex hull (S - {u})))" (is "?lhs = ?rhs") proof assume L: ?lhs then obtain x where "x \ S" and x: "T = convex hull (S - {x})" and "finite S" using assms facet_of_convex_hull_affine_independent aff_independent_finite by blast moreover have "Suc (Suc 0) \ card S" using L x \x \ S\ \finite S\ by (metis Suc_leI assms card.remove convex_hull_eq_empty card_gt_0_iff facet_of_convex_hull_affine_independent finite_Diff not_less_eq_eq) ultimately show ?rhs by auto next assume ?rhs then show ?lhs using assms by (auto simp: facet_of_convex_hull_affine_independent Set.subset_singleton_iff) qed lemma segment_face_of: assumes "(closed_segment a b) face_of S" shows "a extreme_point_of S" "b extreme_point_of S" proof - have as: "{a} face_of S" by (metis (no_types) assms convex_hull_singleton empty_iff extreme_point_of_convex_hull_insert face_of_face face_of_singleton finite.emptyI finite.insertI insert_absorb insert_iff segment_convex_hull) moreover have "{b} face_of S" proof - have "b \ convex hull {a} \ b extreme_point_of convex hull {b, a}" by (meson extreme_point_of_convex_hull_insert finite.emptyI finite.insertI) moreover have "closed_segment a b = convex hull {b, a}" using closed_segment_commute segment_convex_hull by blast ultimately show ?thesis by (metis as assms face_of_face convex_hull_singleton empty_iff face_of_singleton insertE) qed ultimately show "a extreme_point_of S" "b extreme_point_of S" using face_of_singleton by blast+ qed proposition Krein_Milman_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "compact S" shows "S = convex hull (frontier S)" (is "?lhs = ?rhs") proof have "?lhs \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski assms by blast also have "... \ ?rhs" proof (rule hull_mono) show "{x. x extreme_point_of S} \ frontier S" using closure_subset by (auto simp: frontier_def extreme_point_not_in_interior extreme_point_of_def) qed finally show "?lhs \ ?rhs" . next have "?rhs \ convex hull S" by (metis Diff_subset \compact S\ closure_closed compact_eq_bounded_closed frontier_def hull_mono) also have "... \ ?lhs" by (simp add: \convex S\ hull_same) finally show "?rhs \ ?lhs" . qed subsection\Polytopes\ definition\<^marker>\tag important\ polytope where "polytope S \ \v. finite v \ S = convex hull v" lemma polytope_translation_eq: "polytope (image (\x. a + x) S) \ polytope S" proof - have "\a A. polytope A \ polytope ((+) a ` A)" by (metis (no_types) convex_hull_translation finite_imageI polytope_def) then show ?thesis by (metis (no_types) add.left_inverse image_add_0 translation_assoc) qed lemma polytope_linear_image: "\linear f; polytope p\ \ polytope(image f p)" unfolding polytope_def using convex_hull_linear_image by blast lemma polytope_empty: "polytope {}" using convex_hull_empty polytope_def by blast lemma polytope_convex_hull: "finite S \ polytope(convex hull S)" using polytope_def by auto lemma polytope_Times: "\polytope S; polytope T\ \ polytope(S \ T)" unfolding polytope_def by (metis finite_cartesian_product convex_hull_Times) lemma face_of_polytope_polytope: fixes S :: "'a::euclidean_space set" shows "\polytope S; F face_of S\ \ polytope F" unfolding polytope_def by (meson face_of_convex_hull_subset finite_imp_compact finite_subset) lemma finite_polytope_faces: fixes S :: "'a::euclidean_space set" assumes "polytope S" shows "finite {F. F face_of S}" proof - obtain v where "finite v" "S = convex hull v" using assms polytope_def by auto have "finite ((hull) convex ` {T. T \ v})" by (simp add: \finite v\) moreover have "{F. F face_of S} \ ((hull) convex ` {T. T \ v})" by (metis (no_types, lifting) \finite v\ \S = convex hull v\ face_of_convex_hull_subset finite_imp_compact image_eqI mem_Collect_eq subsetI) ultimately show ?thesis by (blast intro: finite_subset) qed lemma finite_polytope_facets: assumes "polytope S" shows "finite {T. T facet_of S}" by (simp add: assms facet_of_def finite_polytope_faces) lemma polytope_scaling: assumes "polytope S" shows "polytope (image (\x. c *\<^sub>R x) S)" by (simp add: assms polytope_linear_image) lemma polytope_imp_compact: fixes S :: "'a::real_normed_vector set" shows "polytope S \ compact S" by (metis finite_imp_compact_convex_hull polytope_def) lemma polytope_imp_convex: "polytope S \ convex S" by (metis convex_convex_hull polytope_def) lemma polytope_imp_closed: fixes S :: "'a::real_normed_vector set" shows "polytope S \ closed S" by (simp add: compact_imp_closed polytope_imp_compact) lemma polytope_imp_bounded: fixes S :: "'a::real_normed_vector set" shows "polytope S \ bounded S" by (simp add: compact_imp_bounded polytope_imp_compact) lemma polytope_interval: "polytope(cbox a b)" unfolding polytope_def by (meson closed_interval_as_convex_hull) lemma polytope_sing: "polytope {a}" using polytope_def by force lemma face_of_polytope_insert: "\polytope S; a \ affine hull S; F face_of S\ \ F face_of convex hull (insert a S)" by (metis (no_types, lifting) affine_hull_convex_hull face_of_convex_hull_insert hull_insert polytope_def) proposition face_of_polytope_insert2: fixes a :: "'a :: euclidean_space" assumes "polytope S" "a \ affine hull S" "F face_of S" shows "convex hull (insert a F) face_of convex hull (insert a S)" proof - obtain V where "finite V" "S = convex hull V" using assms by (auto simp: polytope_def) then have "convex hull (insert a F) face_of convex hull (insert a V)" using affine_hull_convex_hull assms face_of_convex_hull_insert2 by blast then show ?thesis by (metis \S = convex hull V\ hull_insert) qed subsection\Polyhedra\ definition\<^marker>\tag important\ polyhedron where "polyhedron S \ \F. finite F \ S = \ F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b})" lemma polyhedron_Int [intro,simp]: "\polyhedron S; polyhedron T\ \ polyhedron (S \ T)" apply (clarsimp simp add: polyhedron_def) subgoal for F G by (rule_tac x="F \ G" in exI, auto) done lemma polyhedron_UNIV [iff]: "polyhedron UNIV" unfolding polyhedron_def by (rule_tac x="{}" in exI) auto lemma polyhedron_Inter [intro,simp]: "\finite F; \S. S \ F \ polyhedron S\ \ polyhedron(\F)" by (induction F rule: finite_induct) auto lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)" proof - define i::'a where "(i \ SOME i. i \ Basis)" have "\a. a \ 0 \ (\b. {x. i \ x \ -1} = {x. a \ x \ b})" by (rule_tac x="i" in exI) (force simp: i_def SOME_Basis nonzero_Basis) moreover have "\a b. a \ 0 \ {x. -i \ x \ - 1} = {x. a \ x \ b}" apply (rule_tac x="-i" in exI) apply (rule_tac x="-1" in exI) apply (simp add: i_def SOME_Basis nonzero_Basis) done ultimately show ?thesis unfolding polyhedron_def by (rule_tac x="{{x. i \ x \ -1}, {x. -i \ x \ -1}}" in exI) force qed lemma polyhedron_halfspace_le: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis unfolding polyhedron_def by (rule_tac x="{{x. a \ x \ b}}" in exI) auto qed lemma polyhedron_halfspace_ge: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" using polyhedron_halfspace_le [of "-a" "-b"] by simp lemma polyhedron_hyperplane: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x = b}" proof - have "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by force then show ?thesis by (simp add: polyhedron_halfspace_ge polyhedron_halfspace_le) qed lemma affine_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "affine S \ polyhedron S" by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S]) lemma polyhedron_imp_closed: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S" by (metis closed_Inter closed_halfspace_le polyhedron_def) lemma polyhedron_imp_convex: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ convex S" by (metis convex_Inter convex_halfspace_le polyhedron_def) lemma polyhedron_affine_hull: fixes S :: "'a :: euclidean_space set" shows "polyhedron(affine hull S)" by (simp add: affine_imp_polyhedron) subsection\Canonical polyhedron representation making facial structure explicit\ proposition polyhedron_Int_affine: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using hull_subset polyhedron_def by fastforce next assume ?rhs then show ?lhs by (metis polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le) qed proposition rel_interior_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "rel_interior S = {x \ S. \h \ F. a h \ x < b h}" proof - have rels: "\x. x \ rel_interior S \ x \ S" by (meson IntE mem_rel_interior) moreover have "a i \ x < b i" if x: "x \ rel_interior S" and "i \ F" for x i proof - have fif: "F - {i} \ F" using \i \ F\ Diff_insert_absorb Diff_subset set_insert psubsetI by blast then have "S \ affine hull S \ \(F - {i})" by (rule psub) then obtain z where ssub: "S \ \(F - {i})" and zint: "z \ \(F - {i})" and "z \ S" and zaff: "z \ affine hull S" by auto have "z \ x" using \z \ S\ rels x by blast have "z \ affine hull S \ \F" using \z \ S\ seq by auto then have aiz: "a i \ z > b i" using faceq zint zaff by fastforce obtain e where "e > 0" "x \ S" and e: "ball x e \ affine hull S \ S" using x by (auto simp: mem_rel_interior_ball) then have ins: "\y. \norm (x - y) < e; y \ affine hull S\ \ y \ S" by (metis IntI subsetD dist_norm mem_ball) define \ where "\ = min (1/2) (e / 2 / norm(z - x))" have "norm (\ *\<^sub>R x - \ *\<^sub>R z) = norm (\ *\<^sub>R (x - z))" by (simp add: \_def algebra_simps norm_mult) also have "... = \ * norm (x - z)" using \e > 0\ by (simp add: \_def) also have "... < e" using \z \ x\ \e > 0\ by (simp add: \_def min_def field_split_simps norm_minus_commute) finally have lte: "norm (\ *\<^sub>R x - \ *\<^sub>R z) < e" . have \_aff: "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ affine hull S" by (metis \x \ S\ add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff) have "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ S" using ins [OF _ \_aff] by (simp add: algebra_simps lte) then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \ S" using \e > 0\ \z \ x\ by (rule_tac l = \ in that) (auto simp: \_def) then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \ i" using seq \i \ F\ by auto have "b i * l + (a i \ x) * (1 - l) < a i \ (l *\<^sub>R z + (1 - l) *\<^sub>R x)" using l by (simp add: algebra_simps aiz) also have "\ \ b i" using i l using faceq mem_Collect_eq \i \ F\ by blast finally have "(a i \ x) * (1 - l) < b i * (1 - l)" by (simp add: algebra_simps) with l show ?thesis by simp qed moreover have "x \ rel_interior S" if "x \ S" and less: "\h. h \ F \ a h \ x < b h" for x proof - have 1: "\h. h \ F \ x \ interior h" by (metis interior_halfspace_le mem_Collect_eq less faceq) have 2: "\y. \\h\F. y \ interior h; y \ affine hull S\ \ y \ S" by (metis IntI Inter_iff subsetD interior_subset seq) show ?thesis apply (simp add: rel_interior \x \ S\) apply (rule_tac x="\h\F. interior h" in exI) apply (auto simp: \finite F\ open_INT 1 2) done qed ultimately show ?thesis by blast qed lemma polyhedron_Int_affine_parallel: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then obtain F where "finite F" and seq: "S = (affine hull S) \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" by (fastforce simp add: polyhedron_Int_affine) then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?rhs proof - have "\a' b'. a' \ 0 \ affine hull S \ {x. a' \ x \ b'} = affine hull S \ h \ (\w \ affine hull S. (w + a') \ affine hull S)" if "h \ F" "\(affine hull S \ h)" for h proof - have "a h \ 0" and "h = {x. a h \ x \ b h}" "h \ \F = \F" using \h \ F\ ab by auto then have "(affine hull S) \ {x. a h \ x \ b h} \ {}" by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2)) moreover have "\ (affine hull S \ {x. a h \ x \ b h})" using \h = {x. a h \ x \ b h}\ that(2) by blast ultimately show ?thesis using affine_parallel_slice [of "affine hull S"] by (metis \h = {x. a h \ x \ b h}\ affine_affine_hull) qed then obtain a b where ab: "\h. \h \ F; \ (affine hull S \ h)\ \ a h \ 0 \ affine hull S \ {x. a h \ x \ b h} = affine hull S \ h \ (\w \ affine hull S. (w + a h) \ affine hull S)" by metis have seq2: "S = affine hull S \ (\h\{h \ F. \ affine hull S \ h}. {x. a h \ x \ b h})" by (subst seq) (auto simp: ab INT_extend_simps) show ?thesis apply (rule_tac x="(\h. {x. a h \ x \ b h}) ` {h. h \ F \ \(affine hull S \ h)}" in exI) apply (intro conjI seq2) using \finite F\ apply force using ab apply blast done qed next assume ?rhs then show ?lhs by (metis polyhedron_Int_affine) qed proposition polyhedron_Int_affine_parallel_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)) \ (\F'. F' \ F \ S \ (affine hull S) \ (\F')))" (is "?lhs = ?rhs") proof assume ?lhs then obtain f0 where f0: "finite f0" "S = (affine hull S) \ (\f0)" (is "?P f0") "\h \ f0. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)" (is "?Q f0") by (force simp: polyhedron_Int_affine_parallel) define n where "n = (LEAST n. \F. card F = n \ finite F \ ?P F \ ?Q F)" have nf: "\F. card F = n \ finite F \ ?P F \ ?Q F" apply (simp add: n_def) apply (rule LeastI [where k = "card f0"]) using f0 apply auto done then obtain F where F: "card F = n" "finite F" and seq: "?P F" and aff: "?Q F" by blast then have "\ (finite g \ ?P g \ ?Q g)" if "card g < n" for g using that by (auto simp: n_def dest!: not_less_Least) then have *: "\ (?P g \ ?Q g)" if "g \ F" for g using that \finite F\ psubset_card_mono \card F = n\ by (metis finite_Int inf.strict_order_iff) have 1: "\F'. F' \ F \ S \ affine hull S \ \F'" by (subst seq) blast have 2: "S \ affine hull S \ \F'" if "F' \ F" for F' using * [OF that] by (metis IntE aff inf.strict_order_iff that) show ?rhs by (metis \finite F\ seq aff psubsetI 1 2) next assume ?rhs then show ?lhs by (auto simp: polyhedron_Int_affine_parallel) qed lemma polyhedron_Int_affine_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}) \ (\F'. F' \ F \ S \ (affine hull S) \ \F'))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward) qed (auto simp: polyhedron_Int_affine elim!: ex_forward) proposition facet_of_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "C facet_of S \ (\h. h \ F \ C = S \ {x. a h \ x = b h})" proof (cases "S = {}") case True with psub show ?thesis by force next case False have "polyhedron S" unfolding polyhedron_Int_affine by (metis \finite F\ faceq seq) then have "convex S" by (rule polyhedron_imp_convex) with False rel_interior_eq_empty have "rel_interior S \ {}" by blast then obtain x where "x \ rel_interior S" by auto then obtain T where "open T" "x \ T" "x \ S" "T \ affine hull S \ S" by (force simp: mem_rel_interior) then have xaff: "x \ affine hull S" and xint: "x \ \F" using seq hull_inc by auto have "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) with \x \ rel_interior S\ have [simp]: "\h. h\F \ a h \ x < b h" by blast have *: "(S \ {x. a h \ x = b h}) facet_of S" if "h \ F" for h proof - have "S \ affine hull S \ \(F - {h})" using psub that by (metis Diff_disjoint Diff_subset insert_disjoint(2) psubsetI) then obtain z where zaff: "z \ affine hull S" and zint: "z \ \(F - {h})" and "z \ S" by force then have "z \ x" "z \ h" using seq \x \ S\ by auto have "x \ h" using that xint by auto then have able: "a h \ x \ b h" using faceq that by blast also have "... < a h \ z" using \z \ h\ faceq [OF that] xint by auto finally have xltz: "a h \ x < a h \ z" . define l where "l = (b h - a h \ x) / (a h \ z - a h \ x)" define w where "w = (1 - l) *\<^sub>R x + l *\<^sub>R z" have "0 < l" "l < 1" using able xltz \b h < a h \ z\ \h \ F\ by (auto simp: l_def field_split_simps) have awlt: "a i \ w < b i" if "i \ F" "i \ h" for i proof - have "(1 - l) * (a i \ x) < (1 - l) * b i" by (simp add: \l < 1\ \i \ F\) moreover have "l * (a i \ z) \ l * b i" proof (rule mult_left_mono) show "a i \ z \ b i" by (metis Diff_insert_absorb Inter_iff Set.set_insert \h \ F\ faceq insertE mem_Collect_eq that zint) qed (use \0 < l\ in auto) ultimately show ?thesis by (simp add: w_def algebra_simps) qed have weq: "a h \ w = b h" using xltz unfolding w_def l_def by (simp add: algebra_simps) (simp add: field_simps) have faceS: "S \ {x. a h \ x = b h} face_of S" proof (rule face_of_Int_supporting_hyperplane_le) show "\x. x \ S \ a h \ x \ b h" using faceq seq that by fastforce qed fact have "w \ affine hull S" by (simp add: w_def mem_affine xaff zaff) moreover have "w \ \F" using \a h \ w = b h\ awlt faceq less_eq_real_def by blast ultimately have "w \ S" using seq by blast with weq have ne: "S \ {x. a h \ x = b h} \ {}" by blast moreover have "affine hull (S \ {x. a h \ x = b h}) = (affine hull S) \ {x. a h \ x = b h}" proof show "affine hull (S \ {x. a h \ x = b h}) \ affine hull S \ {x. a h \ x = b h}" apply (intro Int_greatest hull_mono Int_lower1) apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2) done next show "affine hull S \ {x. a h \ x = b h} \ affine hull (S \ {x. a h \ x = b h})" proof fix y assume yaff: "y \ affine hull S \ {y. a h \ y = b h}" obtain T where "0 < T" and T: "\j. \j \ F; j \ h\ \ T * (a j \ y - a j \ w) \ b j - a j \ w" proof (cases "F - {h} = {}") case True then show ?thesis by (rule_tac T=1 in that) auto next case False then obtain h' where h': "h' \ F - {h}" by auto let ?body = "(\j. if 0 < a j \ y - a j \ w then (b j - a j \ w) / (a j \ y - a j \ w) else 1) ` (F - {h})" define inff where "inff = Inf ?body" from \finite F\ have "finite ?body" by blast moreover from h' have "?body \ {}" by blast moreover have "j > 0" if "j \ ?body" for j proof - from that obtain x where "x \ F" and "x \ h" and *: "j = (if 0 < a x \ y - a x \ w then (b x - a x \ w) / (a x \ y - a x \ w) else 1)" by blast with awlt [of x] have "a x \ w < b x" by simp with * show ?thesis by simp qed ultimately have "0 < inff" by (simp_all add: finite_less_Inf_iff inff_def) moreover have "inff * (a j \ y - a j \ w) \ b j - a j \ w" if "j \ F" "j \ h" for j proof (cases "a j \ w < a j \ y") case True then have "inff \ (b j - a j \ w) / (a j \ y - a j \ w)" unfolding inff_def using \finite F\ by (auto intro: cInf_le_finite simp add: that split: if_split_asm) then show ?thesis using \0 < inff\ awlt [OF that] mult_strict_left_mono by (fastforce simp add: field_split_simps split: if_split_asm) next case False with \0 < inff\ have "inff * (a j \ y - a j \ w) \ 0" by (simp add: mult_le_0_iff) also have "... < b j - a j \ w" by (simp add: awlt that) finally show ?thesis by simp qed ultimately show ?thesis by (blast intro: that) qed define C where "C = (1 - T) *\<^sub>R w + T *\<^sub>R y" have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ j" if "j \ F" for j proof (cases "j = h") case True have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a h \ x \ b h}" using weq yaff by (auto simp: algebra_simps) with True faceq [OF that] show ?thesis by metis next case False with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a j \ x \ b j}" by (simp add: algebra_simps) with faceq [OF that] show ?thesis by simp qed moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ affine hull S" using yaff \w \ affine hull S\ affine_affine_hull affine_alt by blast ultimately have "C \ S" using seq by (force simp: C_def) moreover have "a h \ C = b h" using yaff by (force simp: C_def algebra_simps weq) ultimately have caff: "C \ affine hull (S \ {y. a h \ y = b h})" by (simp add: hull_inc) have waff: "w \ affine hull (S \ {y. a h \ y = b h})" using \w \ S\ weq by (blast intro: hull_inc) have yeq: "y = (1 - inverse T) *\<^sub>R w + C /\<^sub>R T" using \0 < T\ by (simp add: C_def algebra_simps) show "y \ affine hull (S \ {y. a h \ y = b h})" by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff]) qed qed ultimately have "aff_dim (affine hull (S \ {x. a h \ x = b h})) = aff_dim S - 1" using \b h < a h \ z\ zaff by (force simp: aff_dim_affine_Int_hyperplane) then show ?thesis by (simp add: ne faceS facet_of_def) qed show ?thesis proof show "\h. h \ F \ C = S \ {x. a h \ x = b h} \ C facet_of S" using * by blast next assume "C facet_of S" then have "C face_of S" "convex C" "C \ {}" and affc: "aff_dim C = aff_dim S - 1" by (auto simp: facet_of_def face_of_imp_convex) then obtain x where x: "x \ rel_interior C" by (force simp: rel_interior_eq_empty) then have "x \ C" by (meson subsetD rel_interior_subset) then have "x \ S" using \C facet_of S\ facet_of_imp_subset by blast have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF assms]) have "C \ S" using \C facet_of S\ facet_of_irrefl by blast then have "x \ rel_interior S" by (metis IntI empty_iff \x \ C\ \C \ S\ \C face_of S\ face_of_disjoint_rel_interior) with rels \x \ S\ obtain i where "i \ F" and i: "a i \ x \ b i" by force have "x \ {u. a i \ u \ b i}" by (metis IntD2 InterE \i \ F\ \x \ S\ faceq seq) then have "a i \ x \ b i" by simp then have "a i \ x = b i" using i by auto have "C \ S \ {x. a i \ x = b i}" proof (rule subset_of_face_of [of _ S]) show "S \ {x. a i \ x = b i} face_of S" by (simp add: "*" \i \ F\ facet_of_imp_face_of) show "C \ S" by (simp add: \C face_of S\ face_of_imp_subset) show "S \ {x. a i \ x = b i} \ rel_interior C \ {}" using \a i \ x = b i\ \x \ S\ x by blast qed then have cface: "C face_of (S \ {x. a i \ x = b i})" by (meson \C face_of S\ face_of_subset inf_le1) have con: "convex (S \ {x. a i \ x = b i})" by (simp add: \convex S\ convex_Int convex_hyperplane) show "\h. h \ F \ C = S \ {x. a h \ x = b h}" apply (rule_tac x=i in exI) by (metis (no_types) * \i \ F\ affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface]) qed qed lemma face_of_polyhedron_subset_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and C: "C face_of S" and "C \ {}" "C \ S" obtains h where "h \ F" "C \ S \ {x. a h \ x = b h}" proof - have "C \ S" using \C face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" by (metis \finite F\ faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ {x. a h \ x = b h}) face_of S" if "h \ F" for h using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce have "rel_interior C \ {}" using C \C \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain x where "x \ rel_interior C" by auto have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "x \ rel_interior S" by (metis IntI \x \ rel_interior C\ C \C \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "x \ S" using \C \ S\ \x \ rel_interior C\ rel_interior_subset by auto then have xint: "x \ \F" using seq by blast have "F \ {}" using assms by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial) then obtain i where "i \ F" "\ (a i \ x < b i)" using \x \ S\ rels xnot by auto with xint have "a i \ x = b i" by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq) have face: "S \ {x. a i \ x = b i} face_of S" by (simp add: "*" \i \ F\) show ?thesis proof show "C \ S \ {x. a i \ x = b i}" using subset_of_face_of [OF face \C \ S\] \a i \ x = b i\ \x \ rel_interior C\ \x \ S\ by blast qed fact qed text\Initial part of proof duplicates that above\ proposition face_of_polyhedron_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and C: "C face_of S" and "C \ {}" "C \ S" shows "C = \{S \ {x. a h \ x = b h} | h. h \ F \ C \ S \ {x. a h \ x = b h}}" proof - let ?ab = "\h. {x. a h \ x = b h}" have "C \ S" using \C face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" by (metis \finite F\ faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ ?ab h) face_of S" if "h \ F" for h using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce have "rel_interior C \ {}" using C \C \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain z where z: "z \ rel_interior C" by auto have rels: "rel_interior S = {z \ S. \h\F. a h \ z < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "z \ rel_interior S" by (metis IntI \z \ rel_interior C\ C \C \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "z \ S" using \C \ S\ \z \ rel_interior C\ rel_interior_subset by auto with seq have xint: "z \ \F" by blast have "open (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto simp: \finite F\ open_halfspace_lt open_INT) then obtain e where "0 < e" "ball z e \ (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto intro: openE [of _ z]) then have e: "\h. \h \ F; a h \ z < b h\ \ ball z e \ {w. a h \ w < b h}" by blast have "C \ (S \ ?ab h) \ z \ S \ ?ab h" if "h \ F" for h proof show "z \ S \ ?ab h \ C \ S \ ?ab h" by (metis "*" Collect_cong IntI \C \ S\ empty_iff subset_of_face_of that z) next show "C \ S \ ?ab h \ z \ S \ ?ab h" using \z \ rel_interior C\ rel_interior_subset by force qed then have **: "{S \ ?ab h | h. h \ F \ C \ S \ C \ ?ab h} = {S \ ?ab h |h. h \ F \ z \ S \ ?ab h}" by blast have bsub: "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S \ \F \ \{?ab h |h. h \ F \ a h \ z = b h}" if "i \ F" and i: "a i \ z = b i" for i proof - have sub: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ j" if "j \ F" for j proof - have "a j \ z \ b j" using faceq that xint by auto then consider "a j \ z < b j" | "a j \ z = b j" by linarith then have "\G. G \ {?ab h |h. h \ F \ a h \ z = b h} \ ball z e \ G \ j" proof cases assume "a j \ z < b j" then have "ball z e \ {x. a i \ x = b i} \ j" using e [OF \j \ F\] faceq that by (fastforce simp: ball_def) then show ?thesis by (rule_tac x="{x. a i \ x = b i}" in exI) (force simp: \i \ F\ i) next assume eq: "a j \ z = b j" with faceq that show ?thesis by (rule_tac x="{x. a j \ x = b j}" in exI) (fastforce simp add: \j \ F\) qed then show ?thesis by blast qed have 1: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S" using that \z \ S\ by (intro hull_mono) auto have 2: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{?ab h |h. h \ F \ a h \ z = b h}" by (rule hull_minimal) (auto intro: affine_hyperplane) have 3: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ \F" by (iprover intro: sub Inter_greatest) have *: "\A \ (B :: 'a set); A \ C; E \ C \ D\ \ E \ A \ (B \ D) \ C" for A B C D E by blast show ?thesis by (intro * 1 2 3) qed have "\h. h \ F \ C \ ?ab h" using assms by (metis face_of_polyhedron_subset_explicit [OF \finite F\ seq faceq psub] le_inf_iff) then have fac: "\{S \ ?ab h |h. h \ F \ C \ S \ ?ab h} face_of S" using * by (force simp: \C \ S\ intro: face_of_Inter) have red: "(\a. P a \ T \ S \ \{F X |X. P X}) \ T \ \{S \ F X |X::'a set. P X}" for P T F by blast have "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{S \ ?ab h |h. h \ F \ a h \ z = b h}" by (rule red) (metis seq bsub) with \0 < e\ have zinrel: "z \ rel_interior (\{S \ ?ab h |h. h \ F \ z \ S \ a h \ z = b h})" by (auto simp: mem_rel_interior_ball \z \ S\) show ?thesis using z zinrel by (intro face_of_eq [OF C fac]) (force simp: **) qed subsection\More general corollaries from the explicit representation\ corollary facet_of_polyhedron: assumes "polyhedron S" and "C facet_of S" obtains a b where "a \ 0" "S \ {x. a \ x \ b}" "C = S \ {x. a \ x = b}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis obtain i where "i \ F" and C: "C = S \ {x. a i \ x = b i}" using facet_of_polyhedron_explicit [OF \finite F\ seq ab min] assms by force moreover have ssub: "S \ {x. a i \ x \ b i}" using \i \ F\ ab by (subst seq) auto ultimately show ?thesis by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab) qed corollary face_of_polyhedron: assumes "polyhedron S" and "C face_of S" and "C \ {}" and "C \ S" shows "C = \{F. F facet_of S \ C \ F}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?thesis apply (subst face_of_polyhedron_explicit [OF \finite F\ seq ab min]) apply (auto simp: assms facet_of_polyhedron_explicit [OF \finite F\ seq ab min] cong: Collect_cong) done qed lemma face_of_polyhedron_subset_facet: assumes "polyhedron S" and "C face_of S" and "C \ {}" and "C \ S" obtains F where "F facet_of S" "C \ F" using face_of_polyhedron assms by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq) lemma exposed_face_of_polyhedron: assumes "polyhedron S" shows "F exposed_face_of S \ F face_of S" proof show "F exposed_face_of S \ F face_of S" by (simp add: exposed_face_of_def) next assume "F face_of S" show "F exposed_face_of S" proof (cases "F = {} \ F = S") case True then show ?thesis using \F face_of S\ exposed_face_of by blast next case False then have "{g. g facet_of S \ F \ g} \ {}" by (metis Collect_empty_eq_bot \F face_of S\ assms empty_def face_of_polyhedron_subset_facet) moreover have "\T. \T facet_of S; F \ T\ \ T exposed_face_of S" by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron) ultimately have "\{G. G facet_of S \ F \ G} exposed_face_of S" by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter) then show ?thesis using False \F face_of S\ assms face_of_polyhedron by fastforce qed qed lemma face_of_polyhedron_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" "c face_of S" shows "polyhedron c" by (metis assms face_of_imp_eq_affine_Int polyhedron_Int polyhedron_affine_hull polyhedron_imp_convex) lemma finite_polyhedron_faces: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "finite {F. F face_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have "finite {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" by (simp add: \finite F\) moreover have "{F. F face_of S} - {{}, S} \ {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" apply clarify apply (rename_tac c) apply (drule face_of_polyhedron_explicit [OF \finite F\ seq ab min, simplified], simp_all) apply (rule_tac x="{h \ F. c \ S \ {x. a h \ x = b h}}" in exI, auto) done ultimately show ?thesis by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset) qed lemma finite_polyhedron_exposed_faces: "polyhedron S \ finite {F. F exposed_face_of S}" using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce lemma finite_polyhedron_extreme_points: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "finite {v. v extreme_point_of S}" proof - have "finite {v. {v} face_of S}" using assms by (intro finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto) then show ?thesis by (simp add: face_of_singleton) qed lemma finite_polyhedron_facets: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ finite {F. F facet_of S}" unfolding facet_of_def by (blast intro: finite_subset [OF _ finite_polyhedron_faces]) proposition rel_interior_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_interior S = S - \{F. F facet_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have facet: "(c facet_of S) \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" for c by (rule facet_of_polyhedron_explicit [OF \finite F\ seq ab min]) have rel: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq ab min]) have "a h \ x < b h" if "x \ S" "h \ F" and xnot: "x \ \{F. F facet_of S}" for x h proof - have "x \ \F" using seq that by force with \h \ F\ ab have "a h \ x \ b h" by auto then consider "a h \ x < b h" | "a h \ x = b h" by linarith then show ?thesis proof cases case 1 then show ?thesis . next case 2 have "Collect ((\) x) \ Collect ((\) (\{A. A facet_of S}))" using xnot by fastforce then have "F \ Collect ((\) h)" using 2 \x \ S\ facet by blast with 2 that \x \ \F\ show ?thesis by blast qed qed moreover have "\h\F. a h \ x \ b h" if "x \ \{F. F facet_of S}" for x using that by (force simp: facet) ultimately show ?thesis by (force simp: rel) qed lemma rel_boundary_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "S - rel_interior S = \ {F. F facet_of S}" using facet_of_imp_subset by (fastforce simp add: rel_interior_of_polyhedron assms) lemma rel_frontier_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F facet_of S}" by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron) lemma rel_frontier_of_polyhedron_alt: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F face_of S \ F \ S}" proof show "rel_frontier S \ \ {F. F face_of S \ F \ S}" by (force simp: rel_frontier_of_polyhedron facet_of_def assms) qed (use face_of_subset_rel_frontier in fastforce) text\A characterization of polyhedra as having finitely many faces\ proposition polyhedron_eq_finite_exposed_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F exposed_face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: polyhedron_imp_closed polyhedron_imp_convex finite_polyhedron_exposed_faces) next assume ?rhs then have "closed S" "convex S" and fin: "finite {F. F exposed_face_of S}" by auto show ?lhs proof (cases "S = {}") case True then show ?thesis by auto next case False define F where "F = {h. h exposed_face_of S \ h \ {} \ h \ S}" have "finite F" by (simp add: fin F_def) have hface: "h face_of S" and "\a b. a \ 0 \ S \ {x. a \ x \ b} \ h = S \ {x. a \ x = b}" if "h \ F" for h using exposed_face_of F_def that by blast+ then obtain a b where ab: "\h. h \ F \ a h \ 0 \ S \ {x. a h \ x \ b h} \ h = S \ {x. a h \ x = b h}" by metis have *: "False" if paff: "p \ affine hull S" and "p \ S" and pint: "p \ \{{x. a h \ x \ b h} |h. h \ F}" for p proof - have "rel_interior S \ {}" by (simp add: \S \ {}\ \convex S\ rel_interior_eq_empty) then obtain c where c: "c \ rel_interior S" by auto with rel_interior_subset have "c \ S" by blast have ccp: "closed_segment c p \ affine hull S" by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE) have oS: "openin (top_of_set (closed_segment c p)) (closed_segment c p \ rel_interior S)" by (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp]) obtain x where xcl: "x \ closed_segment c p" and "x \ S" and xnot: "x \ rel_interior S" using connected_openin [of "closed_segment c p"] apply simp apply (drule_tac x="closed_segment c p \ rel_interior S" in spec) apply (drule mp [OF _ oS]) apply (drule_tac x="closed_segment c p \ (- S)" in spec) using rel_interior_subset \closed S\ c \p \ S\ apply blast done then obtain \ where "0 \ \" "\ \ 1" and xeq: "x = (1 - \) *\<^sub>R c + \ *\<^sub>R p" by (auto simp: in_segment) show False proof (cases "\=0 \ \=1") case True with xeq c xnot \x \ S\ \p \ S\ show False by auto next case False then have xos: "x \ open_segment c p" using \x \ S\ c open_segment_def that(2) xcl xnot by auto have xclo: "x \ closure S" using \x \ S\ closure_subset by blast obtain d where "d \ 0" and dle: "\y. y \ closure S \ d \ x \ d \ y" and dless: "\y. y \ rel_interior S \ d \ x < d \ y" by (metis supporting_hyperplane_relative_frontier [OF \convex S\ xclo xnot]) have sex: "S \ {y. d \ y = d \ x} exposed_face_of S" by (simp add: \closed S\ dle exposed_face_of_Int_supporting_hyperplane_ge [OF \convex S\]) have sne: "S \ {y. d \ y = d \ x} \ {}" using \x \ S\ by blast have sns: "S \ {y. d \ y = d \ x} \ S" by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset) obtain h where "h \ F" "x \ h" using F_def \x \ S\ sex sns by blast have abface: "{y. a h \ y = b h} face_of {y. a h \ y \ b h}" using hyperplane_face_of_halfspace_le by blast then have "c \ h" using face_ofD [OF abface xos] \c \ S\ \h \ F\ ab pint \x \ h\ by blast with c have "h \ rel_interior S \ {}" by blast then show False using \h \ F\ F_def face_of_disjoint_rel_interior hface by auto qed qed have "S \ affine hull S \ \{{x. a h \ x \ b h} |h. h \ F}" using ab by (auto simp: hull_subset) moreover have "affine hull S \ \{{x. a h \ x \ b h} |h. h \ F} \ S" using * by blast ultimately have "S = affine hull S \ \ {{x. a h \ x \ b h} |h. h \ F}" .. then show ?thesis apply (rule ssubst) apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le simp: \finite F\) done qed qed corollary polyhedron_eq_finite_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polyhedron_faces polyhedron_imp_closed polyhedron_imp_convex) next assume ?rhs then show ?lhs by (force simp: polyhedron_eq_finite_exposed_faces exposed_face_of intro: finite_subset) qed lemma polyhedron_linear_image_eq: fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "linear h" "bij h" shows "polyhedron (h ` S) \ polyhedron S" proof - have *: "{f. P f} = (image h) ` {f. P (h ` f)}" for P apply safe apply (rule_tac x="inv h ` x" in image_eqI) apply (auto simp: \bij h\ bij_is_surj image_f_inv_f) done have "inj h" using bij_is_inj assms by blast then have injim: "inj_on ((`) h) A" for A by (simp add: inj_on_def inj_image_eq_iff) show ?thesis using \linear h\ \inj h\ apply (simp add: polyhedron_eq_finite_faces closed_injective_linear_image_eq) apply (simp add: * face_of_linear_image [of h _ S, symmetric] finite_image_iff injim) done qed lemma polyhedron_negations: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ polyhedron(image uminus S)" by (subst polyhedron_linear_image_eq) (auto simp: bij_uminus intro!: linear_uminus) subsection\Relation between polytopes and polyhedra\ proposition polytope_eq_bounded_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S \ bounded S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polytope_faces polyhedron_eq_finite_faces polytope_imp_closed polytope_imp_convex polytope_imp_bounded) next assume R: ?rhs then have "finite {v. v extreme_point_of S}" by (simp add: finite_polyhedron_extreme_points) moreover have "S = convex hull {v. v extreme_point_of S}" using R by (simp add: Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex) ultimately show ?lhs unfolding polytope_def by blast qed lemma polytope_Int: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polytope T\ \ polytope(S \ T)" by (simp add: polytope_eq_bounded_polyhedron bounded_Int) lemma polytope_Int_polyhedron: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polyhedron T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polyhedron_Int_polytope: fixes S :: "'a :: euclidean_space set" shows "\polyhedron S; polytope T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polytope_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S" by (simp add: polytope_eq_bounded_polyhedron) lemma polytope_facet_exists: fixes p :: "'a :: euclidean_space set" assumes "polytope p" "0 < aff_dim p" obtains F where "F facet_of p" proof (cases "p = {}") case True with assms show ?thesis by auto next case False then obtain v where "v extreme_point_of p" using extreme_point_exists_convex by (blast intro: \polytope p\ polytope_imp_compact polytope_imp_convex) then show ?thesis by (metis face_of_polyhedron_subset_facet polytope_imp_polyhedron aff_dim_sing all_not_in_conv assms face_of_singleton less_irrefl singletonI that) qed lemma polyhedron_interval [iff]: "polyhedron(cbox a b)" by (metis polytope_imp_polyhedron polytope_interval) lemma polyhedron_convex_hull: fixes S :: "'a :: euclidean_space set" shows "finite S \ polyhedron(convex hull S)" by (simp add: polytope_convex_hull polytope_imp_polyhedron) subsection\Relative and absolute frontier of a polytope\ lemma rel_boundary_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(convex hull S) - rel_interior(convex hull S) = (\a\S. convex hull (S - {a}))" proof - have "finite S" by (metis assms aff_independent_finite) then consider "card S = 0" | "card S = 1" | "2 \ card S" by arith then show ?thesis proof cases case 1 then have "S = {}" by (simp add: \finite S\) then show ?thesis by simp next case 2 show ?thesis by (auto intro: card_1_singletonE [OF \card S = 1\]) next case 3 with assms show ?thesis by (auto simp: polyhedron_convex_hull rel_boundary_of_polyhedron facet_of_convex_hull_affine_independent_alt \finite S\) qed qed proposition frontier_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "card S = Suc (DIM('a))" shows "frontier(convex hull S) = \ {convex hull (S - {a}) | a. a \ S}" proof (cases "affine_dependent S") case True have [iff]: "finite S" using assms using card.infinite by force then have ccs: "closed (convex hull S)" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) { fix x T assume "int (card T) \ aff_dim S + 1" "finite T" "T \ S""x \ convex hull T" then have "S \ T" using True \finite S\ aff_dim_le_card affine_independent_iff_card by fastforce then obtain a where "a \ S" "a \ T" using \T \ S\ by blast then have "\y\S. x \ convex hull (S - {y})" using True affine_independent_iff_card [of S] by (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \a \ T\ \T \ S\ \x \ convex hull T\ hull_mono insert_Diff_single subsetCE) } note * = this have 1: "convex hull S \ (\ a\S. convex hull (S - {a}))" by (subst caratheodory_aff_dim) (blast dest: *) have 2: "\((\a. convex hull (S - {a})) ` S) \ convex hull S" by (rule Union_least) (metis (no_types, lifting) Diff_subset hull_mono imageE) show ?thesis using True apply (simp add: segment_convex_hull frontier_def) using interior_convex_hull_eq_empty [OF assms] apply (simp add: closure_closed [OF ccs]) using "1" "2" by auto next case False then have "frontier (convex hull S) = closure (convex hull S) - interior (convex hull S)" by (simp add: rel_boundary_of_convex_hull frontier_def) also have "\ = (convex hull S) - rel_interior(convex hull S)" by (metis False aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior) also have "\ = \{convex hull (S - {a}) |a. a \ S}" proof - have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)" by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron) then show ?thesis by (simp add: False rel_frontier_convex_hull_cases) qed finally show ?thesis . qed subsection\Special case of a triangle\ proposition frontier_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "frontier(convex hull {a,b,c}) = closed_segment a b \ closed_segment b c \ closed_segment c a" (is "?lhs = ?rhs") proof (cases "b = a \ c = a \ c = b") case True then show ?thesis by (auto simp: assms segment_convex_hull frontier_def empty_interior_convex_hull insert_commute card_insert_le_m1 hull_inc insert_absorb) next case False then have [simp]: "card {a, b, c} = Suc (DIM('a))" by (simp add: card.insert_remove Set.insert_Diff_if assms) show ?thesis proof show "?lhs \ ?rhs" using False by (force simp: segment_convex_hull frontier_of_convex_hull insert_Diff_if insert_commute split: if_split_asm) show "?rhs \ ?lhs" using False apply (simp add: frontier_of_convex_hull segment_convex_hull) apply (intro conjI subsetI) apply (rule_tac X="convex hull {a,b}" in UnionI; force simp: Set.insert_Diff_if) apply (rule_tac X="convex hull {b,c}" in UnionI; force) apply (rule_tac X="convex hull {a,c}" in UnionI; force simp: insert_commute Set.insert_Diff_if) done qed qed corollary inside_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "inside (closed_segment a b \ closed_segment b c \ closed_segment c a) = interior(convex hull {a,b,c})" by (metis assms frontier_of_triangle bounded_empty bounded_insert convex_convex_hull inside_frontier_eq_interior bounded_convex_hull) corollary interior_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "interior(convex hull {a,b,c}) = convex hull {a,b,c} - (closed_segment a b \ closed_segment b c \ closed_segment c a)" using interior_subset by (force simp: frontier_of_triangle [OF assms, symmetric] frontier_def Diff_Diff_Int) subsection\Subdividing a cell complex\ lemma subdivide_interval: fixes x::real assumes "a < \x - y\" "0 < a" obtains n where "n \ \" "x < n * a \ n * a < y \ y < n * a \ n * a < x" proof - consider "a + x < y" | "a + y < x" using assms by linarith then show ?thesis proof cases case 1 let ?n = "of_int (floor (x/a)) + 1" have x: "x < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + x" apply (simp add: algebra_simps) - by (metis \0 < a\ floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right) + by (metis assms(2) floor_divide_lower mult.commute) also have "... < y" by (rule 1) finally have "?n * a < y" . with x show ?thesis using Ints_1 Ints_add Ints_of_int that by blast next case 2 let ?n = "of_int (floor (y/a)) + 1" have y: "y < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + y" apply (simp add: algebra_simps) - by (metis \0 < a\ floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right) + by (metis assms(2) floor_divide_lower mult.commute) also have "... < x" by (rule 2) finally have "?n * a < x" . then show ?thesis using Ints_1 Ints_add Ints_of_int that y by blast qed qed lemma cell_subdivision_lemma: assumes "finite \" and "\X. X \ \ \ polytope X" and "\X. X \ \ \ aff_dim X \ d" and "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and "finite I" shows "\\. \\ = \\ \ finite \ \ (\C \ \. \D. D \ \ \ C \ D) \ (\C \ \. \x \ C. \D. D \ \ \ x \ D \ D \ C) \ (\X \ \. polytope X) \ (\X \ \. aff_dim X \ d) \ (\X \ \. \Y \ \. X \ Y face_of X) \ (\X \ \. \x \ X. \y \ X. \a b. (a,b) \ I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b)" using \finite I\ proof induction case empty then show ?case by (rule_tac x="\" in exI) (auto simp: assms) next case (insert ab I) then obtain \ where eq: "\\ = \\" and "finite \" and sub1: "\C. C \ \ \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" and I: "\X x y a b. \X \ \; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" by (auto simp: that) obtain a b where "ab = (a,b)" by fastforce let ?\ = "(\X. X \ {x. a \ x \ b}) ` \ \ (\X. X \ {x. a \ x \ b}) ` \" have eqInt: "(S \ Collect P) \ (T \ Collect Q) = (S \ T) \ (Collect P \ Collect Q)" for S T::"'a set" and P Q by blast show ?case proof (intro conjI exI) show "\?\ = \\" by (force simp: eq [symmetric]) show "finite ?\" using \finite \\ by force show "\X \ ?\. polytope X" by (force simp: poly polytope_Int_polyhedron polyhedron_halfspace_le polyhedron_halfspace_ge) show "\X \ ?\. aff_dim X \ d" by (auto; metis order_trans aff aff_dim_subset inf_le1) show "\X \ ?\. \x \ X. \y \ X. \a b. (a,b) \ insert ab I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" using \ab = (a, b)\ I by fastforce show "\X \ ?\. \Y \ ?\. X \ Y face_of X" by (auto simp: eqInt halfspace_Int_eq face_of_Int_Int face face_of_halfspace_le face_of_halfspace_ge) show "\C \ ?\. \D. D \ \ \ C \ D" using sub1 by force show "\C\\. \x\C. \D. D \ ?\ \ x \ D \ D \ C" proof (intro ballI) fix C z assume "C \ \" "z \ C" with sub2 obtain D where D: "D \ \" "z \ D" "D \ C" by blast have "D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C \ D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C" using linorder_class.linear [of "a \ z" b] D by blast then show "\D. D \ ?\ \ z \ D \ D \ C" by blast qed qed qed proposition cell_complex_subdivision_exists: fixes \ :: "'a::euclidean_space set set" assumes "0 < e" "finite \" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" obtains "\'" where "finite \'" "\\' = \\" "\X. X \ \' \ diameter X < e" "\X. X \ \' \ polytope X" "\X. X \ \' \ aff_dim X \ d" "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X" "\C. C \ \' \ \D. D \ \ \ C \ D" "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" proof - have "bounded(\\)" by (simp add: \finite \\ poly bounded_Union polytope_imp_bounded) then obtain B where "B > 0" and B: "\x. x \ \\ \ norm x < B" by (meson bounded_pos_less) define C where "C \ {z \ \. \z * e / 2 / real DIM('a)\ \ B}" define I where "I \ \i \ Basis. \j \ C. { (i::'a, j * e / 2 / DIM('a)) }" have "C \ {x \ \. - B / (e / 2 / real DIM('a)) \ x \ x \ B / (e / 2 / real DIM('a))}" using \0 < e\ by (auto simp: field_split_simps C_def) then have "finite C" using finite_int_segment finite_subset by blast then have "finite I" by (simp add: I_def) obtain \' where eq: "\\' = \\" and "finite \'" and poly: "\X. X \ \' \ polytope X" and aff: "\X. X \ \' \ aff_dim X \ d" and face: "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X" and I: "\X x y a b. \X \ \'; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" and sub1: "\C. C \ \' \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" apply (rule exE [OF cell_subdivision_lemma]) using assms \finite I\ by auto show ?thesis proof (rule_tac \'="\'" in that) show "diameter X < e" if "X \ \'" for X proof - have "diameter X \ e/2" proof (rule diameter_le) show "norm (x - y) \ e / 2" if "x \ X" "y \ X" for x y proof - have "norm x < B" "norm y < B" using B \X \ \'\ eq that by blast+ have "norm (x - y) \ (\b\Basis. \(x-y) \ b\)" by (rule norm_le_l1) also have "... \ of_nat (DIM('a)) * (e / 2 / DIM('a))" proof (rule sum_bounded_above) fix i::'a assume "i \ Basis" then have I': "\z b. \z \ C; b = z * e / (2 * real DIM('a))\ \ i \ x \ b \ i \ y \ b \ i \ x \ b \ i \ y \ b" using I[of X x y] \X \ \'\ that unfolding I_def by auto show "\(x - y) \ i\ \ e / 2 / real DIM('a)" proof (rule ccontr) assume "\ \(x - y) \ i\ \ e / 2 / real DIM('a)" then have xyi: "\i \ x - i \ y\ > e / 2 / real DIM('a)" by (simp add: inner_commute inner_diff_right) obtain n where "n \ \" and n: "i \ x < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ y \ i \ y < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ x" using subdivide_interval [OF xyi] DIM_positive \0 < e\ by (auto simp: zero_less_divide_iff) have "\i \ x\ < B" by (metis \i \ Basis\ \norm x < B\ inner_commute norm_bound_Basis_lt) have "\i \ y\ < B" by (metis \i \ Basis\ \norm y < B\ inner_commute norm_bound_Basis_lt) have *: "\n * e\ \ B * (2 * real DIM('a))" if "\ix\ < B" "\iy\ < B" and ix: "ix * (2 * real DIM('a)) < n * e" and iy: "n * e < iy * (2 * real DIM('a))" for ix iy proof (rule abs_leI) have "iy * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\iy\ < B\ in linarith)+ then show "n * e \ B * (2 * real DIM('a))" using iy by linarith next have "- ix * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\ix\ < B\ in linarith)+ then show "- (n * e) \ B * (2 * real DIM('a))" using ix by linarith qed have "n \ C" using \n \ \\ n by (auto simp: C_def divide_simps intro: * \\i \ x\ < B\ \\i \ y\ < B\) show False using I' [OF \n \ C\ refl] n by auto qed qed also have "... = e / 2" by simp finally show ?thesis . qed qed (use \0 < e\ in force) also have "... < e" by (simp add: \0 < e\) finally show ?thesis . qed qed (auto simp: eq poly aff face sub1 sub2 \finite \'\) qed subsection\Simplexes\ text\The notion of n-simplex for integer \<^term>\n \ -1\\ definition\<^marker>\tag important\ simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) where "n simplex S \ \C. \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C" lemma simplex: "n simplex S \ (\C. finite C \ \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C)" by (auto simp add: simplex_def intro: aff_independent_finite) lemma simplex_convex_hull: "\ affine_dependent C \ int(card C) = n + 1 \ n simplex (convex hull C)" by (auto simp add: simplex_def) lemma convex_simplex: "n simplex S \ convex S" by (metis convex_convex_hull simplex_def) lemma compact_simplex: "n simplex S \ compact S" unfolding simplex using finite_imp_compact_convex_hull by blast lemma closed_simplex: "n simplex S \ closed S" by (simp add: compact_imp_closed compact_simplex) lemma simplex_imp_polytope: "n simplex S \ polytope S" unfolding simplex_def polytope_def using aff_independent_finite by blast lemma simplex_imp_polyhedron: "n simplex S \ polyhedron S" by (simp add: polytope_imp_polyhedron simplex_imp_polytope) lemma simplex_dim_ge: "n simplex S \ -1 \ n" by (metis (no_types, hide_lams) aff_dim_geq affine_independent_iff_card diff_add_cancel diff_diff_eq2 simplex_def) lemma simplex_empty [simp]: "n simplex {} \ n = -1" proof assume "n simplex {}" then show "n = -1" unfolding simplex by (metis card.empty convex_hull_eq_empty diff_0 diff_eq_eq of_nat_0) next assume "n = -1" then show "n simplex {}" by (fastforce simp: simplex) qed lemma simplex_minus_1 [simp]: "-1 simplex S \ S = {}" by (metis simplex cancel_comm_monoid_add_class.diff_cancel card_0_eq diff_minus_eq_add of_nat_eq_0_iff simplex_empty) lemma aff_dim_simplex: "n simplex S \ aff_dim S = n" by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card) lemma zero_simplex_sing: "0 simplex {a}" apply (simp add: simplex_def) using affine_independent_1 card_1_singleton_iff convex_hull_singleton by blast lemma simplex_sing [simp]: "n simplex {a} \ n = 0" using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast lemma simplex_zero: "0 simplex S \ (\a. S = {a})" by (metis aff_dim_eq_0 aff_dim_simplex simplex_sing) lemma one_simplex_segment: "a \ b \ 1 simplex closed_segment a b" unfolding simplex_def by (rule_tac x="{a,b}" in exI) (auto simp: segment_convex_hull) lemma simplex_segment_cases: "(if a = b then 0 else 1) simplex closed_segment a b" by (auto simp: one_simplex_segment) lemma simplex_segment: "\n. n simplex closed_segment a b" using simplex_segment_cases by metis lemma polytope_lowdim_imp_simplex: assumes "polytope P" "aff_dim P \ 1" obtains n where "n simplex P" proof (cases "P = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis by (metis assms compact_convex_collinear_segment collinear_aff_dim polytope_imp_compact polytope_imp_convex simplex_segment_cases that) qed lemma simplex_insert_dimplus1: fixes n::int assumes "n simplex S" and a: "a \ affine hull S" shows "(n+1) simplex (convex hull (insert a S))" proof - obtain C where C: "finite C" "\ affine_dependent C" "int(card C) = n+1" and S: "S = convex hull C" using assms unfolding simplex by force show ?thesis unfolding simplex proof (intro exI conjI) have "aff_dim S = n" using aff_dim_simplex assms(1) by blast moreover have "a \ affine hull C" using S a affine_hull_convex_hull by blast moreover have "a \ C" using S a hull_inc by fastforce ultimately show "\ affine_dependent (insert a C)" by (simp add: C S aff_dim_convex_hull aff_dim_insert affine_independent_iff_card) next have "a \ C" using S a hull_inc by fastforce then show "int (card (insert a C)) = n + 1 + 1" by (simp add: C) next show "convex hull insert a S = convex hull (insert a C)" by (simp add: S convex_hull_insert_segments) qed (use C in auto) qed subsection \Simplicial complexes and triangulations\ definition\<^marker>\tag important\ simplicial_complex where "simplicial_complex \ \ finite \ \ (\S \ \. \n. n simplex S) \ (\F S. S \ \ \ F face_of S \ F \ \) \ (\S S'. S \ \ \ S' \ \ \ (S \ S') face_of S)" definition\<^marker>\tag important\ triangulation where "triangulation \ \ finite \ \ (\T \ \. \n. n simplex T) \ (\T T'. T \ \ \ T' \ \ \ (T \ T') face_of T)" subsection\Refining a cell complex to a simplicial complex\ proposition convex_hull_insert_Int_eq: fixes z :: "'a :: euclidean_space" assumes z: "z \ rel_interior S" and T: "T \ rel_frontier S" and U: "U \ rel_frontier S" and "convex S" "convex T" "convex U" shows "convex hull (insert z T) \ convex hull (insert z U) = convex hull (insert z (T \ U))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof (cases "T={} \ U={}") case True then show ?thesis by auto next case False then have "T \ {}" "U \ {}" by auto have TU: "convex (T \ U)" by (simp add: \convex T\ \convex U\ convex_Int) have "(\x\T. closed_segment z x) \ (\x\U. closed_segment z x) \ (if T \ U = {} then {z} else \((closed_segment z) ` (T \ U)))" (is "_ \ ?IF") proof clarify fix x t u assume xt: "x \ closed_segment z t" and xu: "x \ closed_segment z u" and "t \ T" "u \ U" then have ne: "t \ z" "u \ z" using T U z unfolding rel_frontier_def by blast+ show "x \ ?IF" proof (cases "x = z") case True then show ?thesis by auto next case False have t: "t \ closure S" using T \t \ T\ rel_frontier_def by auto have u: "u \ closure S" using U \u \ U\ rel_frontier_def by auto show ?thesis proof (cases "t = u") case True then show ?thesis using \t \ T\ \u \ U\ xt by auto next case False have tnot: "t \ closed_segment u z" proof - have "t \ closure S - rel_interior S" using T \t \ T\ rel_frontier_def by blast then have "t \ open_segment z u" by (meson DiffD2 rel_interior_closure_convex_segment [OF \convex S\ z u] subsetD) then show ?thesis by (simp add: \t \ u\ \t \ z\ open_segment_commute open_segment_def) qed moreover have "u \ closed_segment z t" using rel_interior_closure_convex_segment [OF \convex S\ z t] \u \ U\ \u \ z\ U [unfolded rel_frontier_def] tnot by (auto simp: closed_segment_eq_open) ultimately have "\(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \ z" using that xt xu by (meson between_antisym between_mem_segment between_trans_2 ends_in_segment(2)) then have "\ collinear {t, z, u}" if "x \ z" by (auto simp: that collinear_between_cases between_commute) moreover have "collinear {t, z, x}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt) moreover have "collinear {z, x, u}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xu) ultimately have False using collinear_3_trans [of t z x u] \x \ z\ by blast then show ?thesis by metis qed qed qed then show ?thesis using False \convex T\ \convex U\ TU by (simp add: convex_hull_insert_segments hull_same split: if_split_asm) qed show "?rhs \ ?lhs" by (metis inf_greatest hull_mono inf.cobounded1 inf.cobounded2 insert_mono) qed lemma simplicial_subdivision_aux: assumes "finite \" and "\C. C \ \ \ polytope C" and "\C. C \ \ \ aff_dim C \ of_nat n" and "\C F. \C \ \; F face_of C\ \ F \ \" and "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" shows "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ \\ = \\ \ (\C \ \. \F. finite F \ F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" using assms proof (induction n arbitrary: \ rule: less_induct) case (less n) then have poly\: "\C. C \ \ \ polytope C" and aff\: "\C. C \ \ \ aff_dim C \ of_nat n" and face\: "\C F. \C \ \; F face_of C\ \ F \ \" and intface\: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" by metis+ show ?case proof (cases "n \ 1") case True have "\s. \n \ 1; s \ \\ \ \m. m simplex s" using poly\ aff\ by (force intro: polytope_lowdim_imp_simplex) then show ?thesis unfolding simplicial_complex_def using True by (rule_tac x="\" in exI) (auto simp: less.prems) next case False define \ where "\ \ {C \ \. aff_dim C < n}" have "finite \" "\C. C \ \ \ polytope C" "\C. C \ \ \ aff_dim C \ int (n - 1)" "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" using less.prems by (auto simp: \_def) moreover have \
: "\C F. \C \ \; F face_of C\ \ F \ \" using less.prems unfolding \_def by (metis (no_types, lifting) mem_Collect_eq aff_dim_subset face_of_imp_subset less_le not_le) ultimately obtain \ where "simplicial_complex \" and aff_dim\: "\K. K \ \ \ aff_dim K \ int (n - 1)" and "\\ = \\" and fin\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and C\: "\K. K \ \ \ \C. C \ \ \ K \ C" using less.IH [of "n-1" \] False by auto then have "finite \" and simpl\: "\S. S \ \ \ \n. n simplex S" and face\: "\F S. \S \ \; F face_of S\ \ F \ \" and faceI\: "\S S'. \S \ \; S' \ \\ \ (S \ S') face_of S" by (auto simp: simplicial_complex_def) define \ where "\ \ {C \ \. aff_dim C = n}" have "finite \" by (simp add: \_def less.prems(1)) have poly\: "\C. C \ \ \ polytope C" and convex\: "\C. C \ \ \ convex C" and closed\: "\C. C \ \ \ closed C" by (auto simp: \_def poly\ polytope_imp_convex polytope_imp_closed) have in_rel_interior: "(SOME z. z \ rel_interior C) \ rel_interior C" if "C \ \" for C using that poly\ polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \_def) have *: "\T. \ affine_dependent T \ card T \ n \ aff_dim K < n \ K = convex hull T" if "K \ \" for K proof - obtain r where r: "r simplex K" using \K \ \\ simpl\ by blast have "r = aff_dim K" using \r simplex K\ aff_dim_simplex by blast with r show ?thesis unfolding simplex_def using False \\K. K \ \ \ aff_dim K \ int (n - 1)\ that by fastforce qed have ahK_C_disjoint: "affine hull K \ rel_interior C = {}" if "C \ \" "K \ \" "K \ rel_frontier C" for C K proof - have "convex C" "closed C" by (auto simp: convex\ closed\ \C \ \\) obtain F where F: "F face_of C" and "F \ C" "K \ F" proof - obtain L where "L \ \" "K \ L" using \K \ \\ C\ by blast have "K \ rel_frontier C" by (simp add: \K \ rel_frontier C\) also have "... \ C" by (simp add: \closed C\ rel_frontier_def subset_iff) finally have "K \ C" . have "L \ C face_of C" using \_def \_def \C \ \\ \L \ \\ intface\ by (simp add: inf_commute) moreover have "L \ C \ C" using \C \ \\ \L \ \\ by (metis (mono_tags, lifting) \_def \_def intface\ mem_Collect_eq not_le order_refl \
) moreover have "K \ L \ C" using \C \ \\ \L \ \\ \K \ C\ \K \ L\ by (auto simp: \_def \_def) ultimately show ?thesis using that by metis qed have "affine hull F \ rel_interior C = {}" by (rule affine_hull_face_of_disjoint_rel_interior [OF \convex C\ F \F \ C\]) with hull_mono [OF \K \ F\] show "affine hull K \ rel_interior C = {}" by fastforce qed let ?\ = "(\C \ \. \K \ \ \ Pow (rel_frontier C). {convex hull (insert (SOME z. z \ rel_interior C) K)})" have "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ (\C \ \. \F. F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" proof (rule exI, intro conjI ballI) show "simplicial_complex (\ \ ?\)" unfolding simplicial_complex_def proof (intro conjI impI ballI allI) show "finite (\ \ ?\)" using \finite \\ \finite \\ by simp show "\n. n simplex S" if "S \ \ \ ?\" for S using that ahK_C_disjoint in_rel_interior simpl\ simplex_insert_dimplus1 by fastforce show "F \ \ \ ?\" if S: "S \ \ \ ?\ \ F face_of S" for F S proof - have "F \ \" if "S \ \" using S face\ that by blast moreover have "F \ \ \ ?\" if "F face_of S" "C \ \" "K \ \" and "K \ rel_frontier C" and S: "S = convex hull insert (SOME z. z \ rel_interior C) K" for C K proof - let ?z = "SOME z. z \ rel_interior C" have "?z \ rel_interior C" by (simp add: in_rel_interior \C \ \\) moreover obtain I where "\ affine_dependent I" "card I \ n" "aff_dim K < int n" "K = convex hull I" using * [OF \K \ \\] by auto ultimately have "?z \ affine hull I" using ahK_C_disjoint affine_hull_convex_hull that by blast have "compact I" "finite I" by (auto simp: \\ affine_dependent I\ aff_independent_finite finite_imp_compact) moreover have "F face_of convex hull insert ?z I" by (metis S \F face_of S\ \K = convex hull I\ convex_hull_eq_empty convex_hull_insert_segments hull_hull) ultimately obtain J where "J \ insert ?z I" "F = convex hull J" using face_of_convex_hull_subset [of "insert ?z I" F] by auto show ?thesis proof (cases "?z \ J") case True have "F \ (\K\\ \ Pow (rel_frontier C). {convex hull insert ?z K})" proof have "convex hull (J - {?z}) face_of K" by (metis True \J \ insert ?z I\ \K = convex hull I\ \\ affine_dependent I\ face_of_convex_hull_affine_independent subset_insert_iff) then have "convex hull (J - {?z}) \ \" by (rule face\ [OF \K \ \\]) moreover have "\x. x \ convex hull (J - {?z}) \ x \ rel_frontier C" by (metis True \J \ insert ?z I\ \K = convex hull I\ subsetD hull_mono subset_insert_iff that(4)) ultimately show "convex hull (J - {?z}) \ \ \ Pow (rel_frontier C)" by auto let ?F = "convex hull insert ?z (convex hull (J - {?z}))" have "F \ ?F" apply (clarsimp simp: \F = convex hull J\) by (metis True subsetD hull_mono hull_subset subset_insert_iff) moreover have "?F \ F" apply (clarsimp simp: \F = convex hull J\) by (metis (no_types, lifting) True convex_hull_eq_empty convex_hull_insert_segments hull_hull insert_Diff) ultimately show "F \ {?F}" by auto qed with \C\\\ show ?thesis by auto next case False then have "F \ \" using face_of_convex_hull_affine_independent [OF \\ affine_dependent I\] by (metis Int_absorb2 Int_insert_right_if0 \F = convex hull J\ \J \ insert ?z I\ \K = convex hull I\ face\ inf_le2 \K \ \\) then show "F \ \ \ ?\" by blast qed qed ultimately show ?thesis using that by auto qed have \
: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X \ \" "Y \ ?\" for X Y proof - obtain C K where "C \ \" "K \ \" "K \ rel_frontier C" and Y: "Y = convex hull insert (SOME z. z \ rel_interior C) K" using XY by blast have "convex C" by (simp add: \C \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_iff) let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast obtain D where "D \ \" "X \ D" using C\ \X \ \\ by blast have "D \ rel_interior C = (C \ D) \ rel_interior C" using rel_interior_subset by blast also have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) show "C \ D face_of C" using \_def \_def \C \ \\ \D \ \\ intface\ by blast show "C \ D \ C" by (metis (mono_tags, lifting) Int_lower2 \_def \_def \C \ \\ \D \ \\ aff_dim_subset mem_Collect_eq not_le) qed finally have DC: "D \ rel_interior C = {}" . have eq: "X \ convex hull (insert ?z K) = X \ convex hull K" proof (rule Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) show "disjnt X (rel_interior C)" using DC by (meson \X \ D\ disjnt_def disjnt_subset1) qed obtain I where I: "\ affine_dependent I" and Keq: "K = convex hull I" and [simp]: "convex hull K = K" using "*" \K \ \\ by force then have "?z \ affine hull I" using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull z by blast have "X \ K face_of K" by (simp add: XY(1) \K \ \\ faceI\ inf_commute) also have "... face_of convex hull insert ?z K" by (metis I Keq \?z \ affine hull I\ aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert) finally have "X \ K face_of convex hull insert ?z K" . then show ?thesis by (simp add: XY(1) Y \K \ \\ eq faceI\) qed show "S \ S' face_of S" if "S \ \ \ ?\ \ S' \ \ \ ?\" for S S' using that proof (elim conjE UnE) fix X Y assume "X \ \" and "Y \ \" then show "X \ Y face_of X" by (simp add: faceI\) next fix X Y assume XY: "X \ \" "Y \ ?\" then show "X \ Y face_of X" "Y \ X face_of Y" using \
[OF XY] by (auto simp: Int_commute) next fix X Y assume XY: "X \ ?\" "Y \ ?\" show "X \ Y face_of X" proof - obtain C K D L where "C \ \" "K \ \" "K \ rel_frontier C" and X: "X = convex hull insert (SOME z. z \ rel_interior C) K" and "D \ \" "L \ \" "L \ rel_frontier D" and Y: "Y = convex hull insert (SOME z. z \ rel_interior D) L" using XY by blast let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast have "convex C" by (simp add: \C \ \\ convex\) have "convex K" using "*" \K \ \\ by blast have "convex L" by (meson \L \ \\ convex_simplex simpl\) show ?thesis proof (cases "D=C") case True then have "L \ rel_frontier C" using \L \ rel_frontier D\ by auto have "convex hull insert (SOME z. z \ rel_interior C) (K \ L) face_of convex hull insert (SOME z. z \ rel_interior C) K" by (metis face_of_polytope_insert2 "*" IntI \C \ \\ aff_independent_finite ahK_C_disjoint empty_iff faceI\ polytope_def z \K \ \\ \L \ \\\K \ rel_frontier C\) then show ?thesis using True X Y \K \ rel_frontier C\ \L \ rel_frontier C\ \convex C\ \convex K\ \convex L\ convex_hull_insert_Int_eq z by force next case False have "convex D" by (simp add: \D \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_eq) have "L \ D" by (metis DiffE \D \ \\ \L \ rel_frontier D\ closed\ closure_closed rel_frontier_def subset_eq) let ?w = "(SOME w. w \ rel_interior D)" have w: "?w \ rel_interior D" using \D \ \\ in_rel_interior by blast have "C \ rel_interior D = (D \ C) \ rel_interior D" using rel_interior_subset by blast also have "(D \ C) \ rel_interior D = {}" proof (rule face_of_disjoint_rel_interior) show "D \ C face_of D" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \ \ aff_dim D = int n" using \_def \D \ \\ by blast moreover have "C \ \ \ aff_dim C = int n" using \_def \C \ \\ by blast ultimately show "D \ C \ D" by (metis Int_commute False face_of_aff_dim_lt inf.idem inf_le1 intface\ not_le poly\ polytope_imp_convex) qed finally have CD: "C \ (rel_interior D) = {}" . have zKC: "(convex hull insert ?z K) \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed convex\ hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z) have "disjnt (convex hull insert (SOME z. z \ rel_interior C) K) (rel_interior D)" using zKC CD by (force simp: disjnt_def) then have eq: "convex hull (insert ?z K) \ convex hull (insert ?w L) = convex hull (insert ?z K) \ convex hull L" by (rule Int_convex_hull_insert_rel_exterior [OF \convex D\ \L \ D\ w]) have ch_id: "convex hull K = K" "convex hull L = L" using "*" \K \ \\ \L \ \\ hull_same by auto have "convex C" by (simp add: \C \ \\ convex\) have "convex hull (insert ?z K) \ L = L \ convex hull (insert ?z K)" by blast also have "... = convex hull K \ L" proof (subst Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) show "C \ D face_of C" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \" "aff_dim D = int n" using \_def \D \ \\ by fastforce+ moreover have "C \ \" "aff_dim C = int n" using \_def \C \ \\ by fastforce+ ultimately have "aff_dim D + - 1 * aff_dim C \ 0" by fastforce then have "\ C face_of D" using False \convex D\ face_of_aff_dim_lt by fastforce show "C \ D \ C" by (metis inf_commute \C \ \\ \D \ \\ \\ C face_of D\ intface\) qed then have "D \ rel_interior C = {}" by (metis inf.absorb_iff2 inf_assoc inf_sup_aci(1) rel_interior_subset) then show "disjnt L (rel_interior C)" by (meson \L \ D\ disjnt_def disjnt_subset1) next show "L \ convex hull K = convex hull K \ L" by force qed finally have chKL: "convex hull (insert ?z K) \ L = convex hull K \ L" . have "convex hull insert ?z K \ convex hull L face_of K" by (simp add: \K \ \\ \L \ \\ ch_id chKL faceI\) also have "... face_of convex hull insert ?z K" proof - obtain I where I: "\ affine_dependent I" "K = convex hull I" using * [OF \K \ \\] by auto then have "\a. a \ rel_interior C \ a \ affine hull I" using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull by blast then show ?thesis by (metis I affine_independent_insert face_of_convex_hull_affine_independent hull_insert subset_insertI z) qed finally have 1: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?z K" . have "convex hull insert ?z K \ convex hull L face_of L" by (metis \K \ \\ \L \ \\ chKL ch_id faceI\ inf_commute) also have "... face_of convex hull insert ?w L" proof - obtain I where I: "\ affine_dependent I" "L = convex hull I" using * [OF \L \ \\] by auto then have "\a. a \ rel_interior D \ a \ affine hull I" using \D \ \\ \L \ \\ \L \ rel_frontier D\ affine_hull_convex_hull ahK_C_disjoint by blast then show ?thesis by (metis I aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert w) qed finally have 2: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?w L" . show ?thesis by (simp add: X Y eq 1 2) qed qed qed qed show "\F \ \ \ ?\. C = \F" if "C \ \" for C proof (cases "C \ \") case True then show ?thesis by (meson UnCI fin\ subsetD subsetI) next case False then have "C \ \" by (simp add: \_def \_def aff\ less_le that) let ?z = "SOME z. z \ rel_interior C" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast let ?F = "\K \ \ \ Pow (rel_frontier C). {convex hull (insert ?z K)}" have "?F \ ?\" using \C \ \\ by blast moreover have "C \ \?F" proof fix x assume "x \ C" have "convex C" using \C \ \\ convex\ by blast have "bounded C" using \C \ \\ by (simp add: poly\ polytope_imp_bounded that) have "polytope C" using \C \ \\ poly\ by auto have "\ (?z = x \ C = {?z})" using \C \ \\ aff_dim_sing [of ?z] \\ n \ 1\ by (force simp: \_def) then obtain y where y: "y \ rel_frontier C" and xzy: "x \ closed_segment ?z y" and sub: "open_segment ?z y \ rel_interior C" by (blast intro: segment_to_rel_frontier [OF \convex C\ \bounded C\ z \x \ C\]) then obtain F where "y \ F" "F face_of C" "F \ C" by (auto simp: rel_frontier_of_polyhedron_alt [OF polytope_imp_polyhedron [OF \polytope C\]]) then obtain \ where "finite \" "\ \ \" "F = \\" by (metis (mono_tags, lifting) \_def \C \ \\ \convex C\ aff\ face\ face_of_aff_dim_lt fin\ le_less_trans mem_Collect_eq not_less) then obtain K where "y \ K" "K \ \" using \y \ F\ by blast moreover have x: "x \ convex hull {?z,y}" using segment_convex_hull xzy by auto moreover have "convex hull {?z,y} \ convex hull insert ?z K" by (metis (full_types) \y \ K\ hull_mono empty_subsetI insertCI insert_subset) moreover have "K \ \" using \K \ \\ \\ \ \\ by blast moreover have "K \ rel_frontier C" using \F = \\\ \F \ C\ \F face_of C\ \K \ \\ face_of_subset_rel_frontier by fastforce ultimately show "x \ \?F" by force qed moreover have "convex hull insert (SOME z. z \ rel_interior C) K \ C" if "K \ \" "K \ rel_frontier C" for K proof (rule hull_minimal) show "insert (SOME z. z \ rel_interior C) K \ C" using that \C \ \\ in_rel_interior rel_interior_subset by (force simp: closure_eq rel_frontier_def closed\) show "convex C" by (simp add: \C \ \\ convex\) qed then have "\?F \ C" by auto ultimately show ?thesis by blast qed have "(\C. C \ \ \ L \ C) \ aff_dim L \ int n" if "L \ \ \ ?\" for L using that proof assume "L \ \" then show ?thesis using C\ \_def "*" by fastforce next assume "L \ ?\" then obtain C K where "C \ \" and L: "L = convex hull insert (SOME z. z \ rel_interior C) K" and K: "K \ \" "K \ rel_frontier C" by auto then have "convex hull C = C" by (meson convex\ convex_hull_eq) then have "convex C" by (metis (no_types) convex_convex_hull) have "rel_frontier C \ C" by (metis DiffE closed\ \C \ \\ closure_closed rel_frontier_def subsetI) have "K \ C" using K \rel_frontier C \ C\ by blast have "C \ \" using \_def \C \ \\ by auto moreover have "L \ C" using K L \C \ \\ by (metis \K \ C\ \convex hull C = C\ contra_subsetD hull_mono in_rel_interior insert_subset rel_interior_subset) ultimately show ?thesis using \rel_frontier C \ C\ \L \ C\ aff\ aff_dim_subset \C \ \\ dual_order.trans by blast qed then show "\C. C \ \ \ L \ C" "aff_dim L \ int n" if "L \ \ \ ?\" for L using that by auto qed then show ?thesis apply (rule ex_forward, safe) apply (meson Union_iff subsetCE, fastforce) by (meson infinite_super simplicial_complex_def) qed qed lemma simplicial_subdivision_of_cell_complex_lowdim: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" and aff: "\C. C \ \ \ aff_dim C \ d" obtains \ where "simplicial_complex \" "\K. K \ \ \ aff_dim K \ d" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof (cases "d \ 0") case True then obtain n where n: "d = of_nat n" using zero_le_imp_eq_int by blast have "\\. simplicial_complex \ \ (\K\\. aff_dim K \ int n) \ \\ = \(\C\\. {F. F face_of C}) \ (\C\\C\\. {F. F face_of C}. \F. finite F \ F \ \ \ C = \F) \ (\K\\. \C. C \ (\C\\. {F. F face_of C}) \ K \ C)" proof (rule simplicial_subdivision_aux) show "finite (\C\\. {F. F face_of C})" using \finite \\ poly polyhedron_eq_finite_faces polytope_imp_polyhedron by fastforce show "polytope F" if "F \ (\C\\. {F. F face_of C})" for F using poly that face_of_polytope_polytope by blast show "aff_dim F \ int n" if "F \ (\C\\. {F. F face_of C})" for F using that by clarify (metis n aff_dim_subset aff face_of_imp_subset order_trans) show "F \ (\C\\. {F. F face_of C})" if "G \ (\C\\. {F. F face_of C})" and "F face_of G" for F G using that face_of_trans by blast next fix F1 F2 assume "F1 \ (\C\\. {F. F face_of C})" and "F2 \ (\C\\. {F. F face_of C})" then obtain C1 C2 where "C1 \ \" "C2 \ \" and F: "F1 face_of C1" "F2 face_of C2" by auto show "F1 \ F2 face_of F1" using face_of_Int_subface [OF _ _ F] by (metis \C1 \ \\ \C2 \ \\ face inf_commute) qed moreover have "\(\C\\. {F. F face_of C}) = \\" using face_of_imp_subset face by blast ultimately show ?thesis using face_of_imp_subset n by (fastforce intro!: that simp add: poly face_of_refl polytope_imp_convex) next case False then have m1: "\C. C \ \ \ aff_dim C = -1" by (metis aff aff_dim_empty_eq aff_dim_negative_iff dual_order.trans not_less) then have face\: "\F S. \S \ \; F face_of S\ \ F \ \" by (metis aff_dim_empty face_of_empty) show ?thesis proof have "\S. S \ \ \ \n. n simplex S" by (metis (no_types) m1 aff_dim_empty simplex_minus_1) then show "simplicial_complex \" by (auto simp: simplicial_complex_def \finite \\ face intro: face\) show "aff_dim K \ d" if "K \ \" for K by (simp add: that aff) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C using \C \ \\ equals0I by auto show "\C. C \ \ \ K \ C" if "K \ \" for K using \K \ \\ by blast qed auto qed proposition simplicial_subdivision_of_cell_complex: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "simplicial_complex \" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: simplicial_subdivision_of_cell_complex_lowdim [OF assms aff_dim_le_DIM]) corollary fine_simplicial_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "simplicial_complex \" "\K. K \ \ \ diameter K < e" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof - obtain \ where \: "finite \" "\\ = \\" and diapoly: "\X. X \ \ \ diameter X < e" "\X. X \ \ \ polytope X" and "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" and \covers: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and \covered: "\C. C \ \ \ \D. D \ \ \ C \ D" by (blast intro: cell_complex_subdivision_exists [OF \0 < e\ \finite \\ poly aff_dim_le_DIM face]) then obtain \ where \: "simplicial_complex \" "\\ = \\" and \covers: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and \covered: "\K. K \ \ \ \C. C \ \ \ K \ C" using simplicial_subdivision_of_cell_complex [OF \finite \\] by metis show ?thesis proof show "simplicial_complex \" by (rule \) show "diameter K < e" if "K \ \" for K by (metis le_less_trans diapoly \covered diameter_subset polytope_imp_bounded that) show "\\ = \\" by (simp add: \(2) \\\ = \\\) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C proof - { fix x assume "x \ C" then obtain D where "D \ \" "x \ D" "D \ C" using \covers \C \ \\ \covers by force then have "\X\\ \ Pow C. x \ X" using \D \ \\ \D \ C\ \x \ D\ by blast } moreover have "finite (\ \ Pow C)" using \simplicial_complex \\ simplicial_complex_def by auto ultimately show ?thesis by (rule_tac x="(\ \ Pow C)" in exI) auto qed show "\C. C \ \ \ K \ C" if "K \ \" for K by (meson \covered \covered order_trans that) qed qed subsection\Some results on cell division with full-dimensional cells only\ lemma convex_Union_fulldim_cells: assumes "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" and eq: "\\ = U"and "convex U" shows "\{C \ \. aff_dim C = aff_dim U} = U" (is "?lhs = U") proof - have "closed U" using \finite \\ clo eq by blast have "?lhs \ U" using eq by blast moreover have "U \ ?lhs" proof (cases "\C \ \. aff_dim C = aff_dim U") case True then show ?thesis using eq by blast next case False have "closed ?lhs" by (simp add: \finite \\ clo closed_Union) moreover have "U \ closure ?lhs" proof - have "U \ closure(\{U - C |C. C \ \ \ aff_dim C < aff_dim U})" proof (rule Baire [OF \closed U\]) show "countable {U - C |C. C \ \ \ aff_dim C < aff_dim U}" using \finite \\ uncountable_infinite by fastforce have "\C. C \ \ \ openin (top_of_set U) (U-C)" by (metis Sup_upper clo closed_limpt closedin_limpt eq openin_diff openin_subtopology_self) then show "openin (top_of_set U) T \ U \ closure T" if "T \ {U - C |C. C \ \ \ aff_dim C < aff_dim U}" for T using that dense_complement_convex_closed \closed U\ \convex U\ by auto qed also have "... \ closure ?lhs" proof - obtain C where "C \ \" "aff_dim C < aff_dim U" by (metis False Sup_upper aff_dim_subset eq eq_iff not_le) have "\X. X \ \ \ aff_dim X = aff_dim U \ x \ X" if "\V. (\C. V = U - C \ C \ \ \ aff_dim C < aff_dim U) \ x \ V" for x proof - have "x \ U \ x \ \\" using \C \ \\ \aff_dim C < aff_dim U\ eq that by blast then show ?thesis by (metis Diff_iff Sup_upper Union_iff aff_dim_subset dual_order.order_iff_strict eq that) qed then show ?thesis by (auto intro!: closure_mono) qed finally show ?thesis . qed ultimately show ?thesis using closure_subset_eq by blast qed ultimately show ?thesis by blast qed proposition fine_triangular_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and aff: "\C. C \ \ \ aff_dim C = d" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "triangulation \" "\k. k \ \ \ diameter k < e" "\k. k \ \ \ aff_dim k = d" "\\ = \\" "\C. C \ \ \ \f. finite f \ f \ \ \ C = \f" "\k. k \ \ \ \C. C \ \ \ k \ C" proof - obtain \ where "simplicial_complex \" and dia\: "\K. K \ \ \ diameter K < e" and "\\ = \\" and in\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and in\: "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: fine_simplicial_subdivision_of_cell_complex [OF \e > 0\ \finite \\ poly face]) let ?\ = "{K \ \. aff_dim K = d}" show thesis proof show "triangulation ?\" using \simplicial_complex \\ by (auto simp: triangulation_def simplicial_complex_def) show "diameter L < e" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: dia\) show "aff_dim L = d" if "L \ {K \ \. aff_dim K = d}" for L using that by auto show "\F. finite F \ F \ {K \ \. aff_dim K = d} \ C = \F" if "C \ \" for C proof - obtain F where "finite F" "F \ \" "C = \F" using in\ [OF \C \ \\] by auto show ?thesis proof (intro exI conjI) show "finite {K \ F. aff_dim K = d}" by (simp add: \finite F\) show "{K \ F. aff_dim K = d} \ {K \ \. aff_dim K = d}" using \F \ \\ by blast have "d = aff_dim C" by (simp add: aff that) moreover have "\K. K \ F \ closed K \ convex K" using \simplicial_complex \\ \F \ \\ unfolding simplicial_complex_def by (metis subsetCE \F \ \\ closed_simplex convex_simplex) moreover have "convex (\F)" using \C = \F\ poly polytope_imp_convex that by blast ultimately show "C = \{K \ F. aff_dim K = d}" by (simp add: convex_Union_fulldim_cells \C = \F\ \finite F\) qed qed then show "\{K \ \. aff_dim K = d} = \\" by auto (meson in\ subsetCE) show "\C. C \ \ \ L \ C" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: in\) qed qed end diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy +++ b/src/HOL/Analysis/Starlike.thy @@ -1,6404 +1,6404 @@ (* Title: HOL/Analysis/Starlike.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) chapter \Unsorted\ theory Starlike imports Convex_Euclidean_Space Line_Segment begin lemma affine_hull_closed_segment [simp]: "affine hull (closed_segment a b) = affine hull {a,b}" by (simp add: segment_convex_hull) lemma affine_hull_open_segment [simp]: fixes a :: "'a::euclidean_space" shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})" by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) lemma rel_interior_closure_convex_segment: fixes S :: "_::euclidean_space set" assumes "convex S" "a \ rel_interior S" "b \ closure S" shows "open_segment a b \ rel_interior S" proof fix x have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u by (simp add: algebra_simps) assume "x \ open_segment a b" then show "x \ rel_interior S" unfolding closed_segment_def open_segment_def using assms by (auto intro: rel_interior_closure_convex_shrink) qed lemma convex_hull_insert_segments: "convex hull (insert a S) = (if S = {} then {a} else \x \ convex hull S. closed_segment a x)" by (force simp add: convex_hull_insert_alt in_segment) lemma Int_convex_hull_insert_rel_exterior: fixes z :: "'a::euclidean_space" assumes "convex C" "T \ C" and z: "z \ rel_interior C" and dis: "disjnt S (rel_interior C)" shows "S \ (convex hull (insert z T)) = S \ (convex hull T)" (is "?lhs = ?rhs") proof have "T = {} \ z \ S" using dis z by (auto simp add: disjnt_def) then show "?lhs \ ?rhs" proof (clarsimp simp add: convex_hull_insert_segments) fix x y assume "x \ S" and y: "y \ convex hull T" and "x \ closed_segment z y" have "y \ closure C" by (metis y \convex C\ \T \ C\ closure_subset contra_subsetD convex_hull_eq hull_mono) moreover have "x \ rel_interior C" by (meson \x \ S\ dis disjnt_iff) moreover have "x \ open_segment z y \ {z, y}" using \x \ closed_segment z y\ closed_segment_eq_open by blast ultimately show "x \ convex hull T" using rel_interior_closure_convex_segment [OF \convex C\ z] using y z by blast qed show "?rhs \ ?lhs" by (meson hull_mono inf_mono subset_insertI subset_refl) qed subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto show ?thesis unfolding mem_interior proof (intro exI subsetI conjI) fix y assume "y \ ball (x - e *\<^sub>R (x - c)) (e*d)" then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d" by simp have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "c - ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = (1 / e) *\<^sub>R (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" using \e > 0\ by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) then have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" by (simp add: dist_norm) also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "(1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x \ S" using assms(3-5) d by (intro convexD_alt [OF \convex S\]) (auto intro: convexD_alt [OF \convex S\]) with \e > 0\ show "y \ S" by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) qed (use \e>0\ \d>0\ in auto) qed lemma mem_interior_closure_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ closure S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto have "\y\S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ by force next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis using True \0 < d\ by auto next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using islimpt_approachable x by blast then have "norm (y - x) * (1 - e) < e * d" by (metis "*" dist_norm mult_imp_div_pos_le not_less) then show ?thesis using \y \ S\ by blast qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have "(1 - e) * norm (x - y) / e < d" using y \0 < e\ by (simp add: field_simps norm_minus_commute) then have "z \ interior (ball c d)" using \0 < e\ \e \ 1\ by (simp add: interior_open[OF open_ball] z_def dist_norm) then have "z \ interior S" using d interiorI interior_ball by blast then show ?thesis unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: fixes S :: "'a::euclidean_space set" assumes "convex S" and a: "a \ interior S" and b: "b \ closure S" shows "open_segment a b \ interior S" proof (clarsimp simp: in_segment) fix u::real assume u: "0 < u" "u < 1" have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "... \ interior S" using mem_interior_closure_convex_shrink [OF assms] u by simp finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \ interior S" . qed lemma convex_closure_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" and int: "interior S \ {}" shows "closure(interior S) = closure S" proof - obtain a where a: "a \ interior S" using int by auto have "closure S \ closure(interior S)" proof fix x assume x: "x \ closure S" show "x \ closure (interior S)" proof (cases "x=a") case True then show ?thesis using \a \ interior S\ closure_subset by blast next case False show ?thesis proof (clarsimp simp add: closure_def islimpt_approachable) fix e::real assume xnotS: "x \ interior S" and "0 < e" show "\x'\interior S. x' \ x \ dist x' x < e" proof (intro bexI conjI) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ x" using False \0 < e\ by (auto simp: algebra_simps min_def) show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" using \0 < e\ by (auto simp: dist_norm min_def) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ interior S" using \0 < e\ False by (auto simp add: min_def a intro: mem_interior_closure_convex_shrink [OF \convex S\ a x]) qed qed qed qed then show ?thesis by (simp add: closure_mono interior_subset subset_antisym) qed lemma closure_convex_Int_superset: fixes S :: "'a::euclidean_space set" assumes "convex S" "interior S \ {}" "interior S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(interior S)" by (simp add: convex_closure_interior assms) also have "... \ closure (S \ T)" using interior_subset [of S] assms by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) finally show ?thesis by (simp add: closure_mono dual_order.antisym) qed subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" and "0 \ S" shows "convex hull (insert 0 S) = {y. \u. (\x\S. 0 \ u x) \ sum u S \ 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof (simp add: convex_hull_finite set_eq_iff assms, safe) fix x and u :: "'a \ real" assume "0 \ u 0" "\x\S. 0 \ u x" "u 0 + sum u S = 1" then show "\v. (\x\S. 0 \ v x) \ sum v S \ 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by force next fix x and u :: "'a \ real" assume "\x\S. 0 \ u x" "sum u S \ 1" then show "\v. 0 \ v 0 \ (\x\S. 0 \ v x) \ v 0 + sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by (rule_tac x="\x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult) qed lemma substd_simplex: assumes d: "d \ Basis" shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d \ x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") proof - let ?D = d have "0 \ ?p" using assms by (auto simp: image_def) from d have "finite d" by (blast intro: finite_subset finite_Basis) show ?thesis unfolding simplex[OF \finite d\ \0 \ ?p\] proof (intro set_eqI; safe) fix u :: "'a \ real" assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" let ?x = "(\x\?D. u x *\<^sub>R x)" have ind: "\i\Basis. i \ d \ u i = ?x \ i" and notind: "(\i\Basis. i \ d \ ?x \ i = 0)" using substdbasis_expansion_unique[OF assms] by blast+ then have **: "sum u ?D = sum ((\) ?x) ?D" using assms by (auto intro!: sum.cong) show "0 \ ?x \ i" if "i \ Basis" for i using as(1) ind notind that by fastforce show "sum ((\) ?x) ?D \ 1" using "**" as(2) by linarith show "?x \ i = 0" if "i \ Basis" "i \ d" for i using notind that by blast next fix x assume "\i\Basis. 0 \ x \ i" "sum ((\) x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" with d show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" unfolding substdbasis_expansion_unique[OF assms] by (rule_tac x="inner x" in exI) auto qed qed lemma std_simplex: "convex hull (insert 0 Basis) = {x::'a::euclidean_space. (\i\Basis. 0 \ x\i) \ sum (\i. x\i) Basis \ 1}" using substd_simplex[of Basis] by auto lemma interior_std_simplex: "interior (convex hull (insert 0 Basis)) = {x::'a::euclidean_space. (\i\Basis. 0 < x\i) \ sum (\i. x\i) Basis < 1}" unfolding set_eq_iff mem_interior std_simplex proof (intro allI iffI CollectI; clarify) fix x :: 'a fix e assume "e > 0" and as: "ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" show "(\i\Basis. 0 < x \ i) \ sum ((\) x) Basis < 1" proof safe fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" using as[THEN subsetD[where c="x - (e/2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next have **: "dist x (x + (e/2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) have "\i. i \ Basis \ (x + (e/2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) then have *: "sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) have "sum ((\) x) Basis < sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force finally show "sum ((\) x) Basis < 1" by auto qed next fix x :: 'a assume as: "\i\Basis. 0 < x \ i" "sum ((\) x) Basis < 1" obtain a :: 'b where "a \ UNIV" using UNIV_witness .. let ?d = "(1 - sum ((\) x) Basis) / real (DIM('a))" show "\e>0. ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" proof (rule_tac x="min (Min (((\) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI) fix y assume y: "y \ ball x (min (Min ((\) x ` Basis)) ?d)" have "sum ((\) y) Basis \ sum (\i. x\i + ?d) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" have "\y\i - x\i\ \ norm (y - x)" by (metis Basis_le_norm i inner_commute inner_diff_right) also have "... < ?d" using y by (simp add: dist_norm norm_minus_commute) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant by (auto simp add: Suc_le_eq) finally show "sum ((\) y) Basis \ 1" . show "(\i\Basis. 0 \ y \ i)" proof safe fix i :: 'a assume i: "i \ Basis" have "norm (x - y) < Min (((\) x) ` Basis)" using y by (auto simp: dist_norm less_eq_real_def) also have "... \ x\i" using i by auto finally have "norm (x - y) < x\i" . then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] by (auto simp: inner_simps) qed next have "Min (((\) x) ` Basis) > 0" using as by simp moreover have "?d > 0" using as by (auto simp: Suc_le_eq) ultimately show "0 < min (Min ((\) x ` Basis)) ((1 - sum ((\) x) Basis) / real DIM('a))" by linarith qed qed lemma interior_std_simplex_nonempty: obtains a :: "'a::euclidean_space" where "a \ interior(convex hull (insert 0 Basis))" proof - let ?D = "Basis :: 'a set" let ?a = "sum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" { fix i :: 'a assume i: "i \ Basis" have "?a \ i = inverse (2 * real DIM('a))" by (rule trans[of _ "sum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) (simp_all add: sum.If_cases i) } note ** = this show ?thesis proof show "?a \ interior(convex hull (insert 0 Basis))" unfolding interior_std_simplex mem_Collect_eq proof safe fix i :: 'a assume i: "i \ Basis" show "0 < ?a \ i" unfolding **[OF i] by (auto simp add: Suc_le_eq) next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real DIM('a))) ?D" by (auto intro: sum.cong) also have "\ < 1" unfolding sum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto qed qed qed lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" (is "_ = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast show ?thesis proof (cases "D = {}") case True then show ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False have h0: "affine hull (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" assume x: "x \ rel_interior (convex hull (insert 0 D))" then obtain e where "e > 0" and "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 D)" using mem_rel_interior_ball[of x "convex hull (insert 0 D)"] h0 by auto then have as: "\y. \dist x y < e \ (\i\Basis. i \ D \ y\i = 0)\ \ (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" using assms by (force simp: substd_simplex) have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" then have "\j\D. 0 \ (x - (e/2) *\<^sub>R i) \ j" using D \e > 0\ x0 by (intro as[THEN conjunct1]) (force simp: dist_norm inner_simps inner_Basis) then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto then have **: "dist x (x + (e/2) *\<^sub>R a) < e" using \e > 0\ norm_Basis[of a] D by (auto simp: dist_norm) have "\i. i \ Basis \ (x + (e/2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) then have *: "sum ((\) (x + (e/2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto then have h1: "(\i\Basis. i \ D \ (x + (e/2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) have "sum ((\) x) D < sum ((\) (x + (e/2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" using ** h1 as[rule_format, of "x + (e/2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto qed } moreover { fix x :: "'a::euclidean_space" assume as: "x \ ?s" have "\i. 0 < x\i \ 0 = x\i \ 0 \ x\i" by auto moreover have "\i. i \ D \ i \ D" by auto ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis then have h2: "x \ convex hull (insert 0 D)" using as assms by (force simp add: substd_simplex) obtain a where a: "a \ D" using \D \ {}\ by auto define d where "d \ (1 - sum ((\) x) D) / real (card D)" have "\e>0. ball x e \ {x. \i\Basis. i \ D \ x \ i = 0} \ convex hull insert 0 D" unfolding substd_simplex[OF assms] proof (intro exI; safe) have "0 < card D" using \D \ {}\ \finite D\ by (simp add: card_gt_0_iff) have "Min (((\) x) ` D) > 0" using as \D \ {}\ \finite D\ by (simp) moreover have "d > 0" using as \0 < card D\ by (auto simp: d_def) ultimately show "min (Min (((\) x) ` D)) d > 0" by auto fix y :: 'a assume y2: "\i\Basis. i \ D \ y\i = 0" assume "y \ ball x (min (Min ((\) x ` D)) d)" then have y: "dist x y < min (Min ((\) x ` D)) d" by auto have "sum ((\) y) D \ sum (\i. x\i + d) D" proof (rule sum_mono) fix i assume "i \ D" with D have i: "i \ Basis" by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) also have "... < d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) finally have "\y\i - x\i\ < d" . then show "y \ i \ x \ i + d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant d_def using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . fix i :: 'a assume i: "i \ Basis" then show "0 \ y\i" proof (cases "i\D") case True have "norm (x - y) < x\i" using y Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ by (simp add: dist_norm card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (use y2 in auto) qed then have "x \ rel_interior (convex hull (insert 0 D))" using h0 h2 rel_interior_ball by force } ultimately have "\x. x \ rel_interior (convex hull insert 0 D) \ x \ {x. (\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x \ i = 0)}" by blast then show ?thesis by (rule set_eqI) qed qed lemma rel_interior_substd_simplex_nonempty: assumes "D \ {}" and "D \ Basis" obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) D" have "finite D" using assms finite_Basis infinite_super by blast then have d1: "0 < real (card D)" using \D \ {}\ by auto { fix i assume "i \ D" have "?a \ i = sum (\j. if i = j then inverse (2 * real (card D)) else 0) D" unfolding inner_sum_left using \i \ D\ by (auto simp: inner_Basis subsetD[OF assms(2)] intro: sum.cong) also have "... = inverse (2 * real (card D))" using \i \ D\ \finite D\ by auto finally have "?a \ i = inverse (2 * real (card D))" . } note ** = this show ?thesis proof show "?a \ rel_interior (convex hull (insert 0 D))" unfolding rel_interior_substd_simplex[OF assms(2)] proof safe fix i assume "i \ D" have "0 < inverse (2 * real (card D))" using d1 by auto also have "\ = ?a \ i" using **[of i] \i \ D\ by auto finally show "0 < ?a \ i" by auto next have "sum ((\) ?a) D = sum (\i. inverse (2 * real (card D))) D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) D < 1" by auto next fix i assume "i \ Basis" and "i \ D" have "?a \ span D" proof (rule span_sum[of D "(\b. b /\<^sub>R (2 * real (card D)))" D]) { fix x :: "'a::euclidean_space" assume "x \ D" then have "x \ span D" using span_base[of _ "D"] by auto then have "x /\<^sub>R (2 * real (card D)) \ span D" using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto } then show "\x. x\D \ x /\<^sub>R (2 * real (card D)) \ span D" by auto qed then show "?a \ i = 0 " using \i \ D\ unfolding span_substd_basis[OF assms(2)] using \i \ Basis\ by auto qed qed qed subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" assumes "convex S" and "0 \ S" shows "rel_interior S \ {}" proof (cases "S = {0}") case True then show ?thesis using rel_interior_sing by auto next case False obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" using basis_exists[of S] by metis then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" using subspace_span[of B] subspace_0[of "span B"] span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto obtain d and f :: "'n \ 'n" where fd: "card d = card B" "linear f" "f ` B = d" "f ` span B = {x. \i\Basis. i \ d \ x \ i = (0::real)} \ inj_on f (span B)" and d: "d \ Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto then have "bounded_linear f" using linear_conv_bounded_linear by auto have "d \ {}" using fd B \B \ {}\ by auto have "insert 0 d = f ` (insert 0 B)" using fd linear_0 by auto then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] \linear f\ by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" proof (rule rel_interior_injective_on_span_linear_image[OF \bounded_linear f\]) show "inj_on f (span (convex hull insert 0 B))" using fd * by auto qed ultimately have "rel_interior (convex hull insert 0 B) \ {}" using rel_interior_substd_simplex_nonempty[OF \d \ {}\ d] by fastforce moreover have "convex hull (insert 0 B) \ S" using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto ultimately show ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto qed lemma rel_interior_eq_empty: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior S = {} \ S = {}" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ (+) (-a) ` S" using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto then have "rel_interior ((+) (-a) ` S) \ {}" using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"] convex_translation[of S "-a"] assms by auto then have "rel_interior S \ {}" using rel_interior_translation [of "- a"] by simp } then show ?thesis by auto qed lemma interior_simplex_nonempty: fixes S :: "'N :: euclidean_space set" assumes "independent S" "finite S" "card S = DIM('N)" obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" by (simp add: rel_interior_interior) with that show ?thesis by auto qed lemma convex_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "convex (rel_interior S)" proof - { fix x y and u :: real assume assm: "x \ rel_interior S" "y \ rel_interior S" "0 \ u" "u \ 1" then have "x \ S" using rel_interior_subset by auto have "x - u *\<^sub>R (x-y) \ rel_interior S" proof (cases "0 = u") case False then have "0 < u" using assm by auto then show ?thesis using assm rel_interior_convex_shrink[of S y x u] assms \x \ S\ by auto next case True then show ?thesis using assm by auto qed then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ rel_interior S" by (simp add: algebra_simps) } then show ?thesis unfolding convex_alt by auto qed lemma convex_closure_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "closure (rel_interior S) = closure S" proof - have h1: "closure (rel_interior S) \ closure S" using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto show ?thesis proof (cases "S = {}") case False then obtain a where a: "a \ rel_interior S" using rel_interior_eq_empty assms by auto { fix x assume x: "x \ closure S" { assume "x = a" then have "x \ closure (rel_interior S)" using a unfolding closure_def by auto } moreover { assume "x \ a" { fix e :: real assume "e > 0" define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" using "*" \x \ a\ e1 by force } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto then have "x \ closure(rel_interior S)" unfolding closure_def by auto } ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto qed auto qed lemma rel_interior_same_affine_hull: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "affine hull (rel_interior S) = affine hull S" by (metis assms closure_same_affine_hull convex_closure_rel_interior) lemma rel_interior_aff_dim: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "aff_dim (rel_interior S) = aff_dim S" by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) lemma rel_interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto qed lemma rel_interior_rel_open: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_open (rel_interior S)" unfolding rel_open_def using rel_interior_rel_interior assms by auto lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e < 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" using assms by (simp add: eq_vector_fraction_iff) also have "\ = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def assms by (simp add: divide_simps vector_fraction_eq_iff) (simp add: algebra_simps) finally have "z = y - e *\<^sub>R (y-x)" by auto moreover have "e > 0" "e < 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed lemma convex_rel_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (closure S) = rel_interior S" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_eq_empty by auto next case False have "rel_interior (closure S) \ rel_interior S" using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto moreover { fix z assume z: "z \ rel_interior (closure S)" obtain x where x: "x \ rel_interior S" using \S \ {}\ assms rel_interior_eq_empty by auto have "z \ rel_interior S" proof (cases "x = z") case True then show ?thesis using x by auto next case False obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using y_def dist_norm[of z y] e by auto have "x \ affine hull closure S" using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast moreover have "z \ affine hull closure S" using z rel_interior_subset hull_subset[of "closure S"] by blast ultimately have "y \ affine hull closure S" using y_def affine_affine_hull[of "closure S"] mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto then have "y \ closure S" using e yball by auto have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" using y_def by (simp add: algebra_simps) then obtain e1 where "0 < e1" "e1 < 1" "z = y - e1 *\<^sub>R (y - x)" using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] by (auto simp add: algebra_simps) then show ?thesis using rel_interior_closure_convex_shrink assms x \y \ closure S\ by fastforce qed } ultimately show ?thesis by auto qed lemma convex_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "interior (closure S) = interior S" using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"] convex_rel_interior_closure[of S] assms by auto lemma closure_eq_rel_interior_eq: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 = rel_interior S2" by (metis convex_rel_interior_closure convex_closure_rel_interior assms) lemma closure_eq_between: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 \ S2 \ S2 \ closure S1" (is "?A \ ?B") proof assume ?A then show ?B by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset) next assume ?B then have "closure S1 \ closure S2" by (metis assms(1) convex_closure_rel_interior closure_mono) moreover from \?B\ have "closure S1 \ closure S2" by (metis closed_closure closure_minimal) ultimately show ?A .. qed lemma open_inter_closure_rel_interior: fixes S A :: "'n::euclidean_space set" assumes "convex S" and "open A" shows "A \ closure S = {} \ A \ rel_interior S = {}" by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty) lemma rel_interior_open_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False then have "open_segment a b = affine hull {a, b} \ ball ((a + b) /\<^sub>R 2) (norm (b - a) / 2)" by (simp add: open_segment_as_ball) then show ?thesis unfolding rel_interior_eq openin_open by (metis Elementary_Metric_Spaces.open_ball False affine_hull_open_segment) qed lemma rel_interior_closed_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(closed_segment a b) = (if a = b then {a} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis by simp (metis closure_open_segment convex_open_segment convex_rel_interior_closure rel_interior_open_segment) qed lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment subsection\The relative frontier of a set\ definition\<^marker>\tag important\ "rel_frontier S = closure S - rel_interior S" lemma rel_frontier_empty [simp]: "rel_frontier {} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_eq_empty: fixes S :: "'n::euclidean_space set" shows "rel_frontier S = {} \ affine S" unfolding rel_frontier_def using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric]) lemma rel_frontier_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier {a} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_affine_hull: fixes S :: "'a::euclidean_space set" shows "rel_frontier S \ affine hull S" using closure_affine_hull rel_frontier_def by fastforce lemma rel_frontier_cball [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by (force simp: sphere_def) next case equal then show ?thesis by simp next case greater then show ?thesis by simp (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) qed lemma rel_frontier_translation: fixes a :: "'a::euclidean_space" shows "rel_frontier((\x. a + x) ` S) = (\x. a + x) ` (rel_frontier S)" by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) lemma rel_frontier_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_frontier S = frontier S" by (metis frontier_def interior_rel_interior_gen rel_frontier_def) lemma rel_frontier_frontier: fixes S :: "'n::euclidean_space set" shows "affine hull S = UNIV \ rel_frontier S = frontier S" by (simp add: frontier_def rel_frontier_def rel_interior_interior) lemma closest_point_in_rel_frontier: "\closed S; S \ {}; x \ affine hull S - rel_interior S\ \ closest_point S x \ rel_frontier S" by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) lemma closed_rel_frontier [iff]: fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis proof (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) show "closedin (top_of_set (affine hull S)) (rel_frontier S)" by (simp add: "*" rel_frontier_def) qed simp qed lemma closed_rel_boundary: fixes S :: "'n::euclidean_space set" shows "closed S \ closed(S - rel_interior S)" by (metis closed_rel_frontier closure_closed rel_frontier_def) lemma compact_rel_boundary: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(S - rel_interior S)" by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) lemma bounded_rel_frontier: fixes S :: "'n::euclidean_space set" shows "bounded S \ bounded(rel_frontier S)" by (simp add: bounded_closure bounded_diff rel_frontier_def) lemma compact_rel_frontier_bounded: fixes S :: "'n::euclidean_space set" shows "bounded S \ compact(rel_frontier S)" using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast lemma compact_rel_frontier: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(rel_frontier S)" by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) lemma convex_same_rel_interior_closure: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ closure S = closure T" by (simp add: closure_eq_rel_interior_eq) lemma convex_same_rel_interior_closure_straddle: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ rel_interior S \ T \ T \ closure S" by (simp add: closure_eq_between convex_same_rel_interior_closure) lemma convex_rel_frontier_aff_dim: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" and "S2 \ {}" and "S1 \ rel_frontier S2" shows "aff_dim S1 < aff_dim S2" proof - have "S1 \ closure S2" using assms unfolding rel_frontier_def by auto then have *: "affine hull S1 \ affine hull S2" using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast then have "aff_dim S1 \ aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto moreover { assume eq: "aff_dim S1 = aff_dim S2" then have "S1 \ {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] \S2 \ {}\ by auto have **: "affine hull S1 = affine hull S2" by (simp_all add: * eq \S1 \ {}\ affine_dim_equal) obtain a where a: "a \ rel_interior S1" using \S1 \ {}\ rel_interior_eq_empty assms by auto obtain T where T: "open T" "a \ T \ S1" "T \ affine hull S1 \ S1" using mem_rel_interior[of a S1] a by auto then have "a \ T \ closure S2" using a assms unfolding rel_frontier_def by auto then obtain b where b: "b \ T \ rel_interior S2" using open_inter_closure_rel_interior[of S2 T] assms T by auto then have "b \ affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto then have "b \ S1" using T b by auto then have False using b assms unfolding rel_frontier_def by auto } ultimately show ?thesis using less_le by auto qed lemma convex_rel_interior_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "z \ rel_interior S" shows "\x\affine hull S. \m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" proof - obtain e1 where e1: "e1 > 0 \ cball z e1 \ affine hull S \ S" using mem_rel_interior_cball[of z S] assms by auto { fix x assume x: "x \ affine hull S" { assume "x \ z" define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e assume e: "e > 1 \ e \ m" have "z \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \ affine hull S" using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x by auto have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" by (simp add: algebra_simps) also have "\ = (e - 1) * norm (x-z)" using norm_scaleR e by auto also have "\ \ (m - 1) * norm (x - z)" using e mult_right_mono[of _ _ "norm(x-z)"] by auto also have "\ = (e1 / norm (x - z)) * norm (x - z)" using m_def by auto also have "\ = e1" using \x \ z\ e1 by simp finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \ e1" by auto have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \ cball z e1" using m_def ** unfolding cball_def dist_norm by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \ S" using e * e1 by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" using \m> 1 \ by auto } moreover { assume "x = z" define m where "m = 1 + e1" then have "m > 1" using e1 by auto { fix e assume e: "e > 1 \ e \ m" then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e1 x \x = z\ by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using \m > 1\ by auto } ultimately have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" by blast } then show ?thesis by auto qed lemma convex_rel_interior_if2: fixes S :: "'n::euclidean_space set" assumes "convex S" assumes "z \ rel_interior S" shows "\x\affine hull S. \e. e > 1 \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_if[of S z] assms by auto lemma convex_rel_interior_only_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" assumes "\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" shows "z \ rel_interior S" proof - obtain x where x: "x \ rel_interior S" using rel_interior_eq_empty assms by auto then have "x \ S" using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) then show ?thesis using rel_interior_convex_shrink[of S x y "1-e1"] \0 < e1 \ e1 < 1\ \y \ S\ x assms by auto qed lemma convex_rel_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S "affine"] convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_rel_interior_iff2: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\affine hull S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "z \ interior S \ (\x. \e. e > 0 \ z + e *\<^sub>R x \ S)" proof (cases "aff_dim S = int DIM('n)") case False { assume "z \ interior S" then have False using False interior_rel_interior_gen[of S] by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0 \ z + e1 *\<^sub>R (x - z) \ S" using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" by (simp add: x1_def x2_def algebra_simps) (simp add: "*" pth_8) then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * by auto have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" using x1_def x2_def by (auto simp add: algebra_simps) then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1 e2 by simp then have "x \ affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] x1 x2 z affine_affine_hull[of S] by auto } then have "affine hull S = UNIV" by auto then have "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp) then have False using False by auto } ultimately show ?thesis by auto next case True then have "S \ {}" using aff_dim_empty[of S] by auto have *: "affine hull S = UNIV" using True affine_hull_UNIV by auto { assume "z \ interior S" then have "z \ rel_interior S" using True interior_rel_interior_gen[of S] by auto then have **: "\x. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ * by auto fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" using e1 e_def by auto then have "\e. e > 0 \ z + e *\<^sub>R x \ S" by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using e1 e_def by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" by auto } then have "z \ rel_interior S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ by auto then have "z \ interior S" using True interior_rel_interior_gen[of S] by auto } ultimately show ?thesis by auto qed subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { fix S assume "S \ I" then have "y \ S" using rel_interior_subset y by auto } then have "y \ \I" by auto } then show ?thesis by auto qed lemma convex_closure_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" proof - obtain x where x: "\S\I. x \ rel_interior S" using assms by auto { fix y assume "y \ \{closure S |S. S \ I}" then have y: "\S \ I. y \ closure S" by auto { assume "y = x" then have "y \ closure (\{rel_interior S |S. S \ I})" using x closure_subset[of "\{rel_interior S |S. S \ I}"] by auto } moreover { assume "y \ x" { fix e :: real assume e: "e > 0" define e1 where "e1 = min 1 (e/norm (y - x))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (y - x) \ e" using \y \ x\ \e > 0\ le_divide_eq[of e1 e "norm (y - x)"] by simp_all define z where "z = y - e1 *\<^sub>R (y - x)" { fix S assume "S \ I" then have "z \ rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def by auto } then have *: "z \ \{rel_interior S |S. S \ I}" by auto have "\z. z \ \{rel_interior S |S. S \ I} \ z \ y \ dist z y \ e" using \y \ x\ z_def * e1 e dist_norm[of z y] by (rule_tac x="z" in exI) auto } then have "y islimpt \{rel_interior S |S. S \ I}" unfolding islimpt_approachable_le by blast then have "y \ closure (\{rel_interior S |S. S \ I})" unfolding closure_def by auto } ultimately have "y \ closure (\{rel_interior S |S. S \ I})" by auto } then show ?thesis by auto qed lemma convex_closure_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\I) = \{closure S |S. S \ I}" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_inter_rel_interior_same_closure: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\{rel_interior S |S. S \ I}) = closure (\I)" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "rel_interior (\I) \ \{rel_interior S |S. S \ I}" proof - have "convex (\I)" using assms convex_Inter by auto moreover have "convex (\{rel_interior S |S. S \ I})" using assms convex_rel_interior by (force intro: convex_Inter) ultimately have "rel_interior (\{rel_interior S |S. S \ I}) = rel_interior (\I)" using convex_inter_rel_interior_same_closure assms closure_eq_rel_interior_eq[of "\{rel_interior S |S. S \ I}" "\I"] by blast then show ?thesis using rel_interior_subset[of "\{rel_interior S |S. S \ I}"] by auto qed lemma convex_rel_interior_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" and "finite I" shows "rel_interior (\I) = \{rel_interior S |S. S \ I}" proof - have "\I \ {}" using assms rel_interior_inter_aux[of I] by auto have "convex (\I)" using convex_Inter assms by auto show ?thesis proof (cases "I = {}") case True then show ?thesis using Inter_empty rel_interior_UNIV by auto next case False { fix z assume z: "z \ \{rel_interior S |S. S \ I}" { fix x assume x: "x \ \I" { fix S assume S: "S \ I" then have "z \ rel_interior S" "x \ S" using z x by auto then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S)" using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto } then obtain mS where mS: "\S\I. mS S > 1 \ (\e. e > 1 \ e \ mS S \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" by metis define e where "e = Min (mS ` I)" then have "e \ mS ` I" using assms \I \ {}\ by simp then have "e > 1" using mS by auto moreover have "\S\I. e \ mS S" using e_def assms by auto ultimately have "\e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \ \I" using mS by auto } then have "z \ rel_interior (\I)" using convex_rel_interior_iff[of "\I" z] \\I \ {}\ \convex (\I)\ by auto } then show ?thesis using convex_rel_interior_inter[of I] assms by auto qed qed lemma convex_closure_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" assumes "rel_interior S \ rel_interior T \ {}" shows "closure (S \ T) = closure S \ closure T" using convex_closure_inter[of "{S,T}"] assms by auto lemma convex_rel_interior_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "rel_interior S \ rel_interior T \ {}" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto lemma convex_affine_closure_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "closure (S \ T) = closure S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_closure_inter_two[of S T] assms affine_imp_convex by auto qed lemma connected_component_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected_component S a b \ closed_segment a b \ S" unfolding connected_component_def by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1) ends_in_segment connected_convex_1_gen) lemma connected_component_1: fixes S :: "real set" shows "connected_component S a b \ closed_segment a b \ S" by (simp add: connected_component_1_gen) lemma convex_affine_rel_interior_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "rel_interior (S \ T) = rel_interior S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto qed lemma convex_affine_rel_frontier_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "interior S \ T \ {}" shows "rel_frontier(S \ T) = frontier S \ T" using assms unfolding rel_frontier_def frontier_def using convex_affine_closure_Int convex_affine_rel_interior_Int rel_interior_nonempty_interior by fastforce lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "interior S \ T \ {}" shows "rel_interior(S \ T) = interior S \ T" proof - obtain a where aS: "a \ interior S" and aT:"a \ T" using assms by force have "rel_interior S = interior S" by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior) then show ?thesis by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull) qed lemma closure_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "rel_interior S \ T \ {}" shows "closure(S \ T) = closure S \ T" proof have "closure (S \ T) \ closure T" by (simp add: closure_mono) also have "... \ T" by (simp add: affine_closed assms) finally show "closure(S \ T) \ closure S \ T" by (simp add: closure_mono) next obtain a where "a \ rel_interior S" "a \ T" using assms by auto then have ssT: "subspace ((\x. (-a)+x) ` T)" and "a \ S" using affine_diffs_subspace rel_interior_subset assms by blast+ show "closure S \ T \ closure (S \ T)" proof fix x assume "x \ closure S \ T" show "x \ closure (S \ T)" proof (cases "x = a") case True then show ?thesis using \a \ S\ \a \ T\ closure_subset by fastforce next case False then have "x \ closure(open_segment a x)" by auto then show ?thesis using \x \ closure S \ T\ assms convex_affine_closure_Int by blast qed qed qed lemma subset_rel_interior_convex: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "S \ closure T" and "\ S \ rel_frontier T" shows "rel_interior S \ rel_interior T" proof - have *: "S \ closure T = S" using assms by auto have "\ rel_interior S \ rel_frontier T" using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms by auto then have "rel_interior S \ rel_interior (closure T) \ {}" using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto then have "rel_interior S \ rel_interior T = rel_interior (S \ closure T)" using assms convex_closure convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto also have "\ = rel_interior S" using * by auto finally show ?thesis by auto qed lemma rel_interior_convex_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" shows "f ` (rel_interior S) = rel_interior (f ` S)" proof (cases "S = {}") case True then show ?thesis using assms by auto next case False interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" unfolding image_mono using closure_subset by auto also have "\ = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto also have "\ \ closure (f ` (rel_interior S))" using closure_linear_image_subset assms by auto finally have "closure (f ` S) = closure (f ` rel_interior S)" using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure closure_mono[of "f ` rel_interior S" "f ` S"] * by auto then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior linear_conv_bounded_linear[of f] convex_linear_image[of _ S] convex_linear_image[of _ "rel_interior S"] closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto then have "rel_interior (f ` S) \ f ` rel_interior S" using rel_interior_subset by auto moreover { fix z assume "z \ f ` rel_interior S" then obtain z1 where z1: "z1 \ rel_interior S" "f z1 = z" by auto { fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto qed lemma rel_interior_convex_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - interpret linear f by fact have "S \ {}" using assms by auto have nonemp: "f -` S \ {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) then have "S \ (range f) \ {}" by auto have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" then have z: "f z \ rel_interior S" by auto { fix x assume "x \ f -` S" then have "f x \ S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \ S" using convex_rel_interior_iff[of S "f z"] z assms \S \ {}\ by auto moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" using \linear f\ by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f -` S" using e by auto } then have "z \ rel_interior (f -` S)" using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto } moreover { fix z assume z: "z \ rel_interior (f -` S)" { fix x assume "x \ S \ range f" then obtain y where y: "f y = x" "y \ f -` S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \ f -` S" using convex_rel_interior_iff[of "f -` S" z] z conv by auto moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" using \linear f\ y by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R f z \ S \ range f" using e by auto } then have "f z \ rel_interior (S \ range f)" using \convex (S \ (range f))\ \S \ range f \ {}\ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" by auto } ultimately show ?thesis by auto qed lemma rel_interior_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" proof (cases "S = {} \ T = {}") case True then show ?thesis by auto next case False then have "S \ {}" "T \ {}" by auto then have ri: "rel_interior S \ {}" "rel_interior T \ {}" using rel_interior_eq_empty assms by auto then have "fst -` rel_interior S \ {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto then have "rel_interior ((fst :: 'n * 'm \ 'n) -` S) = fst -` rel_interior S" using linear_fst \convex S\ rel_interior_convex_linear_preimage[of fst S] by auto then have s: "rel_interior (S \ (UNIV :: 'm set)) = rel_interior S \ UNIV" by (simp add: fst_vimage_eq_Times) from ri have "snd -` rel_interior T \ {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto then have "rel_interior ((snd :: 'n * 'm \ 'm) -` T) = snd -` rel_interior T" using linear_snd \convex T\ rel_interior_convex_linear_preimage[of snd T] by auto then have t: "rel_interior ((UNIV :: 'n set) \ T) = UNIV \ rel_interior T" by (simp add: snd_vimage_eq_Times) from s t have *: "rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T) = rel_interior S \ rel_interior T" by auto have "S \ T = S \ (UNIV :: 'm set) \ (UNIV :: 'n set) \ T" by auto then have "rel_interior (S \ T) = rel_interior ((S \ (UNIV :: 'm set)) \ ((UNIV :: 'n set) \ T))" by auto also have "\ = rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T)" using * ri assms convex_Times by (subst convex_rel_interior_inter_two) auto finally show ?thesis using * by auto qed lemma rel_interior_scaleR: fixes S :: "'n::euclidean_space set" assumes "c \ 0" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S] linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms by auto lemma rel_interior_convex_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" by (metis assms linear_scaleR rel_interior_convex_linear_image) lemma convex_rel_open_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" shows "convex (((*\<^sub>R) c) ` S) \ rel_open (((*\<^sub>R) c) ` S)" by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) lemma convex_rel_open_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set) \ rel_open S" and "finite I" shows "convex (\I) \ rel_open (\I)" proof (cases "\{rel_interior S |S. S \ I} = {}") case True then have "\I = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (\I)" using assms unfolding rel_open_def using convex_rel_interior_finite_inter[of I] by auto then show ?thesis using convex_Inter assms by auto qed lemma convex_rel_open_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f ` S) \ rel_open (f ` S)" by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def) lemma convex_rel_open_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f -` S) \ rel_open (f -` S)" proof (cases "f -` (rel_interior S) = {}") case True then have "f -` S = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (f -` S)" using assms unfolding rel_open_def using rel_interior_convex_linear_preimage[of f S] by auto then show ?thesis using convex_linear_vimage assms by auto qed lemma rel_interior_projection: fixes S :: "('m::euclidean_space \ 'n::euclidean_space) set" and f :: "'m::euclidean_space \ 'n::euclidean_space set" assumes "convex S" and "f = (\y. {z. (y, z) \ S})" shows "(y, z) \ rel_interior S \ (y \ rel_interior {y. (f y \ {})} \ z \ rel_interior (f y))" proof - { fix y assume "y \ {y. f y \ {}}" then obtain z where "(y, z) \ S" using assms by auto then have "\x. x \ S \ y = fst x" by auto then obtain x where "x \ S" "y = fst x" by blast then have "y \ fst ` S" unfolding image_def by auto } then have "fst ` S = {y. f y \ {}}" unfolding fst_def using assms by auto then have h1: "fst ` rel_interior S = rel_interior {y. f y \ {}}" using rel_interior_convex_linear_image[of fst S] assms linear_fst by auto { fix y assume "y \ rel_interior {y. f y \ {}}" then have "y \ fst ` rel_interior S" using h1 by auto then have *: "rel_interior S \ fst -` {y} \ {}" by auto moreover have aff: "affine (fst -` {y})" unfolding affine_alt by (simp add: algebra_simps) ultimately have **: "rel_interior (S \ fst -` {y}) = rel_interior S \ fst -` {y}" using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto have conv: "convex (S \ fst -` {y})" using convex_Int assms aff affine_imp_convex by auto { fix x assume "x \ f y" then have "(y, x) \ S \ (fst -` {y})" using assms by auto moreover have "x = snd (y, x)" by auto ultimately have "x \ snd ` (S \ fst -` {y})" by blast } then have "snd ` (S \ fst -` {y}) = f y" using assms by auto then have ***: "rel_interior (f y) = snd ` rel_interior (S \ fst -` {y})" using rel_interior_convex_linear_image[of snd "S \ fst -` {y}"] linear_snd conv by auto { fix z assume "z \ rel_interior (f y)" then have "z \ snd ` rel_interior (S \ fst -` {y})" using *** by auto moreover have "{y} = fst ` rel_interior (S \ fst -` {y})" using * ** rel_interior_subset by auto ultimately have "(y, z) \ rel_interior (S \ fst -` {y})" by force then have "(y,z) \ rel_interior S" using ** by auto } moreover { fix z assume "(y, z) \ rel_interior S" then have "(y, z) \ rel_interior (S \ fst -` {y})" using ** by auto then have "z \ snd ` rel_interior (S \ fst -` {y})" by (metis Range_iff snd_eq_Range) then have "z \ rel_interior (f y)" using *** by auto } ultimately have "\z. (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto } then have h2: "\y z. y \ rel_interior {t. f t \ {}} \ (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto { fix y z assume asm: "(y, z) \ rel_interior S" then have "y \ fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast qed lemma rel_frontier_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_frontier S \ rel_frontier T \ rel_frontier (S \ T)" by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" assumes "cone S" shows "cone ({0} \ rel_interior S)" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_0) next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have *: "0 \ ({0} \ rel_interior S)" and "\c. c > 0 \ (*\<^sub>R) c ` ({0} \ rel_interior S) = ({0} \ rel_interior S)" by (auto simp add: rel_interior_scaleR) then show ?thesis using cone_iff[of "{0} \ rel_interior S"] by auto qed lemma rel_interior_convex_cone_aux: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) \ c > 0 \ x \ (((*\<^sub>R) c) ` (rel_interior S))" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_hull_empty) next case False then obtain s where "s \ S" by auto have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" using convex_cone_hull[of "{(1 :: real)} \ S"] conv by (subst rel_interior_projection) auto { fix y :: real assume "y \ 0" then have "y *\<^sub>R (1,s) \ cone hull ({1 :: real} \ S)" using cone_hull_expl[of "{(1 :: real)} \ S"] \s \ S\ by auto then have "f y \ {}" using f_def by auto } then have "{y. f y \ {}} = {0..}" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have **: "rel_interior {y. f y \ {}} = {0<..}" using rel_interior_real_semiline by auto { fix c :: real assume "c > 0" then have "f c = ((*\<^sub>R) c ` S)" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S" using rel_interior_convex_scaleR[of S c] assms by auto } then show ?thesis using * ** by auto qed lemma rel_interior_convex_cone: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "rel_interior (cone hull ({1 :: real} \ S)) = {(c, c *\<^sub>R x) | c x. c > 0 \ x \ rel_interior S}" (is "?lhs = ?rhs") proof - { fix z assume "z \ ?lhs" have *: "z = (fst z, snd z)" by auto then have "z \ ?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ by fastforce } moreover { fix z assume "z \ ?rhs" then have "z \ ?lhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto } ultimately show ?thesis by blast qed lemma convex_hull_finite_union: assumes "finite I" assumes "\i\I. convex (S i) \ (S i) \ {}" shows "convex hull (\(S ` I)) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)}" (is "?lhs = ?rhs") proof - have "?lhs \ ?rhs" proof fix x assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" using hull_subset[of "\(S ` I)" convex] by auto then show "x \ ?lhs" unfolding *(1)[symmetric] using * assms convex_convex_hull by (subst convex_sum) auto qed { fix i assume "i \ I" with assms have "\p. p \ S i" by auto } then obtain p where p: "\i\I. p i \ S i" by metis { fix i assume "i \ I" { fix x assume "x \ S i" define c where "c j = (if j = i then 1::real else 0)" for j then have *: "sum c I = 1" using \finite I\ \i \ I\ sum.delta[of I i "\j::'a. 1::real"] by auto define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto moreover have "(\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. s i \ S i)" using * c_def s_def p \x \ S i\ by auto ultimately have "x \ ?rhs" by force } then have "?rhs \ S i" by auto } then have *: "?rhs \ \(S ` I)" by auto { fix u v :: real assume uv: "u \ 0 \ v \ 0 \ u + v = 1" fix x y assume xy: "x \ ?rhs \ y \ ?rhs" from xy obtain c s where xc: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)" by auto from xy obtain d t where yc: "y = sum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ sum d I = 1 \ (\i\I. t i \ S i)" by auto define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "sum (\i. u * c i) I = u * sum c I" by (simp add: sum_distrib_left) moreover have "sum (\i. v * d i) I = v * sum d I" by (simp add: sum_distrib_left) ultimately have sum1: "sum e I = 1" using e_def xc yc uv by (simp add: sum.distrib) define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" for i { fix i assume i: "i \ I" have "q i \ S i" proof (cases "e i = 0") case True then show ?thesis using i p q_def by auto next case False then show ?thesis using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] assms q_def e_def i False xc yc uv by (auto simp del: mult_nonneg_nonneg) qed } then have qs: "\i\I. q i \ S i" by auto { fix i assume i: "i \ I" have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" proof (cases "e i = 0") case True have ge: "u * (c i) \ 0 \ v * d i \ 0" using xc yc uv i by simp moreover from ge have "u * c i \ 0 \ v * d i \ 0" using True e_def i by simp ultimately have "u * c i = 0 \ v * d i = 0" by auto with True show ?thesis by auto next case False then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" using q_def by auto then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) = (e i) *\<^sub>R (q i)" by auto with False show ?thesis by (simp add: algebra_simps) qed } then have *: "\i\I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" by auto have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) also have "\ = sum (\i. e i *\<^sub>R q i) I" using * by auto finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (e i) *\<^sub>R (q i)) I" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ ?rhs" using ge0 sum1 qs by auto } then have "convex ?rhs" unfolding convex_def by auto then show ?thesis using \?lhs \ ?rhs\ * hull_minimal[of "\(S ` I)" ?rhs convex] by blast qed lemma convex_hull_union_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "S \ {}" and "convex T" and "T \ {}" shows "convex hull (S \ T) = {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof define I :: "nat set" where "I = {1, 2}" define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(s ` I) = {\ i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)}" using assms s_def I_def by (subst convex_hull_finite_union) auto moreover have "{\i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)} \ ?rhs" using s_def I_def by auto ultimately show "?lhs \ ?rhs" by auto { fix x assume "x \ ?rhs" then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \ u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T" by auto then have "x \ convex hull {s, t}" using convex_hull_2[of s t] by auto then have "x \ convex hull (S \ T)" using * hull_mono[of "{s, t}" "S \ T"] by auto } then show "?lhs \ ?rhs" by blast qed proposition ray_to_rel_frontier: fixes a :: "'a::real_inner" assumes "bounded S" and a: "a \ rel_interior S" and aff: "(a + l) \ affine hull S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ rel_frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ rel_interior S" proof - have aaff: "a \ affine hull S" by (meson a hull_subset rel_interior_subset rev_subsetD) let ?D = "{d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" obtain B where "B > 0" and B: "S \ ball a B" using bounded_subset_ballD [OF \bounded S\] by blast have "a + (B / norm l) *\<^sub>R l \ ball a B" by (simp add: dist_norm \l \ 0\) with B have "a + (B / norm l) *\<^sub>R l \ rel_interior S" using rel_interior_subset subsetCE by blast with \B > 0\ \l \ 0\ have nonMT: "?D \ {}" using divide_pos_pos zero_less_norm_iff by fastforce have bdd: "bdd_below ?D" by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq) have relin_Ex: "\x. x \ rel_interior S \ \e>0. \x'\affine hull S. dist x' x < e \ x' \ rel_interior S" using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff) define d where "d = Inf ?D" obtain \ where "0 < \" and \: "\\. \0 \ \; \ < \\ \ (a + \ *\<^sub>R l) \ rel_interior S" proof - obtain e where "e>0" and e: "\x'. x' \ affine hull S \ dist x' a < e \ x' \ rel_interior S" using relin_Ex a by blast show thesis proof (rule_tac \ = "e / norm l" in that) show "0 < e / norm l" by (simp add: \0 < e\ \l \ 0\) next show "a + \ *\<^sub>R l \ rel_interior S" if "0 \ \" "\ < e / norm l" for \ proof (rule e) show "a + \ *\<^sub>R l \ affine hull S" by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show "dist (a + \ *\<^sub>R l) a < e" using that by (simp add: \l \ 0\ dist_norm pos_less_divide_eq) qed qed qed have inint: "\e. \0 \ e; e < d\ \ a + e *\<^sub>R l \ rel_interior S" unfolding d_def using cInf_lower [OF _ bdd] by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left) have "\ \ d" unfolding d_def using \ dual_order.strict_implies_order le_less_linear by (blast intro: cInf_greatest [OF nonMT]) with \0 < \\ have "0 < d" by simp have "a + d *\<^sub>R l \ rel_interior S" proof assume adl: "a + d *\<^sub>R l \ rel_interior S" obtain e where "e > 0" and e: "\x'. x' \ affine hull S \ dist x' (a + d *\<^sub>R l) < e \ x' \ rel_interior S" using relin_Ex adl by blast have "d + e / norm l \ Inf {d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" proof (rule cInf_greatest [OF nonMT], clarsimp) fix x::real assume "0 < x" and nonrel: "a + x *\<^sub>R l \ rel_interior S" show "d + e / norm l \ x" proof (cases "x < d") case True with inint nonrel \0 < x\ show ?thesis by auto next case False then have dle: "x < d + e / norm l \ dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e" by (simp add: field_simps \l \ 0\) have ain: "a + x *\<^sub>R l \ affine hull S" by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show ?thesis using e [OF ain] nonrel dle by force qed qed then show False using \0 < e\ \l \ 0\ by (simp add: d_def [symmetric] field_simps) qed moreover have "a + d *\<^sub>R l \ closure S" proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" proof (rule subsetD [OF rel_interior_subset inint]) show "d - min d (\ / 2 / norm l) < d" using \l \ 0\ \0 < d\ \0 < \\ by auto qed auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" using 1 2 \0 < d\ \0 < \\ by (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) (auto simp: algebra_simps) qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) show ?thesis by (rule that [OF \0 < d\ infront inint]) qed corollary ray_to_frontier: fixes a :: "'a::euclidean_space" assumes "bounded S" and a: "a \ interior S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - have \
: "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp moreover have "a + l \ affine hull S" using a affine_hull_nonempty_interior by blast ultimately show thesis by (metis \
\bounded S\ \l \ 0\ frontier_def ray_to_rel_frontier rel_frontier_def that) qed lemma segment_to_rel_frontier_aux: fixes x :: "'a::euclidean_space" assumes "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "x \ y" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof - have "x + (y - x) \ affine hull S" using hull_inc [OF y] by auto then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \ rel_frontier S" and di: "\e. \0 \ e; e < d\ \ (x + e *\<^sub>R (y-x)) \ rel_interior S" by (rule ray_to_rel_frontier [OF \bounded S\ x]) (use xy in auto) show ?thesis proof show "x + d *\<^sub>R (y - x) \ rel_frontier S" by (simp add: df) next have "open_segment x y \ rel_interior S" using rel_interior_closure_convex_segment [OF \convex S\ x] closure_subset y by blast moreover have "x + d *\<^sub>R (y - x) \ open_segment x y" if "d < 1" using xy \0 < d\ that by (force simp: in_segment algebra_simps) ultimately have "1 \ d" using df rel_frontier_def by fastforce moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" unfolding in_segment by (rule_tac x="1/d" in exI) (auto simp: algebra_simps) next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" proof (rule rel_interior_closure_convex_segment [OF \convex S\ x]) show "x + d *\<^sub>R (y - x) \ closure S" using df rel_frontier_def by auto qed qed qed lemma segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") case True with xy have "S \ {x}" by blast with True show ?thesis by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y) next case False then show ?thesis using segment_to_rel_frontier_aux [OF S x y] that by blast qed proposition rel_frontier_not_sing: fixes a :: "'a::euclidean_space" assumes "bounded S" shows "rel_frontier S \ {a}" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain z where "z \ S" by blast then show ?thesis proof (cases "S = {z}") case True then show ?thesis by simp next case False then obtain w where "w \ S" "w \ z" using \z \ S\ by blast show ?thesis proof assume "rel_frontier S = {a}" then consider "w \ rel_frontier S" | "z \ rel_frontier S" using \w \ z\ by auto then show False proof cases case 1 then have w: "w \ rel_interior S" using \w \ S\ closure_subset rel_frontier_def by fastforce have "w + (w - z) \ affine hull S" by (metis \w \ S\ \z \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \ rel_frontier S" using \w \ z\ \z \ S\ by (metis assms ray_to_rel_frontier right_minus_eq w) moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ w, of "1 *\<^sub>R (z - w)"] \w \ z\ \z \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) next case 2 then have z: "z \ rel_interior S" using \z \ S\ closure_subset rel_frontier_def by fastforce have "z + (z - w) \ affine hull S" by (metis \z \ S\ \w \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \ rel_frontier S" using \w \ z\ \w \ S\ by (metis assms ray_to_rel_frontier right_minus_eq z) moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ z, of "1 *\<^sub>R (w - z)"] \w \ z\ \w \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) qed qed qed qed subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" shows "closure S + closure T \ closure (S + T)" unfolding set_plus_image closure_Times [symmetric] split_def by (intro closure_bounded_linear_image_subset bounded_linear_add bounded_linear_fst bounded_linear_snd) lemma rel_interior_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S + T) = rel_interior S + rel_interior T" proof - have "rel_interior S + rel_interior T = (\(x,y). x + y) ` (rel_interior S \ rel_interior T)" by (simp add: set_plus_image) also have "\ = (\(x,y). x + y) ` rel_interior (S \ T)" using rel_interior_Times assms by auto also have "\ = rel_interior (S + T)" using fst_snd_linear convex_Times assms rel_interior_convex_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp add: set_plus_image) finally show ?thesis .. qed lemma rel_interior_sum_gen: fixes S :: "'a \ 'n::euclidean_space set" assumes "\i. i\I \ convex (S i)" shows "rel_interior (sum S I) = sum (\i. rel_interior (S i)) I" using rel_interior_sum rel_interior_sing[of "0"] assms by (subst sum_set_cond_linear[of convex], auto simp add: convex_set_plus) lemma convex_rel_open_direct_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S \ T) \ rel_open (S \ T)" by (metis assms convex_Times rel_interior_Times rel_open_def) lemma convex_rel_open_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S + T) \ rel_open (S + T)" by (metis assms convex_set_plus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: assumes "finite I" and "I \ {}" assumes "\i. i\I \ convex (S i) \ cone (S i) \ S i \ {}" shows "convex hull (\(S ` I)) = sum S I" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?lhs" then obtain c xs where x: "x = sum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto define s where "s i = c i *\<^sub>R xs i" for i have "\i\I. s i \ S i" using s_def x assms by (simp add: mem_cone) moreover have "x = sum s I" using x s_def by auto ultimately have "x \ ?rhs" using set_sum_alt[of I S] assms by auto } moreover { fix x assume "x \ ?rhs" then obtain s where x: "x = sum s I \ (\i\I. s i \ S i)" using set_sum_alt[of I S] assms by auto define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = sum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" using x xs_def assms by (simp add: cone_def) moreover have "\i\I. (1 :: real) / of_nat (card I) \ 0" by auto moreover have "sum (\i. (1 :: real) / of_nat (card I)) I = 1" using assms by auto ultimately have "x \ ?lhs" using assms apply (simp add: convex_hull_finite_union[of I S]) by (rule_tac x = "(\i. 1 / (card I))" in exI) auto } ultimately show ?thesis by auto qed lemma convex_hull_union_cones_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "cone S" and "S \ {}" assumes "convex T" and "cone T" and "T \ {}" shows "convex hull (S \ T) = S + T" proof - define I :: "nat set" where "I = {1, 2}" define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(A ` I) = sum A I" using A_def I_def by (metis assms convex_hull_finite_union_cones empty_iff finite.emptyI finite.insertI insertI1) moreover have "sum A I = S + T" using A_def I_def by (force simp add: set_plus_def) ultimately show ?thesis by auto qed lemma rel_interior_convex_hull_union: fixes S :: "'a \ 'n::euclidean_space set" assumes "finite I" and "\i\I. convex (S i) \ S i \ {}" shows "rel_interior (convex hull (\(S ` I))) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior(S i))}" (is "?lhs = ?rhs") proof (cases "I = {}") case True then show ?thesis using convex_hull_empty by auto next case False define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto define K0 where "K0 = cone hull ({1 :: real} \ C0)" define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) have convK: "\i\I. convex (K i)" unfolding K_def by (simp add: assms(2) convex_Times convex_cone_hull) have "K0 \ K i" if "i \ I" for i unfolding K0_def K_def by (simp add: Sigma_mono \\i\I. S i \ C0\ hull_mono that) then have "K0 \ \(K ` I)" by auto moreover have "convex K0" unfolding K0_def by (simp add: C0_def convex_Times convex_cone_hull) ultimately have geq: "K0 \ convex hull (\(K ` I))" using hull_minimal[of _ "K0" "convex"] by blast have "\i\I. K i \ {1 :: real} \ S i" using K_def by (simp add: hull_subset) then have "\(K ` I) \ {1 :: real} \ \(S ` I)" by auto then have "convex hull \(K ` I) \ convex hull ({1 :: real} \ \(S ` I))" by (simp add: hull_mono) then have "convex hull \(K ` I) \ {1 :: real} \ C0" unfolding C0_def using convex_hull_Times[of "{(1 :: real)}" "\(S ` I)"] convex_hull_singleton by auto moreover have "cone (convex hull (\(K ` I)))" by (simp add: K_def cone_Union cone_cone_hull cone_convex_hull) ultimately have "convex hull (\(K ` I)) \ K0" unfolding K0_def using hull_minimal[of _ "convex hull (\(K ` I))" "cone"] by blast then have "K0 = convex hull (\(K ` I))" using geq by auto also have "\ = sum K I" using assms False \\i\I. K i \ {}\ cone_hull_eq convK by (intro convex_hull_finite_union_cones; fastforce simp: K_def) finally have "K0 = sum K I" by auto then have *: "rel_interior K0 = sum (\i. (rel_interior (K i))) I" using rel_interior_sum_gen[of I K] convK by auto { fix x assume "x \ ?lhs" then have "(1::real, x) \ rel_interior K0" using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull by auto then obtain k where k: "(1::real, x) = sum k I \ (\i\I. k i \ rel_interior (K i))" using \finite I\ * set_sum_alt[of I "\i. rel_interior (K i)"] by auto { fix i assume "i \ I" then have "convex (S i) \ k i \ rel_interior (cone hull {1} \ S i)" using k K_def assms by auto then have "\ci si. k i = (ci, ci *\<^sub>R si) \ 0 < ci \ si \ rel_interior (S i)" using rel_interior_convex_cone[of "S i"] by auto } then obtain c s where cs: "\i\I. k i = (c i, c i *\<^sub>R s i) \ 0 < c i \ s i \ rel_interior (S i)" by metis then have "x = (\i\I. c i *\<^sub>R s i) \ sum c I = 1" using k by (simp add: sum_prod) then have "x \ ?rhs" using k cs by auto } moreover { fix x assume "x \ ?rhs" then obtain c s where cs: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto } then have "(1, x) \ rel_interior K0" using * set_sum_alt[of I "(\i. rel_interior (K i))"] assms cs by (simp add: k_def) (metis (mono_tags, lifting) sum_prod) then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto } ultimately show ?thesis by blast qed lemma convex_le_Inf_differential: fixes f :: "real \ real" assumes "convex_on I f" and "x \ interior I" and "y \ I" shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" (is "_ \ _ + Inf (?F x) * (y - x)") proof (cases rule: linorder_cases) assume "x < y" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) then have "K \ I" using \interior I \ I\ e(2) by blast have "Inf (?F x) \ (f x - f y) / (x - y)" proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using \t \ I\ \x < t\ by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" using \convex_on I f\ \x \ I\ \y \ I\ \x < t\ \t < y\ by (rule convex_on_diff) next fix y assume "y \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \K \ I\ _ \K < x\ _]] show "(f K - f x) / (K - x) \ y" by auto qed then show ?thesis using \x < y\ by (simp add: field_simps) next assume "y < x" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using \y < x\ by (auto simp: field_simps) also fix z assume "z \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \y \ I\ _ \y < x\]] have "(f y - f x) / (y - x) \ z" by auto finally show "(f x - f y) / (x - y) \ z" . next have "x + e / 2 \ ball x e" using e by (auto simp: dist_real_def) with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto then show "?F x \ {}" by blast qed then show ?thesis using \y < x\ by (simp add: field_simps) qed simp subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma affine_independent_convex_affine_hull: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" "T \ S" shows "convex hull T = affine hull T \ convex hull S" proof - have fin: "finite S" "finite T" using assms aff_independent_finite finite_subset by auto have "convex hull T \ affine hull T" using convex_hull_subset_affine_hull by blast moreover have "convex hull T \ convex hull S" using assms hull_mono by blast moreover have "affine hull T \ convex hull S \ convex hull T" proof - have 0: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" using affine_dependent_explicit_finite assms(1) fin(1) by auto show ?thesis proof (clarsimp simp add: affine_hull_finite fin) fix u assume S: "(\v\T. u v *\<^sub>R v) \ convex hull S" and T1: "sum u T = 1" then obtain v where v: "\x\S. 0 \ v x" "sum v S = 1" "(\x\S. v x *\<^sub>R x) = (\v\T. u v *\<^sub>R v)" by (auto simp add: convex_hull_finite fin) { fix x assume"x \ T" then have S: "S = (S - T) \ T" \ \split into separate cases\ using assms by auto have [simp]: "(\x\T. v x *\<^sub>R x) + (\x\S - T. v x *\<^sub>R x) = (\x\T. u x *\<^sub>R x)" "sum v T + sum v (S - T) = 1" using v fin S by (auto simp: sum.union_disjoint [symmetric] Un_commute) have "(\x\S. if x \ T then v x - u x else v x) = 0" "(\x\S. (if x \ T then v x - u x else v x) *\<^sub>R x) = 0" using v fin T1 by (subst S, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ } note [simp] = this have "(\x\T. 0 \ u x)" using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ v(1) by fastforce then show "(\v\T. u v *\<^sub>R v) \ convex hull T" using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ T1 by (fastforce simp add: convex_hull_finite fin) qed qed ultimately show ?thesis by blast qed lemma affine_independent_span_eq: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" "card S = Suc (DIM ('a))" shows "affine hull S = UNIV" proof (cases "S = {}") case True then show ?thesis using assms by simp next case False then obtain a T where T: "a \ T" "S = insert a T" by blast then have fin: "finite T" using assms by (metis finite_insert aff_independent_finite) have "UNIV \ (+) a ` span ((\x. x - a) ` T)" proof (intro card_ge_dim_independent Fun.vimage_subsetD) show "independent ((\x. x - a) ` T)" using T affine_dependent_iff_dependent assms(1) by auto show "dim ((+) a -` UNIV) \ card ((\x. x - a) ` T)" using assms T fin by (auto simp: card_image inj_on_def) qed (use surj_plus in auto) then show ?thesis using T(2) affine_hull_insert_span_gen equalityI by fastforce qed lemma affine_independent_span_gt: fixes S :: "'a::euclidean_space set" assumes ind: "\ affine_dependent S" and dim: "DIM ('a) < card S" shows "affine hull S = UNIV" proof (intro affine_independent_span_eq [OF ind] antisym) show "card S \ Suc DIM('a)" using aff_independent_finite affine_dependent_biggerset ind by fastforce show "Suc DIM('a) \ card S" using Suc_leI dim by blast qed lemma empty_interior_affine_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" and dim: "card S \ DIM ('a)" shows "interior(affine hull S) = {}" using assms proof (induct S rule: finite_induct) case (insert x S) then have "dim (span ((\y. y - x) ` S)) < DIM('a)" by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) then show ?case by (simp add: empty_interior_lowdim affine_hull_insert_span_gen interior_translation) qed auto lemma empty_interior_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" and dim: "card S \ DIM ('a)" shows "interior(convex hull S) = {}" by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull interior_mono empty_interior_affine_hull [OF assms]) lemma explicit_subset_rel_interior_convex_hull: fixes S :: "'a::euclidean_space set" shows "finite S \ {y. \u. (\x \ S. 0 < u x \ u x < 1) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y} \ rel_interior (convex hull S)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=S, simplified]) lemma explicit_subset_rel_interior_convex_hull_minimal: fixes S :: "'a::euclidean_space set" shows "finite S \ {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y} \ rel_interior (convex hull S)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=S, simplified]) lemma rel_interior_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_interior(convex hull S) = {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" proof (cases "\a. S = {a}") case True then show "?lhs \ ?rhs" by force next case False have fs: "finite S" using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ S" "b \ S" "a \ b" then have S: "S = (S - {a,b}) \ {a,b}" \ \split into separate cases\ by auto have "(\x\S. if x = a then - d else if x = b then d else 0) = 0" "(\x\S. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs by (subst S, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y assume y: "y \ convex hull S" "y \ ?rhs" have *: False if ua: "\x\S. 0 \ u x" "sum u S = 1" "\ 0 < u a" "a \ S" and yT: "y = (\x\S. u x *\<^sub>R x)" "y \ T" "open T" and sb: "T \ affine hull S \ {w. \u. (\x\S. 0 \ u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = w}" for u T a proof - have ua0: "u a = 0" using ua by auto obtain b where b: "b\S" "a \ b" using ua False by auto obtain e where e: "0 < e" "ball (\x\S. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) have "(\x\S. u x *\<^sub>R x) \ affine hull S" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) then have "(\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull S" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) then have "y - d *\<^sub>R (a - b) \ T \ affine hull S" using d e yT by auto then obtain v where v: "\x\S. 0 \ v x" "sum v S = 1" "(\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto have aff: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" using assms by (simp add: affine_dependent_explicit_finite fs) show False using ua b d v aff [of "\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)"] by (auto simp: algebra_simps sum_subtractf sum.distrib) qed have "y \ rel_interior (convex hull S)" using y apply (simp add: mem_rel_interior) apply (auto simp: convex_hull_finite [OF fs]) apply (drule_tac x=u in spec) apply (auto intro: *) done } with rel_interior_subset show "?lhs \ ?rhs" by blast qed qed lemma interior_convex_hull_explicit_minimal: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "interior(convex hull S) = (if card(S) \ DIM('a) then {} else {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" (is "_ = (if _ then _ else ?rhs)") proof (clarsimp simp: aff_independent_finite empty_interior_convex_hull assms) assume S: "\ card S \ DIM('a)" have "interior (convex hull S) = rel_interior(convex hull S)" using assms S by (simp add: affine_independent_span_gt rel_interior_interior) then show "interior(convex hull S) = ?rhs" by (simp add: assms S rel_interior_convex_hull_explicit) qed lemma interior_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "interior(convex hull S) = (if card(S) \ DIM('a) then {} else {y. \u. (\x \ S. 0 < u x \ u x < 1) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a assume "card Basis < card S" and u: "\x. x\S \ 0 < u x" "sum u S = 1" and a: "a \ S" then have cs: "Suc 0 < card S" by (metis DIM_positive less_trans_Suc) obtain b where b: "b \ S" "a \ b" proof (cases "S \ {a}") case True then show thesis using cs subset_singletonD by fastforce qed blast have "u a + u b \ sum u {a,b}" using a b by simp also have "... \ sum u S" using a b u by (intro Groups_Big.sum_mono2) (auto simp: less_imp_le aff_independent_finite assms) finally have "u a < 1" using \b \ S\ u by fastforce } note [simp] = this show ?thesis using assms by (force simp add: not_le interior_convex_hull_explicit_minimal) qed lemma interior_closed_segment_ge2: fixes a :: "'a::euclidean_space" assumes "2 \ DIM('a)" shows "interior(closed_segment a b) = {}" using assms unfolding segment_convex_hull proof - have "card {a, b} \ DIM('a)" using assms by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) then show "interior (convex hull {a, b}) = {}" by (metis empty_interior_convex_hull finite.insertI finite.emptyI) qed lemma interior_open_segment: fixes a :: "'a::euclidean_space" shows "interior(open_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (simp add: not_le, intro conjI impI) assume "2 \ DIM('a)" then show "interior (open_segment a b) = {}" using interior_closed_segment_ge2 interior_mono segment_open_subset_closed by blast next assume le2: "DIM('a) < 2" show "interior (open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False with le2 have "affine hull (open_segment a b) = UNIV" by (simp add: False affine_independent_span_gt) then show "interior (open_segment a b) = open_segment a b" using rel_interior_interior rel_interior_open_segment by blast qed qed lemma interior_closed_segment: fixes a :: "'a::euclidean_space" shows "interior(closed_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then have "closure (open_segment a b) = closed_segment a b" by simp then show ?thesis by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) qed lemmas interior_segment = interior_closed_segment interior_open_segment lemma closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b = closed_segment c d \ {a,b} = {c,d}" proof assume abcd: "closed_segment a b = closed_segment c d" show "{a,b} = {c,d}" proof (cases "a=b \ c=d") case True with abcd show ?thesis by force next case False then have neq: "a \ b \ c \ d" by force have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)" using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) have "b \ {c, d}" proof - have "insert b (closed_segment c d) = closed_segment c d" using abcd by blast then show ?thesis by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) qed moreover have "a \ {c, d}" by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) ultimately show "{a, b} = {c, d}" using neq by fastforce qed next assume "{a,b} = {c,d}" then show "closed_segment a b = closed_segment c d" by (simp add: segment_convex_hull) qed lemma closed_open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b \ open_segment c d" by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) lemma open_closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d" using closed_open_segment_eq by blast lemma open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b = open_segment c d \ a = b \ c = d \ {a,b} = {c,d}" (is "?lhs = ?rhs") proof assume abcd: ?lhs show ?rhs proof (cases "a=b \ c=d") case True with abcd show ?thesis using finite_open_segment by fastforce next case False then have a2: "a \ b \ c \ d" by force with abcd show ?rhs unfolding open_segment_def by (metis (no_types) abcd closed_segment_eq closure_open_segment) qed next assume ?rhs then show ?lhs by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) qed subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes S :: "'a::euclidean_space set" shows "compact S ==> closure(convex hull S) = convex hull S" by (simp add: compact_imp_closed compact_convex_hull) lemma rel_frontier_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_frontier(convex hull S) = {y. \u. (\x \ S. 0 \ u x) \ (\x \ S. u x = 0) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) have "\u y v. \y \ S; u y = 0; sum u S = 1; \x\S. 0 < v x; sum v S = 1; (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)\ \ \u. sum u S = 0 \ (\v\S. u v \ 0) \ (\v\S. u v *\<^sub>R v) = 0" apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) done then show ?thesis using fs assms apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit) apply (auto simp: convex_hull_finite) apply (metis less_eq_real_def) by (simp add: affine_dependent_explicit_finite) qed lemma frontier_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = {y. \u. (\x \ S. 0 \ u x) \ (DIM ('a) < card S \ (\x \ S. u x = 0)) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) show ?thesis proof (cases "DIM ('a) < card S") case True with assms fs show ?thesis by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) next case False then have "card S \ DIM ('a)" by linarith then show ?thesis using assms fs apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) apply (simp add: convex_hull_finite) done qed qed lemma rel_frontier_convex_hull_cases: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_frontier(convex hull S) = \{convex hull (S - {x}) |x. x \ S}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) { fix u a have "\x\S. 0 \ u x \ a \ S \ u a = 0 \ sum u S = 1 \ \x v. x \ S \ (\x\S - {x}. 0 \ v x) \ sum v (S - {x}) = 1 \ (\x\S - {x}. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" apply (rule_tac x=a in exI) apply (rule_tac x=u in exI) apply (simp add: Groups_Big.sum_diff1 fs) done } moreover { fix a u have "a \ S \ \x\S - {a}. 0 \ u x \ sum u (S - {a}) = 1 \ \v. (\x\S. 0 \ v x) \ (\x\S. v x = 0) \ sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S - {a}. u x *\<^sub>R x)" apply (rule_tac x="\x. if x = a then 0 else u x" in exI) apply (auto simp: sum.If_cases Diff_eq if_smult fs) done } ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) apply (auto simp add: convex_hull_finite fs Union_SetCompr_eq) done qed lemma frontier_convex_hull_eq_rel_frontier: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = (if card S \ DIM ('a) then convex hull S else rel_frontier(convex hull S))" using assms unfolding rel_frontier_def frontier_def by (simp add: affine_independent_span_gt rel_interior_interior finite_imp_compact empty_interior_convex_hull aff_independent_finite) lemma frontier_convex_hull_cases: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = (if card S \ DIM ('a) then convex hull S else \{convex hull (S - {x}) |x. x \ S})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) lemma in_frontier_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" "card S \ Suc (DIM ('a))" "x \ S" shows "x \ frontier(convex hull S)" proof (cases "affine_dependent S") case True with assms obtain y where "y \ S" and y: "y \ affine hull (S - {y})" by (auto simp: affine_dependent_def) moreover have "x \ closure (convex hull S)" by (meson closure_subset hull_inc subset_eq \x \ S\) moreover have "x \ interior (convex hull S)" using assms by (metis Suc_mono affine_hull_convex_hull affine_hull_nonempty_interior \y \ S\ y card.remove empty_iff empty_interior_affine_hull finite_Diff hull_redundant insert_Diff interior_UNIV not_less) ultimately show ?thesis unfolding frontier_def by blast next case False { assume "card S = Suc (card Basis)" then have cs: "Suc 0 < card S" by (simp) with subset_singletonD have "\y \ S. y \ x" by (cases "S \ {x}") fastforce+ } note [dest!] = this show ?thesis using assms unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq by (auto simp: le_Suc_eq hull_inc) qed lemma not_in_interior_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" "card S \ Suc (DIM ('a))" "x \ S" shows "x \ interior(convex hull S)" using in_frontier_convex_hull [OF assms] by (metis Diff_iff frontier_def) lemma interior_convex_hull_eq_empty: fixes S :: "'a::euclidean_space set" assumes "card S = Suc (DIM ('a))" shows "interior(convex hull S) = {} \ affine_dependent S" proof show "affine_dependent S \ interior (convex hull S) = {}" proof (clarsimp simp: affine_dependent_def) fix a b assume "b \ S" "b \ affine hull (S - {b})" then have "interior(affine hull S) = {}" using assms by (metis DIM_positive One_nat_def Suc_mono card.remove card.infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) then show "interior (convex hull S) = {}" using affine_hull_nonempty_interior by fastforce qed next show "interior (convex hull S) = {} \ affine_dependent S" by (metis affine_hull_convex_hull affine_hull_empty affine_independent_span_eq assms convex_convex_hull empty_not_UNIV rel_interior_eq_empty rel_interior_interior) qed subsection \Coplanarity, and collinearity in terms of affine hull\ definition\<^marker>\tag important\ coplanar where "coplanar S \ \u v w. S \ affine hull {u,v,w}" lemma collinear_affine_hull: "collinear S \ (\u v. S \ affine hull {u,v})" proof (cases "S={}") case True then show ?thesis by simp next case False then obtain x where x: "x \ S" by auto { fix u assume *: "\x y. \x\S; y\S\ \ \c. x - y = c *\<^sub>R u" have "\y c. x - y = c *\<^sub>R u \ \a b. y = a *\<^sub>R x + b *\<^sub>R (x + u) \ a + b = 1" by (rule_tac x="1+c" in exI, rule_tac x="-c" in exI, simp add: algebra_simps) then have "\u v. S \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" using * [OF x] by (rule_tac x=x in exI, rule_tac x="x+u" in exI, force) } moreover { fix u v x y assume *: "S \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" have "\c. x - y = c *\<^sub>R (v-u)" if "x\S" "y\S" proof - obtain a r where "a + r = 1" "x = a *\<^sub>R u + r *\<^sub>R v" using "*" \x \ S\ by blast moreover obtain b s where "b + s = 1" "y = b *\<^sub>R u + s *\<^sub>R v" using "*" \y \ S\ by blast ultimately have "x - y = (r-s) *\<^sub>R (v-u)" by (simp add: algebra_simps) (metis scaleR_left.add) then show ?thesis by blast qed } ultimately show ?thesis unfolding collinear_def affine_hull_2 by blast qed lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) lemma collinear_open_segment [simp]: "collinear (open_segment a b)" unfolding open_segment_def by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) lemma collinear_between_cases: fixes c :: "'a::euclidean_space" shows "collinear {a,b,c} \ between (b,c) a \ between (c,a) b \ between (a,b) c" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ {a, b, c} \ \c. x = u + c *\<^sub>R v" by (auto simp: collinear_alt) show ?rhs using uv [of a] uv [of b] uv [of c] by (auto simp: between_1) next assume ?rhs then show ?lhs unfolding between_mem_convex_hull by (metis (no_types, hide_lams) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull) qed lemma subset_continuous_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (closed_segment a b) f" shows "closed_segment (f a) (f b) \ image f (closed_segment a b)" by (metis connected_segment convex_contains_segment ends_in_segment imageI is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) lemma continuous_injective_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof show "closed_segment (f a) (f b) \ f ` closed_segment a b" by (metis subset_continuous_image_segment_1 contf) show "f ` closed_segment a b \ closed_segment (f a) (f b)" proof (cases "a = b") case True then show ?thesis by auto next case False then have fnot: "f a \ f b" using inj_onD injf by fastforce moreover have "f a \ open_segment (f c) (f b)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def) assume fa: "f a \ closed_segment (f c) (f b)" moreover have "closed_segment (f c) (f b) \ f ` closed_segment c b" by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) ultimately have "f a \ f ` closed_segment c b" by blast then have a: "a \ closed_segment c b" by (meson ends_in_segment inj_on_image_mem_iff injf subset_closed_segment that) have cb: "closed_segment c b \ closed_segment a b" by (simp add: closed_segment_subset that) show "f a = f c" proof (rule between_antisym) show "between (f c, f b) (f a)" by (simp add: between_mem_segment fa) show "between (f a, f b) (f c)" by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) qed qed moreover have "f b \ open_segment (f a) (f c)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def fnot eq_commute) assume fb: "f b \ closed_segment (f a) (f c)" moreover have "closed_segment (f a) (f c) \ f ` closed_segment a c" by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) ultimately have "f b \ f ` closed_segment a c" by blast then have b: "b \ closed_segment a c" by (meson ends_in_segment inj_on_image_mem_iff injf subset_closed_segment that) have ca: "closed_segment a c \ closed_segment a b" by (simp add: closed_segment_subset that) show "f b = f c" proof (rule between_antisym) show "between (f c, f a) (f b)" by (simp add: between_commute between_mem_segment fb) show "between (f b, f a) (f c)" by (metis b between_antisym between_commute between_mem_segment between_triv2 that) qed qed ultimately show ?thesis by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) qed qed lemma continuous_injective_image_open_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (open_segment a b) = open_segment (f a) (f b)" proof - have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}" by (metis (no_types, hide_lams) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) also have "... = open_segment (f a) (f b)" using continuous_injective_image_segment_1 [OF assms] by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) finally show ?thesis . qed lemma collinear_imp_coplanar: "collinear s ==> coplanar s" by (metis collinear_affine_hull coplanar_def insert_absorb2) lemma collinear_small: assumes "finite s" "card s \ 2" shows "collinear s" proof - have "card s = 0 \ card s = 1 \ card s = 2" using assms by linarith then show ?thesis using assms using card_eq_SucD numeral_2_eq_2 by (force simp: card_1_singleton_iff) qed lemma coplanar_small: assumes "finite s" "card s \ 3" shows "coplanar s" proof - consider "card s \ 2" | "card s = Suc (Suc (Suc 0))" using assms by linarith then show ?thesis proof cases case 1 then show ?thesis by (simp add: \finite s\ collinear_imp_coplanar collinear_small) next case 2 then show ?thesis using hull_subset [of "{_,_,_}"] by (fastforce simp: coplanar_def dest!: card_eq_SucD) qed qed lemma coplanar_empty: "coplanar {}" by (simp add: coplanar_small) lemma coplanar_sing: "coplanar {a}" by (simp add: coplanar_small) lemma coplanar_2: "coplanar {a,b}" by (auto simp: card_insert_if coplanar_small) lemma coplanar_3: "coplanar {a,b,c}" by (auto simp: card_insert_if coplanar_small) lemma collinear_affine_hull_collinear: "collinear(affine hull s) \ collinear s" unfolding collinear_affine_hull by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \ coplanar s" unfolding coplanar_def by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "coplanar S" "linear f" shows "coplanar(f ` S)" proof - { fix u v w assume "S \ affine hull {u, v, w}" then have "f ` S \ f ` (affine hull {u, v, w})" by (simp add: image_mono) then have "f ` S \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed lemma coplanar_translation_imp: assumes "coplanar S" shows "coplanar ((\x. a + x) ` S)" proof - obtain u v w where "S \ affine hull {u,v,w}" by (meson assms coplanar_def) then have "(+) a ` S \ affine hull {u + a, v + a, w + a}" using affine_hull_translation [of a "{u,v,w}" for u v w] by (force simp: add.commute) then show ?thesis unfolding coplanar_def by blast qed lemma coplanar_translation_eq: "coplanar((\x. a + x) ` S) \ coplanar S" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "coplanar(f ` S) = coplanar S" proof assume "coplanar S" then show "coplanar (f ` S)" using assms(1) coplanar_linear_image by blast next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast assume "coplanar (f ` S)" then show "coplanar S" by (metis coplanar_linear_image g(1) g(2) id_apply image_comp image_id) qed lemma coplanar_subset: "\coplanar t; S \ t\ \ coplanar S" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) lemma collinear_3_imp_in_affine_hull: assumes "collinear {a,b,c}" "a \ b" shows "c \ affine hull {a,b}" proof - obtain u x y where "b - a = y *\<^sub>R u" "c - a = x *\<^sub>R u" using assms unfolding collinear_def by auto with \a \ b\ have "\v. c = (1 - x / y) *\<^sub>R a + v *\<^sub>R b \ 1 - x / y + v = 1" by (simp add: algebra_simps) then show ?thesis by (simp add: hull_inc mem_affine) qed lemma collinear_3_affine_hull: assumes "a \ b" shows "collinear {a,b,c} \ c \ affine hull {a,b}" using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast lemma collinear_3_eq_affine_dependent: "collinear{a,b,c} \ a = b \ a = c \ b = c \ affine_dependent {a,b,c}" proof (cases "a = b \ a = c \ b = c") case True then show ?thesis by (auto simp: insert_commute) next case False then have "collinear{a,b,c}" if "affine_dependent {a,b,c}" using that unfolding affine_dependent_def by (auto simp: insert_Diff_if; metis affine_hull_3_imp_collinear insert_commute) moreover have "affine_dependent {a,b,c}" if "collinear{a,b,c}" using False that by (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) ultimately show ?thesis using False by blast qed lemma affine_dependent_imp_collinear_3: "affine_dependent {a,b,c} \ collinear{a,b,c}" by (simp add: collinear_3_eq_affine_dependent) lemma collinear_3: "NO_MATCH 0 x \ collinear {x,y,z} \ collinear {0, x-y, z-y}" by (auto simp add: collinear_def) lemma collinear_3_expand: "collinear{a,b,c} \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" proof - have "collinear{a,b,c} = collinear{a,c,b}" by (simp add: insert_commute) also have "... = collinear {0, a - c, b - c}" by (simp add: collinear_3) also have "... \ (a = c \ b = c \ (\ca. b - c = ca *\<^sub>R (a - c)))" by (simp add: collinear_lemma) also have "... \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" by (cases "a = c \ b = c") (auto simp: algebra_simps) finally show ?thesis . qed lemma collinear_aff_dim: "collinear S \ aff_dim S \ 1" proof assume "collinear S" then obtain u and v :: "'a" where "aff_dim S \ aff_dim {u,v}" by (metis \collinear S\ aff_dim_affine_hull aff_dim_subset collinear_affine_hull) then show "aff_dim S \ 1" using order_trans by fastforce next assume "aff_dim S \ 1" then have le1: "aff_dim (affine hull S) \ 1" by simp obtain B where "B \ S" and B: "\ affine_dependent B" "affine hull S = affine hull B" using affine_basis_exists [of S] by auto then have "finite B" "card B \ 2" using B le1 by (auto simp: affine_independent_iff_card) then have "collinear B" by (rule collinear_small) then show "collinear S" by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed lemma collinear_midpoint: "collinear{a, midpoint a b, b}" proof - have \
: "\a \ midpoint a b; b - midpoint a b \ - 1 *\<^sub>R (a - midpoint a b)\ \ b = midpoint a b" by (simp add: algebra_simps) show ?thesis by (auto simp: collinear_3 collinear_lemma intro: \
) qed lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" assumes "a \ c" shows "b = midpoint a c \ collinear{a,b,c} \ dist a b = dist b c" proof - have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) have "b = midpoint a c \ collinear{a,b,c}" using collinear_midpoint by blast moreover have "b = midpoint a c \ dist a b = dist b c" if "collinear{a,b,c}" proof - consider "a = c" | u where "b = u *\<^sub>R a + (1 - u) *\<^sub>R c" using \collinear {a,b,c}\ unfolding collinear_3_expand by blast then show ?thesis proof cases case 2 with assms have "dist a b = dist b c \ b = midpoint a c" by (simp add: dist_norm * midpoint_def scaleR_add_right del: divide_const_simps) then show ?thesis by (auto simp: dist_midpoint) qed (use assms in auto) qed ultimately show ?thesis by blast qed lemma between_imp_collinear: fixes x :: "'a :: euclidean_space" assumes "between (a,b) x" shows "collinear {a,x,b}" proof (cases "x = a \ x = b \ a = b") case True with assms show ?thesis by (auto simp: dist_commute) next case False then have False if "\c. b - x \ c *\<^sub>R (a - x)" using that [of "-(norm(b - x) / norm(x - a))"] assms by (simp add: between_norm vector_add_divide_simps flip: real_vector.scale_minus_right) then show ?thesis by (auto simp: collinear_3 collinear_lemma) qed lemma midpoint_between: fixes a b :: "'a::euclidean_space" shows "b = midpoint a c \ between (a,c) b \ dist a b = dist b c" proof (cases "a = c") case False show ?thesis using False between_imp_collinear between_midpoint(1) midpoint_collinear by blast qed (auto simp: dist_commute) lemma collinear_triples: assumes "a \ b" shows "collinear(insert a (insert b S)) \ (\x \ S. collinear{a,b,x})" (is "?lhs = ?rhs") proof safe fix x assume ?lhs and "x \ S" then show "collinear {a, b, x}" using collinear_subset by force next assume ?rhs then have "\x \ S. collinear{a,x,b}" by (simp add: insert_commute) then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ insert a (insert b S)" for x using that assms collinear_3_expand by fastforce+ have "\c. x - y = c *\<^sub>R (b - a)" if x: "x \ insert a (insert b S)" and y: "y \ insert a (insert b S)" for x y proof - obtain u v where "x = u *\<^sub>R a + (1 - u) *\<^sub>R b" "y = v *\<^sub>R a + (1 - v) *\<^sub>R b" using "*" x y by presburger then have "x - y = (v - u) *\<^sub>R (b - a)" by (simp add: scale_left_diff_distrib scale_right_diff_distrib) then show ?thesis .. qed then show ?lhs unfolding collinear_def by metis qed lemma collinear_4_3: assumes "a \ b" shows "collinear {a,b,c,d} \ collinear{a,b,c} \ collinear{a,b,d}" using collinear_triples [OF assms, of "{c,d}"] by (force simp:) lemma collinear_3_trans: assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \ c" shows "collinear{a,b,d}" proof - have "collinear{b,c,a,d}" by (metis (full_types) assms collinear_4_3 insert_commute) then show ?thesis by (simp add: collinear_subset) qed lemma affine_hull_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" proof - have 1: "u *\<^sub>R a + v *\<^sub>R b = a + v *\<^sub>R (b - a)" if "u + v = 1" for u v using that by (simp add: algebra_simps flip: scaleR_add_left) have 2: "a + u *\<^sub>R (b - a) = (1 - u) *\<^sub>R a + u *\<^sub>R b" for u by (auto simp: algebra_simps) show ?thesis by (force simp add: affine_hull_2 dest: 1 intro!: 2) qed lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" assumes "\ collinear{a,b,c}" and 2: "DIM('a) = 2" shows "interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" (is "?lhs = ?rhs") proof have abc: "a \ b" "a \ c" "b \ c" "\ affine_dependent {a, b, c}" using assms by (auto simp: collinear_3_eq_affine_dependent) with 2 show "?lhs \ ?rhs" by (fastforce simp add: interior_convex_hull_explicit_minimal) show "?rhs \ ?lhs" using abc 2 apply (clarsimp simp add: interior_convex_hull_explicit_minimal) subgoal for x y z by (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) auto done qed subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" by auto lemma hyperplane_eq_Ex: assumes "a \ 0" obtains x where "a \ x = b" by (rule_tac x = "(b / (a \ a)) *\<^sub>R a" in that) (simp add: assms) lemma hyperplane_eq_empty: "{x. a \ x = b} = {} \ a = 0 \ b \ 0" using hyperplane_eq_Ex by (metis (mono_tags, lifting) empty_Collect_eq inner_zero_left) lemma hyperplane_eq_UNIV: "{x. a \ x = b} = UNIV \ a = 0 \ b = 0" proof - have "a = 0 \ b = 0" if "UNIV \ {x. a \ x = b}" using subsetD [OF that, where c = "((b+1) / (a \ a)) *\<^sub>R a"] by (simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_lt: "{x. a \ x < b} = {} \ a = 0 \ b \ 0" proof - have "a = 0 \ b \ 0" if "{x. a \ x < b} \ {}" using subsetD [OF that, where c = "((b-1) / (a \ a)) *\<^sub>R a"] by (force simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_gt: "{x. a \ x > b} = {} \ a = 0 \ b \ 0" using halfspace_eq_empty_lt [of "-a" "-b"] by simp lemma halfspace_eq_empty_le: "{x. a \ x \ b} = {} \ a = 0 \ b < 0" proof - have "a = 0 \ b < 0" if "{x. a \ x \ b} \ {}" using subsetD [OF that, where c = "((b-1) / (a \ a)) *\<^sub>R a"] by (force simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_ge: "{x. a \ x \ b} = {} \ a = 0 \ b > 0" using halfspace_eq_empty_le [of "-a" "-b"] by simp subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" proof (cases "S = {} \ T = {}") case True with that show ?thesis by auto next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on UNIV f" unfolding f_def by (intro continuous_intros continuous_on_setdist) show ?thesis proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that) show "{x. 0 < f x} \ {x. f x < 0} = {}" by auto show "open {x. 0 < f x}" by (simp add: open_Collect_less contf) show "open {x. f x < 0}" by (simp add: open_Collect_less contf) have "\x. x \ S \ setdist {x} T \ 0" "\x. x \ T \ setdist {x} S \ 0" by (meson False assms disjoint_iff setdist_eq_0_sing_1)+ then show "S \ {x. 0 < f x}" "T \ {x. f x < 0}" using less_eq_real_def by (fastforce simp add: f_def setdist_sing_in_set)+ qed qed lemma separation_normal: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "S \ T = {}" obtains U V where "open U" "open V" "S \ U" "T \ V" "U \ V = {}" using separation_closures [of S T] by (metis assms closure_closed disjnt_def inf_commute) lemma separation_normal_local: fixes S :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains S' T' where "openin (top_of_set U) S'" "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis using UT US by (blast dest: closedin_subset) next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on U f" unfolding f_def by (intro continuous_intros) show ?thesis proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" using closedin_imp_subset assms by blast+ then show "S \ U \ f -` {0<..}" "T \ U \ f -` {..<0}" using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+ qed qed lemma separation_normal_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" "closed T" "S \ T = {}" obtains U V where "open U" "compact(closure U)" "open V" "S \ U" "T \ V" "U \ V = {}" proof - have "closed S" "bounded S" using assms by (auto simp: compact_eq_bounded_closed) then obtain r where "r>0" and r: "S \ ball 0 r" by (auto dest!: bounded_subset_ballD) have **: "closed (T \ - ball 0 r)" "S \ (T \ - ball 0 r) = {}" using assms r by blast+ then obtain U V where UV: "open U" "open V" "S \ U" "T \ - ball 0 r \ V" "U \ V = {}" by (meson \closed S\ separation_normal) then have "compact(closure U)" by (meson bounded_ball bounded_subset compact_closure compl_le_swap2 disjoint_eq_subset_Compl le_sup_iff) with UV show thesis using that by auto qed subsection\Connectedness of the intersection of a chain\ proposition connected_chain: fixes \ :: "'a :: euclidean_space set set" assumes cc: "\S. S \ \ \ compact S \ connected S" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof (cases "\ = {}") case True then show ?thesis by auto next case False then have cf: "compact(\\)" by (simp add: cc compact_Inter) have False if AB: "closed A" "closed B" "A \ B = {}" and ABeq: "A \ B = \\" and "A \ {}" "B \ {}" for A B proof - obtain U V where "open U" "open V" "A \ U" "B \ V" "U \ V = {}" using separation_normal [OF AB] by metis obtain K where "K \ \" "compact K" using cc False by blast then obtain N where "open N" and "K \ N" by blast let ?\ = "insert (U \ V) ((\S. N - S) ` \)" obtain \ where "\ \ ?\" "finite \" "K \ \\" proof (rule compactE [OF \compact K\]) show "K \ \(insert (U \ V) ((-) N ` \))" using \K \ N\ ABeq \A \ U\ \B \ V\ by auto show "\B. B \ insert (U \ V) ((-) N ` \) \ open B" by (auto simp: \open U\ \open V\ open_Un \open N\ cc compact_imp_closed open_Diff) qed then have "finite(\ - {U \ V})" by blast moreover have "\ - {U \ V} \ (\S. N - S) ` \" using \\ \ ?\\ by blast ultimately obtain \ where "\ \ \" "finite \" and Deq: "\ - {U \ V} = (\S. N-S) ` \" using finite_subset_image by metis obtain J where "J \ \" and J: "(\S\\. N - S) \ N - J" proof (cases "\ = {}") case True with \\ \ {}\ that show ?thesis by auto next case False have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (meson \\ \ \\ in_mono local.linear) with \finite \\ \\ \ {}\ have "\J \ \. (\S\\. N - S) \ N - J" proof induction case (insert X \) show ?case proof (cases "\ = {}") case True then show ?thesis by auto next case False then have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (simp add: insert.prems) with insert.IH False obtain J where "J \ \" and J: "(\Y\\. N - Y) \ N - J" by metis have "N - J \ N - X \ N - X \ N - J" by (meson Diff_mono \J \ \\ insert.prems(2) insert_iff order_refl) then show ?thesis proof assume "N - J \ N - X" with J show ?thesis by auto next assume "N - X \ N - J" with J have "N - X \ \ ((-) N ` \) \ N - J" by auto with \J \ \\ show ?thesis by blast qed qed qed simp with \\ \ \\ show ?thesis by (blast intro: that) qed have "K \ \(insert (U \ V) (\ - {U \ V}))" using \K \ \\\ by auto also have "... \ (U \ V) \ (N - J)" by (metis (no_types, hide_lams) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1) finally have "J \ K \ U \ V" by blast moreover have "connected(J \ K)" by (metis Int_absorb1 \J \ \\ \K \ \\ cc inf.orderE local.linear) moreover have "U \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \A \ {}\ \A \ U\ by blast moreover have "V \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \B \ {}\ \B \ V\ by blast ultimately show False using connectedD [of "J \ K" U V] \open U\ \open V\ \U \ V = {}\ by auto qed with cf show ?thesis by (auto simp: connected_closed_set compact_imp_closed) qed lemma connected_chain_gen: fixes \ :: "'a :: euclidean_space set set" assumes X: "X \ \" "compact X" and cc: "\T. T \ \ \ closed T \ connected T" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof - have "\\ = (\T\\. X \ T)" using X by blast moreover have "connected (\T\\. X \ T)" proof (rule connected_chain) show "\T. T \ (\) X ` \ \ compact T \ connected T" using cc X by auto (metis inf.absorb2 inf.orderE local.linear) show "\S T. S \ (\) X ` \ \ T \ (\) X ` \ \ S \ T \ T \ S" using local.linear by blast qed ultimately show ?thesis by metis qed lemma connected_nest: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" proof (rule connected_chain) show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" by (metis image_iff le_cases nest) qed (use S in blast) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" proof (rule connected_chain_gen [of "S k"]) show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" by (metis imageE le_cases nest) qed (use S in auto) subsection\Proper maps, including projections out of compact sets\ lemma finite_indexed_bound: assumes A: "finite A" "\x. x \ A \ \n::'a::linorder. P x n" shows "\m. \x \ A. \k\m. P x k" using A proof (induction A) case empty then show ?case by force next case (insert a A) then obtain m n where "\x \ A. \k\m. P x k" "P a n" by force then show ?case by (metis dual_order.trans insert_iff le_cases) qed proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis obtain C where "closed C" and Keq: "K = S \ C" using assms by (auto simp: closedin_closed) have *: "y \ f ` K" if "y \ T" and y: "y islimpt f ` K" for y proof - obtain h where "\n. (\x\K. h n = f x) \ h n \ y" "inj h" and hlim: "(h \ y) sequentially" using \y \ T\ y by (force simp: limpt_sequential_inj) then obtain X where X: "\n. X n \ K \ h n = f (X n) \ h n \ y" by metis then have fX: "\n. f (X n) = h n" by metis define \ where "\ \ \n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))}" have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n proof (intro closed_Int_compact [OF \closed C\ com] compact_sequence_with_limit) show "insert y (range (\i. f (X (n + i)))) \ T" using X \K \ S\ \f ` S \ T\ \y \ T\ by blast show "(\i. f (X (n + i))) \ y" by (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) qed then have comf: "compact (\ n)" for n by (simp add: Keq Int_def \_def conj_commute) have ne: "\\ \ {}" if "finite \" and \: "\t. t \ \ \ (\n. t = \ n)" for \ proof - obtain m where m: "\t. t \ \ \ \k\m. t = \ k" by (rule exE [OF finite_indexed_bound [OF \finite \\ \]], force+) have "X m \ \\" using X le_Suc_ex by (fastforce simp: \_def dest: m) then show ?thesis by blast qed have "(\n. \ n) \ {}" proof (rule compact_fip_Heine_Borel) show "\\'. \finite \'; \' \ range \\ \ \ \' \ {}" by (meson ne rangeE subset_eq) qed (use comf in blast) then obtain x where "x \ K" "\n. (f x = y \ (\u. f x = h (n + u)))" by (force simp add: \_def fX) then show ?thesis unfolding image_iff by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) qed lemma compact_continuous_image_eq: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes f: "inj_on f S" shows "continuous_on S f \ (\T. compact T \ T \ S \ compact(f ` T))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_on_subset compact_continuous_image) next assume RHS: ?rhs obtain g where gf: "\x. x \ S \ g (f x) = x" by (metis inv_into_f_f f) then have *: "(S \ f -` U) = g ` U" if "U \ f ` S" for U using that by fastforce have gfim: "g ` f ` S \ S" using gf by auto have **: "compact (f ` S \ g -` C)" if C: "C \ S" "compact C" for C proof - obtain h where "h C \ C \ h C \ S \ compact (f ` C)" by (force simp: C RHS) moreover have "f ` C = (f ` S \ g -` C)" using C gf by auto ultimately show ?thesis using C by auto qed show ?lhs using proper_map [OF _ _ gfim] ** by (simp add: continuous_on_closed * closedin_imp_subset) qed subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" assumes "collinear S" shows "convex S \ connected S" proof assume "convex S" then show "connected S" using convex_connected by blast next assume S: "connected S" show "convex S" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - a = c *\<^sub>R z" by (meson \a \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - a = f x *\<^sub>R z" by metis then have inj_f: "inj_on f (affine hull S)" by (metis diff_add_cancel inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - a" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - a" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" proof (clarsimp simp: dist_norm continuous_on_iff diff) show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" - by (metis \z \ 0\ mult_pos_pos real_mult_less_iff1 zero_less_norm_iff) + by (metis \z \ 0\ mult_pos_pos mult_less_iff1 zero_less_norm_iff) qed then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis proof (clarsimp simp: convex_contains_segment) fix x y z assume "x \ S" "y \ S" "z \ closed_segment x y" have False if "z \ S" proof - have "f ` (closed_segment x y) = closed_segment (f x) (f y)" proof (rule continuous_injective_image_segment_1) show "continuous_on (closed_segment x y) f" by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) show "inj_on f (closed_segment x y)" by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) qed then have fz: "f z \ closed_segment (f x) (f y)" using \z \ closed_segment x y\ by blast have "z \ affine hull S" by (meson \x \ S\ \y \ S\ \z \ closed_segment x y\ convex_affine_hull convex_contains_segment hull_inc subset_eq) then have fz_notin: "f z \ f ` S" using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - consider "f x \ f z \ f z \ f y" | "f y \ f z \ f z \ f x" using fz by (auto simp add: closed_segment_eq_real_ivl split: if_split_asm) then have "{.. f ` {x,y} \ {} \ {f z<..} \ f ` {x,y} \ {}" by cases (use fz_notin \x \ S\ \y \ S\ in \auto simp: image_iff\) then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed ultimately show False using connectedD [OF conn_fS, of "{.. S" by meson qed qed qed lemma compact_convex_collinear_segment_alt: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "connected S" "collinear S" obtains a b where "S = closed_segment a b" proof - obtain \ where "\ \ S" using \S \ {}\ by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - \ = c *\<^sub>R z" by (meson \\ \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - \ = f x *\<^sub>R z" by metis let ?g = "\r. r *\<^sub>R z + \" have gf: "?g (f x) = x" if "x \ affine hull S" for x by (metis diff_add_cancel f that) then have inj_f: "inj_on f (affine hull S)" by (metis inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - \" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - \" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" proof (clarsimp simp: dist_norm continuous_on_iff diff) show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" - by (metis \z \ 0\ mult_pos_pos real_mult_less_iff1 zero_less_norm_iff) + by (metis \z \ 0\ mult_pos_pos mult_less_iff1 zero_less_norm_iff) qed then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" by (meson \compact S\ compact_continuous_image_eq cont_f hull_subset inj_f) ultimately obtain x y where "f ` S = {x..y}" by (meson connected_compact_interval_1) then have fS_eq: "f ` S = closed_segment x y" using \S \ {}\ closed_segment_eq_real_ivl by auto obtain a b where "a \ S" "f a = x" "b \ S" "f b = y" by (metis (full_types) ends_in_segment fS_eq imageE) have "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof (rule continuous_injective_image_segment_1) show "continuous_on (closed_segment a b) f" by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) show "inj_on f (closed_segment a b)" by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) qed then have "f ` (closed_segment a b) = f ` S" by (simp add: \f a = x\ \f b = y\ fS_eq) then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" unfolding image_def using \a \ S\ \b \ S\ by (safe; metis (mono_tags, lifting) convex_affine_hull convex_contains_segment gf hull_subset subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis using that by blast qed lemma compact_convex_collinear_segment: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "convex S" "collinear S" obtains a b where "S = closed_segment a b" using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ lemma proper_map_fst: assumes "compact T" "K \ S" "compact K" shows "compact (S \ T \ fst -` K)" proof - have "(S \ T \ fst -` K) = K \ T" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact T" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) qed lemma proper_map_snd: assumes "compact S" "K \ T" "compact K" shows "compact (S \ T \ snd -` K)" proof - have "(S \ T \ snd -` K) = S \ K" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) qed lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp qed lemma closed_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "('a * 'b::euclidean_space) set" assumes "compact S" and clo: "closed T" shows "closed {y. \x. x \ S \ (x, y) \ T}" proof - have *: "{y. \x. x \ S \ Pair x y \ T} = {y. \x. x \ S \ Pair x y \ ((S \ UNIV) \ T)}" by auto show ?thesis unfolding * by (intro clo closedin_closed_Int closedin_closed_trans [OF _ closed_UNIV] closedin_compact_projection [OF \compact S\]) qed subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" proof show "affine hull (S \ T) \ affine hull S" by (simp add: hull_mono) next obtain a where "a \ S" "a \ T" and at: "a \ interior T" using assms interior_subset by blast then obtain e where "e > 0" and e: "cball a e \ T" using mem_interior_cball by blast have *: "x \ (+) a ` span ((\x. x - a) ` (S \ T))" if "x \ S" for x proof (cases "x = a") case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis by blast next case False define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" by simp have "e / norm (x - a) \ k" using k_def by linarith then have "a + k *\<^sub>R (x - a) \ cball a e" using \0 < k\ False by (simp add: dist_norm) (simp add: field_simps) then have T: "a + k *\<^sub>R (x - a) \ T" using e by blast have S: "a + k *\<^sub>R (x - a) \ S" using k \a \ S\ convexD [OF \convex S\ \a \ S\ \x \ S\, of "1-k" k] by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" by (intro span_mul [OF span_base] image_eqI [where x = "a + k *\<^sub>R (x - a)"]) (auto simp: S T) with xa image_iff show ?thesis by fastforce qed have "S \ affine hull (S \ T)" by (force simp: * \a \ S\ \a \ T\ hull_inc affine_hull_span_gen [of a]) then show "affine hull S \ affine hull (S \ T)" by (simp add: subset_hull) qed corollary affine_hull_convex_Int_open: fixes S :: "'a::real_normed_vector set" assumes "convex S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast corollary affine_hull_affine_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "affine S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) corollary affine_hull_affine_Int_open: fixes S :: "'a::real_normed_vector set" assumes "affine S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) corollary affine_hull_open: fixes S :: "'a::real_normed_vector set" assumes "open S" "S \ {}" shows "affine hull S = UNIV" by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) lemma aff_dim_convex_Int_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ interior T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast lemma aff_dim_convex_Int_open: fixes S :: "'a::euclidean_space set" shows "\convex S; open T; S \ T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_convex_Int_nonempty_interior interior_eq by blast lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto ultimately show ?thesis by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) qed lemma affine_hull_halfspace_lt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x < r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_lt [of a r] by (simp add: open_halfspace_lt affine_hull_open) lemma affine_hull_halfspace_le: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r < 0 then {} else UNIV)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "affine hull closure {x. a \ x < r} = UNIV" using affine_hull_halfspace_lt closure_same_affine_hull by fastforce moreover have "{x. a \ x < r} \ {x. a \ x \ r}" by (simp add: Collect_mono) ultimately show ?thesis using False antisym_conv hull_mono top_greatest by (metis affine_hull_halfspace_lt) qed lemma affine_hull_halfspace_gt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x > r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_gt [of r a] by (simp add: open_halfspace_gt affine_hull_open) lemma affine_hull_halfspace_ge: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r > 0 then {} else UNIV)" using affine_hull_halfspace_le [of "-a" "-r"] by simp lemma aff_dim_halfspace_lt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x < r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) lemma aff_dim_halfspace_le: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r < 0 then -1 else DIM('a))" proof - have "int (DIM('a)) = aff_dim (UNIV::'a set)" by (simp) then have "aff_dim (affine hull {x. a \ x \ r}) = DIM('a)" if "(a = 0 \ r \ 0)" using that by (simp add: affine_hull_halfspace_le not_less) then show ?thesis by (force) qed lemma aff_dim_halfspace_gt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x > r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) lemma aff_dim_halfspace_ge: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r > 0 then -1 else DIM('a))" using aff_dim_halfspace_le [of "-a" "-r"] by simp proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) next case False then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") case True have "?lhs \ (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) also have "... \ ?rhs" using span_zero [of S] True \c \ S\ affine_hull_span_0 hull_inc by (fastforce simp add: affine_hull_span_gen [of c] \c = 0\) finally show ?thesis . next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x proof - have "\y. a \ y = 0 \ c + y = x" by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) then show "x \ (+) c ` {y. a \ y = 0}" by blast qed have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" if "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = b}" for a b proof - have "b = a \ c" using span_0 that by fastforce with that have "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = a \ c}" by simp then have "span ((\x. x - c) ` S) = (\x. x - c) ` {x. a \ x = a \ c}" by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) also have "... = {x. a \ x = 0}" by (force simp: inner_distrib inner_diff_right intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed have "?lhs = (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) also have "... = ?rhs" by (fastforce simp add: affine_hull_span_gen [of c] \c \ S\ hull_inc inner_distrib intro: xc_im intro!: 2) finally show ?thesis . qed qed corollary aff_dim_hyperplane [simp]: fixes a :: "'a::euclidean_space" shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain x s' where S: "S = insert x s'" "x \ s'" by (meson Set.set_insert all_not_in_conv) show ?thesis using S by (force simp add: affine_hull_insert_span_gen span_zero insert_commute [of a] aff_dim_eq_dim [of x] dim_insert) qed lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe assume "affine_dependent (insert a S)" and "a \ S" then show "False" using \a \ S\ assms insert_absorb by fastforce next assume lhs: "affine_dependent (insert a S)" then have "a \ S" by (metis (no_types) assms insert_absorb) moreover have "finite S" using affine_independent_iff_card assms by blast moreover have "aff_dim (insert a S) \ int (card S)" using \finite S\ affine_independent_iff_card \a \ S\ lhs by fastforce ultimately show "a \ affine hull S" by (metis aff_dim_affine_independent aff_dim_insert assms) next assume "a \ S" and "a \ affine hull S" show "affine_dependent (insert a S)" by (simp add: \a \ affine hull S\ \a \ S\ affine_dependent_def) qed lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "subspace S" shows "bounded S \ S = {0}" proof - have "False" if "bounded S" "x \ S" "x \ 0" for x proof - obtain B where B: "\y. y \ S \ norm y < B" "B > 0" using \bounded S\ by (force simp: bounded_pos_less) have "(B / norm x) *\<^sub>R x \ S" using assms subspace_mul \x \ S\ by auto moreover have "norm ((B / norm x) *\<^sub>R x) = B" using that B by (simp add: algebra_simps) ultimately show False using B by force qed then have "bounded S \ S = {0}" using assms subspace_0 by fastforce then show ?thesis by blast qed lemma affine_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "affine S" shows "bounded S \ S = {} \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain b where "b \ S" by blast with False assms have "bounded S \ S = {b}" using affine_diffs_subspace [OF assms \b \ S\] by (metis (no_types, lifting) ab_group_add_class.ab_left_minus bounded_translation image_empty image_insert subspace_bounded_eq_trivial translation_invert) then show ?thesis by auto qed lemma affine_bounded_eq_lowdim: fixes S :: "'a::euclidean_space set" assumes "affine S" shows "bounded S \ aff_dim S \ 0" proof show "aff_dim S \ 0 \ bounded S" by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) qed (use affine_bounded_eq_trivial assms in fastforce) lemma bounded_hyperplane_eq_trivial_0: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "bounded {x. a \ x = 0} \ DIM('a) = 1" proof assume "bounded {x. a \ x = 0}" then have "aff_dim {x. a \ x = 0} \ 0" by (simp add: affine_bounded_eq_lowdim affine_hyperplane) with assms show "DIM('a) = 1" by (simp add: le_Suc_eq) next assume "DIM('a) = 1" then show "bounded {x. a \ x = 0}" by (simp add: affine_bounded_eq_lowdim affine_hyperplane assms) qed lemma bounded_hyperplane_eq_trivial: fixes a :: "'a::euclidean_space" shows "bounded {x. a \ x = r} \ (if a = 0 then r \ 0 else DIM('a) = 1)" proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) assume "r \ 0" "a \ 0" have "aff_dim {x. y \ x = 0} = aff_dim {x. a \ x = r}" if "y \ 0" for y::'a by (metis that \a \ 0\ aff_dim_hyperplane) then show "bounded {x. a \ x = r} = (DIM('a) = Suc 0)" by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" proof - obtain y where "y \ S" and y: "\u. u \ S \ dist z y \ dist z u" using distance_attains_inf [of S z] assms by auto then have *: "(y - z) \ z < (y - z) \ z + (norm (y - z))\<^sup>2 / 2" using \y \ S\ \z \ S\ by auto show ?thesis proof (rule that [OF \y \ S\ *]) fix x assume "x \ S" have yz: "0 < (y - z) \ (y - z)" using \y \ S\ \z \ S\ by auto { assume 0: "0 < ((z - y) \ (x - y))" with any_closest_point_dot [OF \convex S\ \closed S\] have False using y \x \ S\ \y \ S\ not_less by blast } then have "0 \ ((y - z) \ (x - y))" by (force simp: not_less inner_diff_left) with yz have "0 < 2 * ((y - z) \ (x - y)) + (y - z) \ (y - z)" by (simp add: algebra_simps) then show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) qed qed lemma separating_hyperplane_closed_0_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "0 \ S" obtains a b where "a \ S" "a \ 0" "0 < b" "\x. x \ S \ a \ x > b" using separating_hyperplane_closed_point_inset [OF assms] by simp (metis \0 \ S\) proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - obtain C where "C \ S" "finite C" and C: "f' = k ` C" using finite_subset_image [OF f'] by blast obtain a where "a \ S" "a \ 0" using \S \ {}\ \0 \ S\ ex_in_conv by blast then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis proof (cases "C = {}") case True with C ass show ?thesis by auto next case False have "closed (convex hull C)" using \finite C\ compact_eq_bounded_closed finite_imp_compact_convex_hull by auto moreover have "convex hull C \ {}" by (simp add: False) moreover have "0 \ convex hull C" by (metis \C \ S\ \convex S\ \0 \ S\ convex_hull_subset hull_same insert_absorb insert_subset) ultimately obtain a b where "a \ convex hull C" "a \ 0" "0 < b" and ab: "\x. x \ convex hull C \ a \ x > b" using separating_hyperplane_closed_0_inset by blast then have "a \ S" by (metis \C \ S\ assms(1) subsetCE subset_hull) moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have "\x. \a \ 0; x \ C\ \ 0 \ x \ a" using ab \0 < b\ by (metis hull_inc inner_commute less_eq_real_def less_trans) then have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" by (auto simp add: field_split_simps) show ?thesis unfolding C k_def using ass aa Int_iff empty_iff by force qed qed moreover have "\T. T \ k ` S \ closed T" using closed_halfspace_ge k_def by blast ultimately have "(span S \ frontier(cball 0 1)) \ (\ (k ` S)) \ {}" by (metis compact_imp_fip closed_Int_compact closed_span compact_cball compact_frontier) then show ?thesis unfolding set_eq_iff k_def by simp (metis inner_commute norm_eq_zero that zero_neq_one) qed lemma separating_hyperplane_set_point_inaff: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and zno: "z \ S" obtains a b where "(z + a) \ affine hull (insert z S)" and "a \ 0" and "a \ z \ b" and "\x. x \ S \ a \ x \ b" proof - from separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] have "convex ((+) (- z) ` S)" using \convex S\ by simp moreover have "(+) (- z) ` S \ {}" by (simp add: \S \ {}\) moreover have "0 \ (+) (- z) ` S" using zno by auto ultimately obtain a where "a \ span ((+) (- z) ` S)" "a \ 0" and a: "\x. x \ ((+) (- z) ` S) \ 0 \ a \ x" using separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] by blast then have szx: "\x. x \ S \ a \ z \ a \ x" by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) moreover have "z + a \ affine hull insert z S" using \a \ span ((+) (- z) ` S)\ affine_hull_insert_span_gen by blast ultimately show ?thesis using \a \ 0\ szx that by auto qed proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" proof - obtain a b where aff: "(x + a) \ affine hull (insert x (rel_interior S))" and "a \ 0" and "a \ x \ b" and ageb: "\u. u \ (rel_interior S) \ a \ u \ b" using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms by (auto simp: rel_interior_eq_empty convex_rel_interior) have le_ay: "a \ x \ a \ y" if "y \ S" for y proof - have con: "continuous_on (closure (rel_interior S)) ((\) a)" by (rule continuous_intros continuous_on_subset | blast)+ have y: "y \ closure (rel_interior S)" using \convex S\ closure_def convex_closure_rel_interior \y \ S\ by fastforce show ?thesis using continuous_ge_on_closure [OF con y] ageb \a \ x \ b\ by fastforce qed have 3: "a \ x < a \ y" if "y \ rel_interior S" for y proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" unfolding y'_def by (metis \x \ S\ \y \ S\ \convex S\ aff affine_affine_hull hull_redundant rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) ultimately have "y' \ S" using e by auto have "a \ x \ a \ y" using le_ay \a \ 0\ \y \ S\ by blast moreover have "a \ x \ a \ y" using le_ay [OF \y' \ S\] \a \ 0\ \0 < e\ not_le by (fastforce simp add: y'_def inner_diff dot_square_norm power2_eq_square) ultimately show ?thesis by force qed show ?thesis by (rule that [OF \a \ 0\ le_ay 3]) qed lemma supporting_hyperplane_relative_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ closure S" "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ closure S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" using supporting_hyperplane_rel_boundary [of "closure S" x] by (metis assms convex_closure convex_rel_interior_closure) subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows convex_hull_Int_subset: "convex hull S \ convex hull T \ convex hull (S \ T)" (is ?C) and affine_hull_Int_subset: "affine hull S \ affine hull T \ affine hull (S \ T)" (is ?A) proof - have [simp]: "finite S" "finite T" using aff_independent_finite assms by blast+ have "sum u (S \ T) = 1 \ (\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" if [simp]: "sum u S = 1" "sum v T = 1" and eq: "(\x\T. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" for u v proof - define f where "f x = (if x \ S then u x else 0) - (if x \ T then v x else 0)" for x have "sum f (S \ T) = 0" by (simp add: f_def sum_Un sum_subtractf flip: sum.inter_restrict) moreover have "(\x\(S \ T). f x *\<^sub>R x) = 0" by (simp add: eq f_def sum_Un scaleR_left_diff_distrib sum_subtractf if_smult flip: sum.inter_restrict cong: if_cong) ultimately have "\v. v \ S \ T \ f v = 0" using aff_independent_finite assms unfolding affine_dependent_explicit by blast then have u [simp]: "\x. x \ S \ u x = (if x \ T then v x else 0)" by (simp add: f_def) presburger have "sum u (S \ T) = sum u S" by (simp add: sum.inter_restrict) then have "sum u (S \ T) = 1" using that by linarith moreover have "(\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force qed then show ?A ?C by (auto simp: convex_hull_finite affine_hull_finite) qed proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows "affine hull (S \ T) = affine hull S \ affine hull T" by (simp add: affine_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ convex_hull_Int: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows "convex hull (S \ T) = convex hull S \ convex hull T" by (simp add: convex_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ fixes S :: "'a::euclidean_space set set" assumes "\ affine_dependent (\S)" shows affine_hull_Inter: "affine hull (\S) = (\T\S. affine hull T)" (is "?A") and convex_hull_Inter: "convex hull (\S) = (\T\S. convex hull T)" (is "?C") proof - have "finite S" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms proof (induction S rule: finite_induct) case empty then show ?case by auto next case (insert T F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False with "insert.prems" have [simp]: "\ affine_dependent (T \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce show ?thesis by (simp add: affine_hull_Int convex_hull_Int insert.IH) qed qed then show "?A" "?C" by auto qed proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" shows "x \ convex hull (insert a (T \ T'))" proof (cases "a \ S") case True then have "\ affine_dependent (insert a T \ insert a T')" using affine_dependent_subset assms by auto then have "x \ convex hull (insert a T \ insert a T')" by (metis IntI convex_hull_Int x x') then show ?thesis by simp next case False then have anot: "a \ T" "a \ T'" using assms by auto have [simp]: "finite S" by (simp add: aff_independent_finite assms) then obtain b where b0: "\s. s \ S \ 0 \ b s" and b1: "sum b S = 1" and aeq: "a = (\s\S. b s *\<^sub>R s)" using a by (auto simp: convex_hull_finite) have fin [simp]: "finite T" "finite T'" using assms infinite_super \finite S\ by blast+ then obtain c c' where c0: "\t. t \ insert a T \ 0 \ c t" and c1: "sum c (insert a T) = 1" and xeq: "x = (\t \ insert a T. c t *\<^sub>R t)" and c'0: "\t. t \ insert a T' \ 0 \ c' t" and c'1: "sum c' (insert a T') = 1" and x'eq: "x = (\t \ insert a T'. c' t *\<^sub>R t)" using x x' by (auto simp: convex_hull_finite) with fin anot have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" and wsumT: "(\t \ T. c t *\<^sub>R t) = x - c a *\<^sub>R a" by simp_all have wsumT': "(\t \ T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" using x'eq fin anot by simp define cc where "cc \ \x. if x \ T then c x else 0" define cc' where "cc' \ \x. if x \ T' then c' x else 0" define dd where "dd \ \x. cc x - cc' x + (c a - c' a) * b x" have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') have wsumSS: "(\t \ S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\t \ S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) have sum_dd0: "sum dd S = 0" unfolding dd_def using S by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf algebra_simps sum_distrib_right [symmetric] b1) have "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\v\S. b v *\<^sub>R v)" for x by (simp add: pth_5 real_vector.scale_sum_right mult.commute) then have *: "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x using aeq by blast have "(\v \ S. dd v *\<^sub>R v) = 0" unfolding dd_def using S by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) then have dd0: "dd v = 0" if "v \ S" for v using naff [unfolded affine_dependent_explicit not_ex, rule_format, of S dd] using that sum_dd0 by force consider "c' a \ c a" | "c a \ c' a" by linarith then show ?thesis proof cases case 1 then have "sum cc S \ sum cc' S" by (simp add: sumSS') then have le: "cc x \ cc' x" if "x \ S" for x using dd0 [OF that] 1 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc(a := c a)) x" by (simp add: c0 cc_def) show "0 \ (cc(a := c a)) a" by (simp add: c0) have "sum (cc(a := c a)) (insert a (T \ T')) = c a + sum (cc(a := c a)) (T \ T')" by (simp add: anot) also have "... = c a + sum (cc(a := c a)) S" using \T \ S\ False cc0 cc_def \a \ S\ by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a + (1 - c a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS'(1)) finally show "sum (cc(a := c a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\x \ T \ T'. (cc(a := c a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c a *\<^sub>R a + (\x \ S. (cc(a := c a)) x *\<^sub>R x)" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = x" by simp qed next case 2 then have "sum cc' S \ sum cc S" by (simp add: sumSS') then have le: "cc' x \ cc x" if "x \ S" for x using dd0 [OF that] 2 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc' x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c'0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc'(a := c' a)) x" by (simp add: c'0 cc'_def) show "0 \ (cc'(a := c' a)) a" by (simp add: c'0) have "sum (cc'(a := c' a)) (insert a (T \ T')) = c' a + sum (cc'(a := c' a)) (T \ T')" by (simp add: anot) also have "... = c' a + sum (cc'(a := c' a)) S" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c' a + (1 - c' a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc'(a := c' a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\x \ T \ T'. (cc'(a := c' a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c' a *\<^sub>R a + (\x \ S. (cc'(a := c' a)) x *\<^sub>R x)" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = x" by simp qed qed qed corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" (is "?lhs = ?rhs") proof (rule subset_antisym) show "?lhs \ ?rhs" using in_convex_hull_exchange_unique assms by blast show "?rhs \ ?lhs" by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) qed lemma Int_closed_segment: fixes b :: "'a::euclidean_space" assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True then show ?thesis using assms collinear_3_eq_affine_dependent by fastforce next case False from assms show ?thesis proof assume "b \ closed_segment a c" moreover have "\ affine_dependent {a, c}" by (simp) ultimately show ?thesis using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"] by (simp add: segment_convex_hull insert_commute) next assume ncoll: "\ collinear {a, b, c}" have False if "closed_segment a b \ closed_segment b c \ {b}" proof - have "b \ closed_segment a b" and "b \ closed_segment b c" by auto with that obtain d where "b \ d" "d \ closed_segment a b" "d \ closed_segment b c" by force then have d: "collinear {a, d, b}" "collinear {b, d, c}" by (auto simp: between_mem_segment between_imp_collinear) have "collinear {a, b, c}" by (metis (full_types) \b \ d\ collinear_3_trans d insert_commute) with ncoll show False .. qed then show ?thesis by blast qed qed lemma affine_hull_finite_intersection_hyperplanes: fixes S :: "'a::euclidean_space set" obtains \ where "finite \" "of_nat (card \) + aff_dim S = DIM('a)" "affine hull S = \\" "\h. h \ \ \ \a b. a \ 0 \ h = {x. a \ x = b}" proof - obtain b where "b \ S" and indb: "\ affine_dependent b" and eq: "affine hull S = affine hull b" using affine_basis_exists by blast obtain c where indc: "\ affine_dependent c" and "b \ c" and affc: "affine hull c = UNIV" by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV) then have "finite c" by (simp add: aff_independent_finite) then have fbc: "finite b" "card b \ card c" using \b \ c\ infinite_super by (auto simp: card_mono) have imeq: "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b)) = ((\a. affine hull (c - {a})) ` (c - b))" by blast have card_cb: "(card (c - b)) + aff_dim S = DIM('a)" proof - have aff: "aff_dim (UNIV::'a set) = aff_dim c" by (metis aff_dim_affine_hull affc) have "aff_dim b = aff_dim S" by (metis (no_types) aff_dim_affine_hull eq) then have "int (card b) = 1 + aff_dim S" by (simp add: aff_dim_affine_independent indb) then show ?thesis using fbc aff by (simp add: \\ affine_dependent c\ \b \ c\ aff_dim_affine_independent card_Diff_subset of_nat_diff) qed show ?thesis proof (cases "c = b") case True show ?thesis proof show "int (card {}) + aff_dim S = int DIM('a)" using True card_cb by auto show "affine hull S = \ {}" using True affc eq by blast qed auto next case False have ind: "\ affine_dependent (\a\c - b. c - {a})" by (rule affine_independent_subset [OF indc]) auto have *: "1 + aff_dim (c - {t}) = int (DIM('a))" if t: "t \ c" for t proof - have "insert t c = c" using t by blast then show ?thesis by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) qed let ?\ = "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b))" show ?thesis proof have "card ((\a. affine hull (c - {a})) ` (c - b)) = card (c - b)" proof (rule card_image) show "inj_on (\a. affine hull (c - {a})) (c - b)" unfolding inj_on_def by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) qed then show "int (card ?\) + aff_dim S = int DIM('a)" by (simp add: imeq card_cb) show "affine hull S = \ ?\" by (metis Diff_eq_empty_iff INT_simps(4) UN_singleton order_refl \b \ c\ False eq double_diff affine_hull_Inter [OF ind]) have "\a. \a \ c; a \ b\ \ aff_dim (c - {a}) = int (DIM('a) - Suc 0)" by (metis "*" DIM_ge_Suc0 One_nat_def add_diff_cancel_left' int_ops(2) of_nat_diff) then show "\h. h \ ?\ \ \a b. a \ 0 \ h = {x. a \ x = b}" by (auto simp only: One_nat_def aff_dim_eq_hyperplane [symmetric]) qed (use \finite c\ in auto) qed qed lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "0 \ S" and "w \ S" and "a \ w \ 0" shows "{x + y| x y. x \ S \ a \ y = 0} = UNIV" proof - have "subspace S" by (simp add: assms subspace_affine) have span1: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" using \0 \ S\ add.left_neutral by (intro span_mono) force have "w \ span {y. a \ y = 0}" using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast then have "DIM('a) - 1 = dim {y. a \ y = 0}" by (simp add: dim_hyperplane) also have "... < dim {x + y |x y. x \ S \ a \ y = 0}" using span1 span2 by (blast intro: dim_psubset) finally have "DIM('a) - 1 < dim {x + y |x y. x \ S \ a \ y = 0}" . then have DD: "dim {x + y |x y. x \ S \ a \ y = 0} = DIM('a)" using antisym dim_subset_UNIV lowdim_subset_hyperplane not_le by fastforce have subs: "subspace {x + y| x y. x \ S \ a \ y = 0}" using subspace_sums [OF \subspace S\ subspace_hyperplane] by simp moreover have "span {x + y| x y. x \ S \ a \ y = 0} = UNIV" using DD dim_eq_full by blast ultimately show ?thesis by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition\<^marker>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" and "S - {v. a \ v = b} \ {}" shows "{x + y| x y. x \ S \ a \ y = b} = UNIV" proof (cases "a = 0") case True with assms show ?thesis by (auto simp: if_splits) next case False obtain c where "c \ S" and c: "a \ c = b" using assms by force with affine_diffs_subspace [OF \affine S\] have "subspace ((+) (- c) ` S)" by blast then have aff: "affine ((+) (- c) ` S)" by (simp add: subspace_imp_affine) have 0: "0 \ (+) (- c) ` S" by (simp add: \c \ S\) obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ (+) (- c) ` S" using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) define U where "U \ {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" have "u + v \ (+) (c+c) ` U" if "u \ S" "b = a \ v" for u v proof show "u + v = c + c + (u + v - c - c)" by (simp add: algebra_simps) have "\x y. u + v - c - c = x + y \ (\xa\S. x = xa - c) \ a \ y = 0" proof (intro exI conjI) show "u + v - c - c = (u-c) + (v-c)" "a \ (v - c) = 0" by (simp_all add: algebra_simps c that) qed (use that in auto) then show "u + v - c - c \ U" by (auto simp: U_def image_def) qed moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) ultimately have "{x + y |x y. x \ S \ a \ y = b} = (+) (c+c) ` U" by (fastforce simp: algebra_simps U_def) also have "... = range ((+) (c + c))" by (simp only: U_def affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . qed lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" and "0 \ S" "0 \ T" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - have "0 \ {x + y |x y. x \ S \ y \ T}" using assms by force then have 0: "0 \ affine hull {x + y |x y. x \ S \ y \ T}" by (metis (lifting) hull_inc) have sub: "subspace S" "subspace T" using assms by (auto simp: subspace_affine) show ?thesis using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) qed proposition aff_dim_sums_Int: assumes "affine S" and "affine T" and "S \ T \ {}" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - obtain a where a: "a \ S" "a \ T" using assms by force have aff: "affine ((+) (-a) ` S)" "affine ((+) (-a) ` T)" using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp) have zero: "0 \ ((+) (-a) ` S)" "0 \ ((+) (-a) ` T)" using a assms by auto have "{x + y |x y. x \ (+) (- a) ` S \ y \ (+) (- a) ` T} = (+) (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" by (force simp: algebra_simps scaleR_2) moreover have "(+) (- a) ` S \ (+) (- a) ` T = (+) (- a) ` (S \ T)" by auto ultimately show ?thesis using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq by (metis (lifting)) qed lemma aff_dim_affine_Int_hyperplane: fixes a :: "'a::euclidean_space" assumes "affine S" shows "aff_dim(S \ {x. a \ x = b}) = (if S \ {v. a \ v = b} = {} then - 1 else if S \ {v. a \ v = b} then aff_dim S else aff_dim S - 1)" proof (cases "a = 0") case True with assms show ?thesis by auto next case False then have "aff_dim (S \ {x. a \ x = b}) = aff_dim S - 1" if "x \ S" "a \ x \ b" and non: "S \ {v. a \ v = b} \ {}" for x proof - have [simp]: "{x + y| x y. x \ S \ a \ y = b} = UNIV" using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast show ?thesis using aff_dim_sums_Int [OF assms affine_hyperplane non] by (simp add: of_nat_diff False) qed then show ?thesis by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) qed lemma aff_dim_lt_full: fixes S :: "'a::euclidean_space set" shows "aff_dim S < DIM('a) \ (affine hull S \ UNIV)" by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis proof (rule order_antisym) show "aff_dim S \ aff_dim T" by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) next obtain a where "a \ S" using \S \ {}\ by blast have "S \ T" using ope openin_imp_subset by auto then have "a \ T" using \a \ S\ by auto then have subT': "subspace ((\x. - a + x) ` T)" using affine_diffs_subspace \affine T\ by auto then obtain B where Bsub: "B \ ((\x. - a + x) ` T)" and po: "pairwise orthogonal B" and eq1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim ((\x. - a + x) ` T)" and spanB: "span B = ((\x. - a + x) ` T)" by (rule orthonormal_basis_subspace) auto obtain e where "0 < e" and e: "cball a e \ T \ S" by (meson \a \ S\ openin_contains_cball ope) have "aff_dim T = aff_dim ((\x. - a + x) ` T)" by (metis aff_dim_translation_eq) also have "... = dim ((\x. - a + x) ` T)" using aff_dim_subspace subT' by blast also have "... = card B" by (simp add: cardB) also have "... = card ((\x. e *\<^sub>R x) ` B)" using \0 < e\ by (force simp: inj_on_def card_image) also have "... \ dim ((\x. - a + x) ` S)" proof (simp, rule independent_card_le_dim) have e': "cball 0 e \ (\x. x - a) ` T \ (\x. x - a) ` S" using e by (auto simp: dist_norm norm_minus_commute subset_eq) have "(\x. e *\<^sub>R x) ` B \ cball 0 e \ (\x. x - a) ` T" using Bsub \0 < e\ eq1 subT' \a \ T\ by (auto simp: subspace_def) then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast have "inj_on ((*\<^sub>R) e) (span B)" using \0 < e\ inj_on_def by fastforce then show "independent ((\x. e *\<^sub>R x) ` B)" using linear_scale_self \independent B\ linear_dependent_inj_imageD by blast qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by (force cong: image_cong_simp) finally show "aff_dim T \ aff_dim S" . qed qed lemma dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) next have "dim T = aff_dim S" using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" by (metis aff_dim_subset aff_dim_subspace dim_span span_superset subspace_span) finally show "dim T \ dim S" by simp qed subsection\Lower-dimensional affine subsets are nowhere dense\ proposition dense_complement_subspace: fixes S :: "'a :: euclidean_space set" assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S" proof - have "closure(S - U) = S" if "dim U < dim S" "U \ S" for U proof - have "span U \ span S" by (metis neq_iff psubsetI span_eq_dim span_mono that) then obtain a where "a \ 0" "a \ span S" and a: "\y. y \ span U \ orthogonal a y" using orthogonal_to_subspace_exists_gen by metis show ?thesis proof have "closed S" by (simp add: \subspace S\ closed_subspace) then show "closure (S - U) \ S" by (simp add: closure_minimal) show "S \ closure (S - U)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" show "\y\S - U. dist y x < e" proof (cases "x \ U") case True let ?y = "x + (e/2 / norm a) *\<^sub>R a" show ?thesis proof show "dist ?y x < e" using \0 < e\ by (simp add: dist_norm) next have "?y \ S" by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base) qed ultimately show "?y \ S - U" by blast qed next case False with \0 < e\ \x \ S\ show ?thesis by force qed qed qed qed moreover have "S - S \ T = S-T" by blast moreover have "dim (S \ T) < dim S" by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le) ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ dense_complement_affine: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S" proof (cases "S \ T = {}") case True then show ?thesis by (metis Diff_triv affine_hull_eq \affine S\ closure_same_affine_hull closure_subset hull_subset subset_antisym) next case False then obtain z where z: "z \ S \ T" by blast then have "subspace ((+) (- z) ` S)" by (meson IntD1 affine_diffs_subspace \affine S\) moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))" thm aff_dim_eq_dim using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp) ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)" by (simp add: dense_complement_subspace) then show ?thesis by (metis closure_translation translation_diff translation_invert) qed corollary\<^marker>\tag unimportant\ dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" by blast then have "closure (S \ closure (affine hull S - T)) = closure (S \ (affine hull S - T))" by (rule closure_openin_Int_closure [OF ope]) then show ?thesis by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less) qed corollary\<^marker>\tag unimportant\ dense_complement_convex: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" shows "closure(S - T) = closure S" proof show "closure (S - T) \ closure S" by (simp add: closure_mono) have "closure (rel_interior S - T) = closure (rel_interior S)" by (simp add: assms dense_complement_openin_affine_hull openin_rel_interior rel_interior_aff_dim rel_interior_same_affine_hull) then show "closure S \ closure (S - T)" by (metis Diff_mono \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed corollary\<^marker>\tag unimportant\ dense_complement_convex_closed: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" "closed S" shows "closure(S - T) = S" by (simp add: assms dense_complement_convex) subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ If we take a slice out of a set, we can do it perpendicularly, with the normal vector to the slice parallel to the affine hull.\ proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" "\w. w \ S \ (w + a') \ S" proof (cases "S \ {x. a \ x = b} = {}") case True then obtain u v where "u \ S" "v \ S" "a \ u \ b" "a \ v > b" using assms by (auto simp: not_le) define \ where "\ = u + ((b - a \ u) / (a \ v - a \ u)) *\<^sub>R (v - u)" have "\ \ S" by (simp add: \_def \u \ S\ \v \ S\ \affine S\ mem_affine_3_minus) moreover have "a \ \ = b" using \a \ u \ b\ \b < a \ v\ by (simp add: \_def algebra_simps) (simp add: field_simps) ultimately have False using True by force then show ?thesis .. next case False then obtain z where "z \ S" and z: "a \ z = b" using assms by auto with affine_diffs_subspace [OF \affine S\] have sub: "subspace ((+) (- z) ` S)" by blast then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)" by (auto simp: subspace_imp_affine) obtain a' a'' where a': "a' \ span ((+) (- z) ` S)" and a: "a = a' + a''" and "\w. w \ span ((+) (- z) ` S) \ orthogonal a'' w" using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis then have "\w. w \ S \ a'' \ (w-z) = 0" by (simp add: span_base orthogonal_def) then have a'': "\w. w \ S \ a'' \ w = (a - a') \ z" by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) show ?thesis proof (cases "a' = 0") case True with a assms True a'' diff_zero less_irrefl show ?thesis by auto next case False show ?thesis proof show "S \ {x. a' \ x \ a' \ z} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = a' \ z} = S \ {x. a \ x = b}" by (auto simp: a ba'' inner_left_distrib) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" by (metis subspace_add a' span_eq_iff sub) then show "\w. w \ S \ (w + a') \ S" by fastforce qed (use False in auto) qed qed lemma diffs_affine_hull_span: assumes "a \ S" shows "(\x. x - a) ` (affine hull S) = span ((\x. x - a) ` S)" proof - have *: "((\x. x - a) ` (S - {a})) = ((\x. x - a) ` S) - {0}" by (auto simp: algebra_simps) show ?thesis by (auto simp add: algebra_simps affine_hull_span2 [OF assms] *) qed lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" shows "aff_dim S = dim ((\x. x - a) ` S)" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" and card: "of_nat (card B) = aff_dim S + 1" using aff_dim_basis_exists by blast then have "B \ {}" using assms by (metis affine_hull_eq_empty ex_in_conv) then obtain c where "c \ B" by auto then have "c \ S" by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) have *: "(\x. x - c) ` S = (\x. x - a) ` S" using assms \c \ S\ by (auto simp: image_iff xy; metis mem_affine_3_minus pth_1) have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp also have "... = dim ((\x. x - c) ` B)" using affine_independent_card_dim_diffs [OF ind \c \ B\] by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim ((\x. x - c) ` (affine hull B))" by (simp add: diffs_affine_hull_span \c \ B\) also have "... = dim ((\x. x - a) ` S)" by (simp add: affS aff *) finally show ?thesis . qed lemma aff_dim_linear_image_le: assumes "linear f" shows "aff_dim(f ` S) \ aff_dim S" proof - have "aff_dim (f ` T) \ aff_dim T" if "affine T" for T proof (cases "T = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto have 2: "{x - f a| x. x \ f ` T} = f ` ((\x. x - a) ` T)" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) also have "... = int (dim (f ` ((\x. x - a) ` T)))" by (force simp: linear_diff [OF assms] 2) also have "... \ int (dim ((\x. x - a) ` T))" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) finally show ?thesis . qed then have "aff_dim (f ` (affine hull S)) \ aff_dim (affine hull S)" using affine_affine_hull [of S] by blast then show ?thesis using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce qed lemma aff_dim_injective_linear_image [simp]: assumes "linear f" "inj f" shows "aff_dim (f ` S) = aff_dim S" proof (rule antisym) show "aff_dim (f ` S) \ aff_dim S" by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" by (simp add: \linear g\ aff_dim_linear_image_le) finally show "aff_dim S \ aff_dim (f ` S)" . qed lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" proof (cases "d = -1 \ S={}") case True with assms show ?thesis by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) next case False with assms obtain a where "a \ S" "0 \ d" by auto with assms have ss: "subspace ((+) (- a) ` S)" by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp) have "nat d \ dim ((+) (- a) ` S)" by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) then obtain T where "subspace T" and Tsb: "T \ span ((+) (- a) ` S)" and Tdim: "dim T = nat d" using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast then have "affine T" using subspace_affine by blast then have "affine ((+) a ` T)" by (metis affine_hull_eq affine_hull_translation) moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed moreover have "aff_dim ((+) a ` T) = d" by (simp add: aff_dim_subspace Tdim \0 \ d\ \subspace T\ aff_dim_translation_eq) ultimately show ?thesis by (rule that) qed subsection\Paracompactness\ proposition paracompact: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" obtains \' where "S \ \ \'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. x \ S \ \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" proof (cases "S = {}") case True with that show ?thesis by blast next case False have "\T U. x \ U \ open U \ closure U \ T \ T \ \" if "x \ S" for x proof - obtain T where "x \ T" "T \ \" "open T" using assms \x \ S\ by blast then obtain e where "e > 0" "cball x e \ T" by (force simp: open_contains_cball) then show ?thesis by (meson open_ball \T \ \\ ball_subset_cball centre_in_ball closed_cball closure_minimal dual_order.trans) qed then obtain F G where Gin: "x \ G x" and oG: "open (G x)" and clos: "closure (G x) \ F x" and Fin: "F x \ \" if "x \ S" for x by metis then obtain \ where "\ \ G ` S" "countable \" "\\ = \(G ` S)" using Lindelof [of "G ` S"] by (metis image_iff) then obtain K where K: "K \ S" "countable K" and eq: "\(G ` K) = \(G ` S)" by (metis countable_subset_image) with False Gin have "K \ {}" by force then obtain a :: "nat \ 'a" where "range a = K" by (metis range_from_nat_into \countable K\) then have odif: "\n. open (F (a n) - \{closure (G (a m)) |m. m < n})" using \K \ S\ Fin opC by (fastforce simp add:) let ?C = "range (\n. F(a n) - \{closure(G(a m)) |m. m < n})" have enum_S: "\n. x \ F(a n) \ x \ G(a n)" if "x \ S" for x proof - have "\y \ K. x \ G y" using eq that Gin by fastforce then show ?thesis using clos K \range a = K\ closure_subset by blast qed show ?thesis proof show "S \ Union ?C" proof fix x assume "x \ S" define n where "n \ LEAST n. x \ F(a n)" have n: "x \ F(a n)" using enum_S [OF \x \ S\] by (force simp: n_def intro: LeastI) have notn: "x \ F(a m)" if "m < n" for m using that not_less_Least by (force simp: n_def) then have "x \ \{closure (G (a m)) |m. m < n}" using n \K \ S\ \range a = K\ clos notn by fastforce with n show "x \ Union ?C" by blast qed show "\U. U \ ?C \ open U \ (\T. T \ \ \ U \ T)" using Fin \K \ S\ \range a = K\ by (auto simp: odif) show "\V. open V \ x \ V \ finite {U. U \ ?C \ (U \ V \ {})}" if "x \ S" for x proof - obtain n where n: "x \ F(a n)" "x \ G(a n)" using \x \ S\ enum_S by auto have "{U \ ?C. U \ G (a n) \ {}} \ (\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n" proof clarsimp fix k assume "(F (a k) - \{closure (G (a m)) |m. m < k}) \ G (a n) \ {}" then have "k \ n" by auto (metis closure_subset not_le subsetCE) then show "F (a k) - \{closure (G (a m)) |m. m < k} \ (\n. F (a n) - \{closure (G (a m)) |m. m < n}) ` {..n}" by force qed moreover have "finite ((\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n)" by force ultimately have *: "finite {U \ ?C. U \ G (a n) \ {}}" using finite_subset by blast have "a n \ S" using \K \ S\ \range a = K\ by blast then show ?thesis by (blast intro: oG n *) qed qed qed corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes cin: "closedin (top_of_set U) S" and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T using oin [OF that] by (auto simp: openin_open) then obtain F where opF: "open (F T)" and intF: "U \ F T = T" if "T \ \" for T by metis obtain K where K: "closed K" "U \ K = S" using cin by (auto simp: closedin_closed) have 1: "U \ \(insert (- K) (F ` \))" by clarsimp (metis Int_iff Union_iff \U \ K = S\ \S \ \\\ subsetD intF) have 2: "\T. T \ insert (- K) (F ` \) \ open T" using \closed K\ by (auto simp: opF) obtain \ where "U \ \\" and D1: "\U. U \ \ \ open U \ (\T. T \ insert (- K) (F ` \) \ U \ T)" and D2: "\x. x \ U \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" by (blast intro: paracompact [OF 1 2]) let ?C = "{U \ V |V. V \ \ \ (V \ K \ {})}" show ?thesis proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x proof - from D2 [OF that] obtain V where "open V" "x \ V" "finite {U \ \. U \ V \ {}}" by auto with * show ?thesis by (rule_tac x="U \ V" in exI) (auto intro: that finite_subset [OF *]) qed qed qed corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" and "S \ \\" obtains \' where "S \ \\'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) = (S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis unfolding eq by (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) auto qed lemma continuous_closed_graph_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast qed lemma continuous_closed_graph: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "closed S" and contf: "continuous_on S f" shows "closed ((\x. Pair x (f x)) ` S)" proof (rule closedin_closed_trans) show "closedin (top_of_set (S \ UNIV)) ((\x. (x, f x)) ` S)" by (rule continuous_closed_graph_gen [OF contf subset_UNIV]) qed (simp add: \closed S\ closed_Times) lemma continuous_from_closed_graph: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" and clo: "closed ((\x. Pair x (f x)) ` S)" shows "continuous_on S f" using fim clo by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" by (simp_all add: fg) then have "continuous_on S (\x. if P x then f x else g x)" "continuous_on T (\x. if P x then f x else g x)" by (simp_all add: contf contg cong: continuous_on_cong) then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" proof (cases "a \ S") case True with xS insert_Diff that show ?thesis by fastforce next case False show ?thesis proof (cases "finite S \ card S \ Suc (DIM('a))") case True then obtain u where u0: "\i. i \ S \ 0 \ u i" and u1: "sum u S = 1" and ua: "(\i\S. u i *\<^sub>R i) = a" using a by (auto simp: convex_hull_finite) obtain v where v0: "\i. i \ S \ 0 \ v i" and v1: "sum v S = 1" and vx: "(\i\S. v i *\<^sub>R i) = x" using True xS by (auto simp: convex_hull_finite) show ?thesis proof (cases "\b. b \ S \ v b = 0") case True then obtain b where b: "b \ S" "v b = 0" by blast show ?thesis proof have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x = a then 0 else v x) = (\x \ S - {b}. if x = a then 0 else v x)" using fin by (force intro: sum.mono_neutral_right) also have "... = (\x \ S - {b}. v x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x)" by (metis \v b = 0\ diff_zero sum.infinite sum_diff1 u1 zero_neq_one) finally show "(\x\insert a (S - {b}). if x = a then 0 else v x) = 1" by (simp add: v1) show "\x. x \ insert a (S - {b}) \ 0 \ (if x = a then 0 else v x)" by (auto simp: v0) have "(\x \ insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = (\x \ S - {b}. (if x = a then 0 else v x) *\<^sub>R x)" using fin by (force intro: sum.mono_neutral_right) also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) next case False have le_Max: "u i / v i \ Max ((\i. u i / v i) ` S)" if "i \ S" for i by (simp add: True that) have "Max ((\i. u i / v i) ` S) \ (\i. u i / v i) ` S" using True v1 by (auto intro: Max_in) then obtain b where "b \ S" and beq: "Max ((\b. u b / v b) ` S) = u b / v b" by blast then have "0 \ u b / v b" using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1 by (metis False eq_iff v0) then have "0 < u b" "0 < v b" using False \b \ S\ u0 v0 by force+ have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show ?thesis proof show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) = v b / u b + (\x \ S - {b}. v x - (v b / u b) * u x)" using \a \ S\ \b \ S\ True by (auto intro!: sum.cong split: if_split_asm) also have "... = v b / u b + (\x \ S - {b}. v x) - (v b / u b) * (\x \ S - {b}. u x)" by (simp add: Groups_Big.sum_subtractf sum_distrib_left) also have "... = (\x\S. v x)" using \0 < u b\ True by (simp add: Groups_Big.sum_diff1 u1 field_simps) finally show "sum (\x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1" by (simp add: v1) show "0 \ (if i = a then v b / u b else v i - v b / u b * u i)" if "i \ insert a (S - {b})" for i using \0 < u b\ \0 < v b\ v0 [of i] le_Max [of i] beq that False by (auto simp: field_simps split: if_split_asm) have "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = (v b / u b) *\<^sub>R a + (\x\S - {b}. (v x - v b / u b * u x) *\<^sub>R x)" using \a \ S\ \b \ S\ True by (auto intro!: sum.cong split: if_split_asm) also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally show "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) qed next case False obtain T where "finite T" "T \ S" and caT: "card T \ Suc (DIM('a))" and xT: "x \ convex hull T" using xS by (auto simp: caratheodory [of S]) with False obtain b where b: "b \ S" "b \ T" by (metis antisym subsetI) show ?thesis proof show "x \ convex hull insert a (S - {b})" using \T \ S\ b by (blast intro: subsetD [OF hull_mono xT]) qed (rule \b \ S\) qed qed lemma convex_hull_exchange_Union: fixes a :: "'a::euclidean_space" assumes "a \ convex hull S" shows "convex hull S = (\b \ S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (blast intro: in_convex_hull_exchange [OF assms]) show "?rhs \ ?lhs" proof clarify fix x b assume"b \ S" "x \ convex hull insert a (S - {b})" then show "x \ convex hull S" if "b \ S" by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE) qed qed lemma Un_closed_segment: fixes a :: "'a::euclidean_space" assumes "b \ closed_segment a c" shows "closed_segment a b \ closed_segment b c = closed_segment a c" proof (cases "c = a") case True with assms show ?thesis by simp next case False with assms have "convex hull {a, b} \ convex hull {b, c} = (\ba\{a, c}. convex hull insert b ({a, c} - {ba}))" by (auto simp: insert_Diff_if insert_commute) then show ?thesis using convex_hull_exchange_Union by (metis assms segment_convex_hull) qed lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" (is "?lhs = ?rhs") proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) have *: "?rhs \ insert b (open_segment a b \ open_segment b c)" if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ ?rhs" proof - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (?rhs))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis proof show "?lhs \ ?rhs" by (simp add: assms b subset_open_segment) show "?rhs \ ?lhs" using Un_closed_segment [OF b] * by (simp add: closed_segment_eq_open insert_commute) qed qed subsection\Covering an open set by a countable chain of compact sets\ proposition open_Union_compact_subsets: fixes S :: "'a::euclidean_space set" assumes "open S" obtains C where "\n. compact(C n)" "\n. C n \ S" "\n. C n \ interior(C(Suc n))" "\(range C) = S" "\K. \compact K; K \ S\ \ \N. \n\N. K \ (C n)" proof (cases "S = {}") case True then show ?thesis by (rule_tac C = "\n. {}" in that) auto next case False then obtain a where "a \ S" by auto let ?C = "\n. cball a (real n) - (\x \ -S. \e \ ball 0 (1 / real(Suc n)). {x + e})" have "\N. \n\N. K \ (f n)" if "\n. compact(f n)" and sub_int: "\n. f n \ interior (f(Suc n))" and eq: "\(range f) = S" and "compact K" "K \ S" for f K proof - have *: "\n. f n \ (\n. interior (f n))" by (meson Sup_upper2 UNIV_I \\n. f n \ interior (f (Suc n))\ image_iff) have mono: "\m n. m \ n \f m \ f n" by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int) obtain I where "finite I" and I: "K \ (\i\I. interior (f i))" proof (rule compactE_image [OF \compact K\]) show "K \ (\n. interior (f n))" using \K \ S\ \\(f ` UNIV) = S\ * by blast qed auto { fix n assume n: "Max I \ n" have "(\i\I. interior (f i)) \ f n" by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \finite I\] n) then have "K \ f n" using I by auto } then show ?thesis by blast qed moreover have "\f. (\n. compact(f n)) \ (\n. (f n) \ S) \ (\n. (f n) \ interior(f(Suc n))) \ ((\(range f) = S))" proof (intro exI conjI allI) show "\n. compact (?C n)" by (auto simp: compact_diff open_sums) show "\n. ?C n \ S" by auto show "?C n \ interior (?C (Suc n))" for n proof (simp add: interior_diff, rule Diff_mono) show "cball a (real n) \ ball a (1 + real n)" by (simp add: cball_subset_ball_iff) have cl: "closed (\x\- S. \e\cball 0 (1 / (2 + real n)). {x + e})" using assms by (auto intro: closed_compact_sums) have "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" by (simp add: cball_subset_ball_iff field_split_simps UN_mono) finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed have "S \ \ (range ?C)" proof fix x assume x: "x \ S" then obtain e where "e > 0" and e: "ball x e \ S" using assms open_contains_ball by blast then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e" using reals_Archimedean2 by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff) obtain N2 where N2: "norm(x - a) \ real N2" by (meson real_arch_simple) have N12: "inverse((N1 + N2) + 1) \ inverse(N1)" using \N1 > 0\ by (auto simp: field_split_simps) have "x \ y + z" if "y \ S" "norm z < 1 / (1 + (real N1 + real N2))" for y z proof - have "e * real N1 < e * (1 + (real N1 + real N2))" by (simp add: \0 < e\) then have "1 / (1 + (real N1 + real N2)) < e" using N1 \e > 0\ by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc) then have "x - z \ ball x e" using that by simp then have "x - z \ S" using e by blast with that show ?thesis by auto qed with N2 show "x \ \ (range ?C)" by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) qed then show "\ (range ?C) = S" by auto qed ultimately show ?thesis using that by metis qed subsection\Orthogonal complement\ definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" unfolding subspace_def orthogonal_comp_def orthogonal_def by (auto simp: inner_right_distrib) lemma orthogonal_comp_anti_mono: assumes "A \ B" shows "B\<^sup>\ \ A\<^sup>\" proof fix x assume x: "x \ B\<^sup>\" show "x \ orthogonal_comp A" using x unfolding orthogonal_comp_def by (simp add: orthogonal_def, metis assms in_mono) qed lemma orthogonal_comp_null [simp]: "{0}\<^sup>\ = UNIV" by (auto simp: orthogonal_comp_def orthogonal_def) lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\ = {0}" unfolding orthogonal_comp_def orthogonal_def by auto (use inner_eq_zero_iff in blast) lemma orthogonal_comp_subset: "U \ U\<^sup>\\<^sup>\" by (auto simp: orthogonal_comp_def orthogonal_def inner_commute) lemma subspace_sum_minimal: assumes "S \ U" "T \ U" "subspace U" shows "S + T \ U" proof fix x assume "x \ S + T" then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" by (meson set_plus_elim) then show "x \ U" by (meson assms subsetCE subspace_add) qed proposition subspace_sum_orthogonal_comp: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U + U\<^sup>\ = UNIV" proof - obtain B where "B \ U" and ortho: "pairwise orthogonal B" "\x. x \ B \ norm x = 1" and "independent B" "card B = dim U" "span B = U" using orthonormal_basis_subspace [OF assms] by metis then have "finite B" by (simp add: indep_card_eq_dim_span) have *: "\x\B. \y\B. x \ y = (if x=y then 1 else 0)" using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def) { fix v let ?u = "\b\B. (v \ b) *\<^sub>R b" have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" by (metis (no_types, lifting) \span B = U\ assms subspace_sum span_base span_mul) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y assume "y \ U" with \span B = U\ span_finite [OF \finite B\] obtain u where u: "y = (\b\B. u b *\<^sub>R b)" by auto have "b \ (v - ?u) = 0" if "b \ B" for b using that \finite B\ by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong) then show "y \ (v - ?u) = 0" by (simp add: u inner_sum_left) qed ultimately have "v \ U + U\<^sup>\" using set_plus_intro by fastforce } then show ?thesis by auto qed lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" using orthogonal_comp_def orthogonal_self by (force simp: assms subspace_0 subspace_orthogonal_comp) lemma orthogonal_comp_self: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U\<^sup>\\<^sup>\ = U" proof have ssU': "subspace (U\<^sup>\)" by (simp add: subspace_orthogonal_comp) have "u \ U" if "u \ U\<^sup>\\<^sup>\" for u proof - obtain v w where "u = v+w" "v \ U" "w \ U\<^sup>\" using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast then have "u-v \ U\<^sup>\" by simp moreover have "v \ U\<^sup>\\<^sup>\" using \v \ U\ orthogonal_comp_subset by blast then have "u-v \ U\<^sup>\\<^sup>\" by (simp add: subspace_diff subspace_orthogonal_comp that) ultimately have "u-v = 0" using orthogonal_Int_0 ssU' by blast with \v \ U\ show ?thesis by auto qed then show "U\<^sup>\\<^sup>\ \ U" by auto qed (use orthogonal_comp_subset in auto) lemma ker_orthogonal_comp_adjoint: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" proof - have "\x. \\y. y \ f x = 0\ \ f x = 0" using assms inner_commute all_zero_iff by metis then show ?thesis using assms by (auto simp: orthogonal_comp_def orthogonal_def adjoint_works inner_commute) qed subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" "surj (adjoint f)" shows "inj f" proof - have "\x. y = adjoint f x" for y using assms by (simp add: surjD) then show "inj f" using assms unfolding inj_on_def image_def by (metis (no_types) adjoint_works euclidean_eqI) qed \ \\<^url>\https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\\ lemma surj_adjoint_iff_inj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "surj (adjoint f) \ inj f" proof assume "surj (adjoint f)" then show "inj f" by (simp add: assms linear_surj_adj_imp_inj) next assume "inj f" have "f -` {0} = {0}" using assms \inj f\ linear_0 linear_injective_0 by fastforce moreover have "f -` {0} = range (adjoint f)\<^sup>\" by (intro ker_orthogonal_comp_adjoint assms) ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" using adjoint_linear \linear f\ by (subst (asm) orthogonal_comp_self) (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "inj (adjoint f) \ surj f" proof assume "inj (adjoint f)" have "(adjoint f) -` {0} = {0}" by (metis \inj (adjoint f)\ adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV) then have "(range(f))\<^sup>\ = {0}" by (metis (no_types, hide_lams) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero) then show "surj f" by (metis \inj (adjoint f)\ adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj) next assume "surj f" then have "range f = (adjoint f -` {0})\<^sup>\" by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint) then have "{0} = adjoint f -` {0}" using \surj f\ adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force then show "inj (adjoint f)" by (simp add: \surj f\ adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj) qed lemma linear_singular_into_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" shows "\ inj f \ (\a. a \ 0 \ (\x. a \ f x = 0))" (is "_ = ?rhs") proof assume "\inj f" then show ?rhs using all_zero_iff by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" by (metis assms linear_injective_isomorphism all_zero_iff) qed lemma linear_singular_image_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\inj f" obtains a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms by (fastforce simp add: linear_singular_into_hyperplane) end diff --git a/src/HOL/Analysis/Tagged_Division.thy b/src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy +++ b/src/HOL/Analysis/Tagged_Division.thy @@ -1,2602 +1,2602 @@ (* Title: HOL/Analysis/Tagged_Division.thy Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP *) section \Tagged Divisions for Henstock-Kurzweil Integration\ theory Tagged_Division imports Topology_Euclidean_Space begin lemma sum_Sigma_product: assumes "finite S" and "\i. i \ S \ finite (T i)" shows "(\i\S. sum (x i) (T i)) = (\(i, j)\Sigma S T. x i j)" using assms proof induct case empty then show ?case by simp next case (insert a S) show ?case by (simp add: insert.hyps(1) insert.prems sum.Sigma) qed lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one subsection\<^marker>\tag unimportant\ \Sundries\ text\A transitive relation is well-founded if all initial segments are finite.\ lemma wf_finite_segments: assumes "irrefl r" and "trans r" and "\x. finite {y. (y, x) \ r}" shows "wf (r)" apply (simp add: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms) using acyclic_def assms irrefl_def trans_Restr by fastforce text\For creating values between \<^term>\u\ and \<^term>\v\.\ lemma scaling_mono: fixes u::"'a::linordered_field" assumes "u \ v" "0 \ r" "r \ s" shows "u + r * (v - u) / s \ v" proof - have "r/s \ 1" using assms using divide_le_eq_1 by fastforce then have "(r/s) * (v - u) \ 1 * (v - u)" by (meson diff_ge_0_iff_ge mult_right_mono \u \ v\) then show ?thesis by (simp add: field_simps) qed subsection \Some useful lemmas about intervals\ lemma interior_subset_union_intervals: assumes "i = cbox a b" and "j = cbox c d" and "interior j \ {}" and "i \ j \ S" and "interior i \ interior j = {}" shows "interior i \ interior S" proof - have "box a b \ cbox c d = {}" using Int_interval_mixed_eq_empty[of c d a b] assms unfolding interior_cbox by auto moreover have "box a b \ cbox c d \ S" apply (rule order_trans,rule box_subset_cbox) using assms by auto ultimately show ?thesis unfolding assms interior_cbox by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI) qed lemma interior_Union_subset_cbox: assumes "finite f" assumes f: "\s. s \ f \ \a b. s = cbox a b" "\s. s \ f \ interior s \ t" and t: "closed t" shows "interior (\f) \ t" proof - have [simp]: "s \ f \ closed s" for s using f by auto define E where "E = {s\f. interior s = {}}" then have "finite E" "E \ {s\f. interior s = {}}" using \finite f\ by auto then have "interior (\f) = interior (\(f - E))" proof (induction E rule: finite_subset_induct') case (insert s f') have "interior (\(f - insert s f') \ s) = interior (\(f - insert s f'))" using insert.hyps \finite f\ by (intro interior_closed_Un_empty_interior) auto also have "\(f - insert s f') \ s = \(f - f')" using insert.hyps by auto finally show ?case by (simp add: insert.IH) qed simp also have "\ \ \(f - E)" by (rule interior_subset) also have "\ \ t" proof (rule Union_least) fix s assume "s \ f - E" with f[of s] obtain a b where s: "s \ f" "s = cbox a b" "box a b \ {}" by (fastforce simp: E_def) have "closure (interior s) \ closure t" by (intro closure_mono f \s \ f\) with s \closed t\ show "s \ t" by simp qed finally show ?thesis . qed lemma Int_interior_Union_intervals: "\finite \; open S; \T. T\\ \ \a b. T = cbox a b; \T. T\\ \ S \ (interior T) = {}\ \ S \ interior (\\) = {}" using interior_Union_subset_cbox[of \ "UNIV - S"] by auto lemma interval_split: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "cbox a b \ {x. x\k \ c} = cbox a (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)" "cbox a b \ {x. x\k \ c} = cbox (\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) b" using assms by (rule_tac set_eqI; auto simp: mem_box)+ lemma interval_not_empty: "\i\Basis. a\i \ b\i \ cbox a b \ {}" by (simp add: box_ne_empty) subsection \Bounds on intervals where they exist\ definition\<^marker>\tag important\ interval_upperbound :: "('a::euclidean_space) set \ 'a" where "interval_upperbound s = (\i\Basis. (SUP x\s. x\i) *\<^sub>R i)" definition\<^marker>\tag important\ interval_lowerbound :: "('a::euclidean_space) set \ 'a" where "interval_lowerbound s = (\i\Basis. (INF x\s. x\i) *\<^sub>R i)" lemma interval_upperbound[simp]: "\i\Basis. a\i \ b\i \ interval_upperbound (cbox a b) = (b::'a::euclidean_space)" unfolding interval_upperbound_def euclidean_representation_sum cbox_def by (safe intro!: cSup_eq) auto lemma interval_lowerbound[simp]: "\i\Basis. a\i \ b\i \ interval_lowerbound (cbox a b) = (a::'a::euclidean_space)" unfolding interval_lowerbound_def euclidean_representation_sum cbox_def by (safe intro!: cInf_eq) auto lemmas interval_bounds = interval_upperbound interval_lowerbound lemma fixes X::"real set" shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X" and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X" by (auto simp: interval_upperbound_def interval_lowerbound_def) lemma interval_bounds'[simp]: assumes "cbox a b \ {}" shows "interval_upperbound (cbox a b) = b" and "interval_lowerbound (cbox a b) = a" using assms unfolding box_ne_empty by auto lemma interval_upperbound_Times: assumes "A \ {}" and "B \ {}" shows "interval_upperbound (A \ B) = (interval_upperbound A, interval_upperbound B)" proof- from assms have fst_image_times': "A = fst ` (A \ B)" by simp have "(\i\Basis. (SUP x\A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (SUP x\A. x \ i) *\<^sub>R i)" by (subst (2) fst_image_times') (simp del: fst_image_times add: image_comp inner_Pair_0) moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp have "(\i\Basis. (SUP x\A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (SUP x\B. x \ i) *\<^sub>R i)" by (subst (2) snd_image_times') (simp del: snd_image_times add: image_comp inner_Pair_0) ultimately show ?thesis unfolding interval_upperbound_def by (subst sum_Basis_prod_eq) (auto simp add: sum_prod) qed lemma interval_lowerbound_Times: assumes "A \ {}" and "B \ {}" shows "interval_lowerbound (A \ B) = (interval_lowerbound A, interval_lowerbound B)" proof- from assms have fst_image_times': "A = fst ` (A \ B)" by simp have "(\i\Basis. (INF x\A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (INF x\A. x \ i) *\<^sub>R i)" by (subst (2) fst_image_times') (simp del: fst_image_times add: image_comp inner_Pair_0) moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp have "(\i\Basis. (INF x\A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (INF x\B. x \ i) *\<^sub>R i)" by (subst (2) snd_image_times') (simp del: snd_image_times add: image_comp inner_Pair_0) ultimately show ?thesis unfolding interval_lowerbound_def by (subst sum_Basis_prod_eq) (auto simp add: sum_prod) qed subsection \The notion of a gauge --- simply an open set containing the point\ definition\<^marker>\tag important\ "gauge \ \ (\x. x \ \ x \ open (\ x))" lemma gaugeI: assumes "\x. x \ \ x" and "\x. open (\ x)" shows "gauge \" using assms unfolding gauge_def by auto lemma gaugeD[dest]: assumes "gauge \" shows "x \ \ x" and "open (\ x)" using assms unfolding gauge_def by auto lemma gauge_ball_dependent: "\x. 0 < e x \ gauge (\x. ball x (e x))" unfolding gauge_def by auto lemma gauge_ball[intro]: "0 < e \ gauge (\x. ball x e)" unfolding gauge_def by auto lemma gauge_trivial[intro!]: "gauge (\x. ball x 1)" by (rule gauge_ball) auto lemma gauge_Int[intro]: "gauge \1 \ gauge \2 \ gauge (\x. \1 x \ \2 x)" unfolding gauge_def by auto lemma gauge_reflect: fixes \ :: "'a::euclidean_space \ 'a set" shows "gauge \ \ gauge (\x. uminus ` \ (- x))" using equation_minus_iff by (auto simp: gauge_def surj_def intro!: open_surjective_linear_image linear_uminus) lemma gauge_Inter: assumes "finite S" and "\u. u\S \ gauge (f u)" shows "gauge (\x. \{f u x | u. u \ S})" proof - have *: "\x. {f u x |u. u \ S} = (\u. f u x) ` S" by auto show ?thesis unfolding gauge_def unfolding * using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto qed lemma gauge_existence_lemma: "(\x. \d :: real. p x \ 0 < d \ q d x) \ (\x. \d>0. p x \ q d x)" by (metis zero_less_one) subsection \Attempt a systematic general set of "offset" results for components\ (*FIXME Restructure, the subsection consists only of 1 lemma *) lemma gauge_modify: assumes "(\S. open S \ open {x. f(x) \ S})" "gauge d" shows "gauge (\x. {y. f y \ d (f x)})" using assms unfolding gauge_def by force subsection \Divisions\ definition\<^marker>\tag important\ division_of (infixl "division'_of" 40) where "s division_of i \ finite s \ (\K\s. K \ i \ K \ {} \ (\a b. K = cbox a b)) \ (\K1\s. \K2\s. K1 \ K2 \ interior(K1) \ interior(K2) = {}) \ (\s = i)" lemma division_ofD[dest]: assumes "s division_of i" shows "finite s" and "\K. K \ s \ K \ i" and "\K. K \ s \ K \ {}" and "\K. K \ s \ \a b. K = cbox a b" and "\K1 K2. K1 \ s \ K2 \ s \ K1 \ K2 \ interior(K1) \ interior(K2) = {}" and "\s = i" using assms unfolding division_of_def by auto lemma division_ofI: assumes "finite s" and "\K. K \ s \ K \ i" and "\K. K \ s \ K \ {}" and "\K. K \ s \ \a b. K = cbox a b" and "\K1 K2. K1 \ s \ K2 \ s \ K1 \ K2 \ interior K1 \ interior K2 = {}" and "\s = i" shows "s division_of i" using assms unfolding division_of_def by auto lemma division_of_finite: "s division_of i \ finite s" by auto lemma division_of_self[intro]: "cbox a b \ {} \ {cbox a b} division_of (cbox a b)" unfolding division_of_def by auto lemma division_of_trivial[simp]: "s division_of {} \ s = {}" unfolding division_of_def by auto lemma division_of_sing[simp]: "s division_of cbox a (a::'a::euclidean_space) \ s = {cbox a a}" (is "?l = ?r") proof assume ?r moreover { fix K assume "s = {{a}}" "K\s" then have "\x y. K = cbox x y" apply (rule_tac x=a in exI)+ apply force done } ultimately show ?l unfolding division_of_def cbox_sing by auto next assume ?l have "x = {a}" if "x \ s" for x by (metis \s division_of cbox a a\ cbox_sing division_ofD(2) division_ofD(3) subset_singletonD that) moreover have "s \ {}" using \s division_of cbox a a\ by auto ultimately show ?r unfolding cbox_sing by auto qed lemma elementary_empty: obtains p where "p division_of {}" unfolding division_of_trivial by auto lemma elementary_interval: obtains p where "p division_of (cbox a b)" by (metis division_of_trivial division_of_self) lemma division_contains: "s division_of i \ \x\i. \k\s. x \ k" unfolding division_of_def by auto lemma forall_in_division: "d division_of i \ (\x\d. P x) \ (\a b. cbox a b \ d \ P (cbox a b))" unfolding division_of_def by fastforce lemma cbox_division_memE: assumes \: "K \ \" "\ division_of S" obtains a b where "K = cbox a b" "K \ {}" "K \ S" using assms unfolding division_of_def by metis lemma division_of_subset: assumes "p division_of (\p)" and "q \ p" shows "q division_of (\q)" proof (rule division_ofI) note * = division_ofD[OF assms(1)] show "finite q" using "*"(1) assms(2) infinite_super by auto { fix k assume "k \ q" then have kp: "k \ p" using assms(2) by auto show "k \ \q" using \k \ q\ by auto show "\a b. k = cbox a b" using *(4)[OF kp] by auto show "k \ {}" using *(3)[OF kp] by auto } fix k1 k2 assume "k1 \ q" "k2 \ q" "k1 \ k2" then have **: "k1 \ p" "k2 \ p" "k1 \ k2" using assms(2) by auto show "interior k1 \ interior k2 = {}" using *(5)[OF **] by auto qed auto lemma division_of_union_self[intro]: "p division_of s \ p division_of (\p)" unfolding division_of_def by auto lemma division_inter: fixes s1 s2 :: "'a::euclidean_space set" assumes "p1 division_of s1" and "p2 division_of s2" shows "{k1 \ k2 | k1 k2. k1 \ p1 \ k2 \ p2 \ k1 \ k2 \ {}} division_of (s1 \ s2)" (is "?A' division_of _") proof - let ?A = "{s. s \ (\(k1,k2). k1 \ k2) ` (p1 \ p2) \ s \ {}}" have *: "?A' = ?A" by auto show ?thesis unfolding * proof (rule division_ofI) have "?A \ (\(x, y). x \ y) ` (p1 \ p2)" by auto moreover have "finite (p1 \ p2)" using assms unfolding division_of_def by auto ultimately show "finite ?A" by auto have *: "\s. \{x\s. x \ {}} = \s" by auto show "\?A = s1 \ s2" unfolding * using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto { fix k assume "k \ ?A" then obtain k1 k2 where k: "k = k1 \ k2" "k1 \ p1" "k2 \ p2" "k \ {}" by auto then show "k \ {}" by auto show "k \ s1 \ s2" using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)] unfolding k by auto obtain a1 b1 where k1: "k1 = cbox a1 b1" using division_ofD(4)[OF assms(1) k(2)] by blast obtain a2 b2 where k2: "k2 = cbox a2 b2" using division_ofD(4)[OF assms(2) k(3)] by blast show "\a b. k = cbox a b" unfolding k k1 k2 unfolding Int_interval by auto } fix k1 k2 assume "k1 \ ?A" then obtain x1 y1 where k1: "k1 = x1 \ y1" "x1 \ p1" "y1 \ p2" "k1 \ {}" by auto assume "k2 \ ?A" then obtain x2 y2 where k2: "k2 = x2 \ y2" "x2 \ p1" "y2 \ p2" "k2 \ {}" by auto assume "k1 \ k2" then have th: "x1 \ x2 \ y1 \ y2" unfolding k1 k2 by auto have *: "interior x1 \ interior x2 = {} \ interior y1 \ interior y2 = {} \ interior (x1 \ y1) \ interior x1 \ interior (x1 \ y1) \ interior y1 \ interior (x2 \ y2) \ interior x2 \ interior (x2 \ y2) \ interior y2 \ interior (x1 \ y1) \ interior (x2 \ y2) = {}" by auto show "interior k1 \ interior k2 = {}" unfolding k1 k2 apply (rule *) using assms division_ofD(5) k1 k2(2) k2(3) th apply auto done qed qed lemma division_inter_1: assumes "d division_of i" and "cbox a (b::'a::euclidean_space) \ i" shows "{cbox a b \ k | k. k \ d \ cbox a b \ k \ {}} division_of (cbox a b)" proof (cases "cbox a b = {}") case True show ?thesis unfolding True and division_of_trivial by auto next case False have *: "cbox a b \ i = cbox a b" using assms(2) by auto show ?thesis using division_inter[OF division_of_self[OF False] assms(1)] unfolding * by auto qed lemma elementary_Int: fixes s t :: "'a::euclidean_space set" assumes "p1 division_of s" and "p2 division_of t" shows "\p. p division_of (s \ t)" using assms division_inter by blast lemma elementary_Inter: assumes "finite f" and "f \ {}" and "\s\f. \p. p division_of (s::('a::euclidean_space) set)" shows "\p. p division_of (\f)" using assms proof (induct f rule: finite_induct) case (insert x f) show ?case proof (cases "f = {}") case True then show ?thesis unfolding True using insert by auto next case False obtain p where "p division_of \f" using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] .. moreover obtain px where "px division_of x" using insert(5)[rule_format,OF insertI1] .. ultimately show ?thesis by (simp add: elementary_Int Inter_insert) qed qed auto lemma division_disjoint_union: assumes "p1 division_of s1" and "p2 division_of s2" and "interior s1 \ interior s2 = {}" shows "(p1 \ p2) division_of (s1 \ s2)" proof (rule division_ofI) note d1 = division_ofD[OF assms(1)] note d2 = division_ofD[OF assms(2)] show "finite (p1 \ p2)" using d1(1) d2(1) by auto show "\(p1 \ p2) = s1 \ s2" using d1(6) d2(6) by auto { fix k1 k2 assume as: "k1 \ p1 \ p2" "k2 \ p1 \ p2" "k1 \ k2" moreover let ?g="interior k1 \ interior k2 = {}" { assume as: "k1\p1" "k2\p2" have ?g using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]] using assms(3) by blast } moreover { assume as: "k1\p2" "k2\p1" have ?g using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]] using assms(3) by blast } ultimately show ?g using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto } fix k assume k: "k \ p1 \ p2" show "k \ s1 \ s2" using k d1(2) d2(2) by auto show "k \ {}" using k d1(3) d2(3) by auto show "\a b. k = cbox a b" using k d1(4) d2(4) by auto qed lemma partial_division_extend_1: fixes a b c d :: "'a::euclidean_space" assumes incl: "cbox c d \ cbox a b" and nonempty: "cbox c d \ {}" obtains p where "p division_of (cbox a b)" "cbox c d \ p" proof let ?B = "\f::'a\'a \ 'a. cbox (\i\Basis. (fst (f i) \ i) *\<^sub>R i) (\i\Basis. (snd (f i) \ i) *\<^sub>R i)" define p where "p = ?B ` (Basis \\<^sub>E {(a, c), (c, d), (d, b)})" show "cbox c d \ p" unfolding p_def by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) have ord: "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" if "i \ Basis" for i using incl nonempty that unfolding box_eq_empty subset_box by (auto simp: not_le) show "p division_of (cbox a b)" proof (rule division_ofI) show "finite p" unfolding p_def by (auto intro!: finite_PiE) { fix k assume "k \ p" then obtain f where f: "f \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" by (auto simp: p_def) then show "\a b. k = cbox a b" by auto have "k \ cbox a b \ k \ {}" proof (simp add: k box_eq_empty subset_box not_less, safe) fix i :: 'a assume i: "i \ Basis" with f have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" by (auto simp: PiE_iff) with i ord[of i] show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" by auto qed then show "k \ {}" "k \ cbox a b" by auto { fix l assume "l \ p" then obtain g where g: "g \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g" by (auto simp: p_def) assume "l \ k" have "\i\Basis. f i \ g i" proof (rule ccontr) assume "\ ?thesis" with f g have "f = g" by (auto simp: PiE_iff extensional_def fun_eq_iff) with \l \ k\ show False by (simp add: l k) qed then obtain i where *: "i \ Basis" "f i \ g i" .. then have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" using f g by (auto simp: PiE_iff) with * ord[of i] show "interior l \ interior k = {}" by (auto simp add: l k disjoint_interval intro!: bexI[of _ i]) } note \k \ cbox a b\ } moreover { fix x assume x: "x \ cbox a b" have "\i\Basis. \l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" proof fix i :: 'a assume "i \ Basis" with x ord[of i] have "(a \ i \ x \ i \ x \ i \ c \ i) \ (c \ i \ x \ i \ x \ i \ d \ i) \ (d \ i \ x \ i \ x \ i \ b \ i)" by (auto simp: cbox_def) then show "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" by auto qed then obtain f where f: "\i\Basis. x \ i \ {fst (f i) \ i..snd (f i) \ i} \ f i \ {(a, c), (c, d), (d, b)}" unfolding bchoice_iff .. moreover from f have "restrict f Basis \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" by auto moreover from f have "x \ ?B (restrict f Basis)" by (auto simp: mem_box) ultimately have "\k\p. x \ k" unfolding p_def by blast } ultimately show "\p = cbox a b" by auto qed qed proposition partial_division_extend_interval: assumes "p division_of (\p)" "(\p) \ cbox a b" obtains q where "p \ q" "q division_of cbox a (b::'a::euclidean_space)" proof (cases "p = {}") case True obtain q where "q division_of (cbox a b)" by (rule elementary_interval) then show ?thesis using True that by blast next case False note p = division_ofD[OF assms(1)] have div_cbox: "\k\p. \q. q division_of cbox a b \ k \ q" proof fix k assume kp: "k \ p" obtain c d where k: "k = cbox c d" using p(4)[OF kp] by blast have *: "cbox c d \ cbox a b" "cbox c d \ {}" using p(2,3)[OF kp, unfolded k] using assms(2) by (blast intro: order.trans)+ obtain q where "q division_of cbox a b" "cbox c d \ q" by (rule partial_division_extend_1[OF *]) then show "\q. q division_of cbox a b \ k \ q" unfolding k by auto qed obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" using bchoice[OF div_cbox] by blast have "q x division_of \(q x)" if x: "x \ p" for x apply (rule division_ofI) using division_ofD[OF q(1)[OF x]] by auto then have di: "\x. x \ p \ \d. d division_of \(q x - {x})" by (meson Diff_subset division_of_subset) have "\d. d division_of \((\i. \(q i - {i})) ` p)" apply (rule elementary_Inter [OF finite_imageI[OF p(1)]]) apply (auto dest: di simp: False elementary_Inter [OF finite_imageI[OF p(1)]]) done then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" .. have "d \ p division_of cbox a b" proof - have te: "\S f T. S \ {} \ \i\S. f i \ i = T \ T = \(f ` S) \ \S" by auto have cbox_eq: "cbox a b = \((\i. \(q i - {i})) ` p) \ \p" proof (rule te[OF False], clarify) fix i assume i: "i \ p" show "\(q i - {i}) \ i = cbox a b" using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto qed { fix K assume K: "K \ p" note qk=division_ofD[OF q(1)[OF K]] have *: "\u T S. T \ S = {} \ u \ S \ u \ T = {}" by auto have "interior (\i\p. \(q i - {i})) \ interior K = {}" proof (rule *[OF Int_interior_Union_intervals]) show "\T. T\q K - {K} \ interior K \ interior T = {}" using qk(5) using q(2)[OF K] by auto show "interior (\i\p. \(q i - {i})) \ interior (\(q K - {K}))" apply (rule interior_mono)+ using K by auto qed (use qk in auto)} note [simp] = this show "d \ p division_of (cbox a b)" unfolding cbox_eq apply (rule division_disjoint_union[OF d assms(1)]) apply (rule Int_interior_Union_intervals) apply (rule p open_interior ballI)+ apply simp_all done qed then show ?thesis by (meson Un_upper2 that) qed lemma elementary_bounded[dest]: fixes S :: "'a::euclidean_space set" shows "p division_of S \ bounded S" unfolding division_of_def by (metis bounded_Union bounded_cbox) lemma elementary_subset_cbox: "p division_of S \ \a b. S \ cbox a (b::'a::euclidean_space)" by (meson bounded_subset_cbox_symmetric elementary_bounded) proposition division_union_intervals_exists: fixes a b :: "'a::euclidean_space" assumes "cbox a b \ {}" obtains p where "(insert (cbox a b) p) division_of (cbox a b \ cbox c d)" proof (cases "cbox c d = {}") case True with assms that show ?thesis by force next case False show ?thesis proof (cases "cbox a b \ cbox c d = {}") case True then show ?thesis by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty) next case False obtain u v where uv: "cbox a b \ cbox c d = cbox u v" unfolding Int_interval by auto have uv_sub: "cbox u v \ cbox c d" using uv by auto obtain p where pd: "p division_of cbox c d" and "cbox u v \ p" by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]]) note p = this division_ofD[OF pd] have "interior (cbox a b \ \(p - {cbox u v})) = interior(cbox u v \ \(p - {cbox u v}))" apply (rule arg_cong[of _ _ interior]) using p(8) uv by auto also have "\ = {}" unfolding interior_Int apply (rule Int_interior_Union_intervals) using p(6) p(7)[OF p(2)] \finite p\ apply auto done finally have [simp]: "interior (cbox a b) \ interior (\(p - {cbox u v})) = {}" by simp have cbe: "cbox a b \ cbox c d = cbox a b \ \(p - {cbox u v})" using p(8) unfolding uv[symmetric] by auto have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" proof - have "{cbox a b} division_of cbox a b" by (simp add: assms division_of_self) then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" by (metis (no_types) Diff_subset \interior (cbox a b) \ interior (\(p - {cbox u v})) = {}\ division_disjoint_union division_of_subset insert_is_Un p(1) p(8)) qed with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe) qed qed lemma division_of_Union: assumes "finite f" and "\p. p \ f \ p division_of (\p)" and "\k1 k2. k1 \ \f \ k2 \ \f \ k1 \ k2 \ interior k1 \ interior k2 = {}" shows "\f division_of \(\f)" using assms by (auto intro!: division_ofI) lemma elementary_union_interval: fixes a b :: "'a::euclidean_space" assumes "p division_of \p" obtains q where "q division_of (cbox a b \ \p)" proof (cases "p = {} \ cbox a b = {}") case True obtain p where "p division_of (cbox a b)" by (rule elementary_interval) then show thesis using True assms that by auto next case False then have "p \ {}" "cbox a b \ {}" by auto note pdiv = division_ofD[OF assms] show ?thesis proof (cases "interior (cbox a b) = {}") case True show ?thesis apply (rule that [of "insert (cbox a b) p", OF division_ofI]) using pdiv(1-4) True False apply auto apply (metis IntI equals0D pdiv(5)) by (metis disjoint_iff_not_equal pdiv(5)) next case False have "\K\p. \q. (insert (cbox a b) q) division_of (cbox a b \ K)" proof fix K assume kp: "K \ p" from pdiv(4)[OF kp] obtain c d where "K = cbox c d" by blast then show "\q. (insert (cbox a b) q) division_of (cbox a b \ K)" by (meson \cbox a b \ {}\ division_union_intervals_exists) qed from bchoice[OF this] obtain q where "\x\p. insert (cbox a b) (q x) division_of (cbox a b) \ x" .. note q = division_ofD[OF this[rule_format]] let ?D = "\{insert (cbox a b) (q K) | K. K \ p}" show thesis proof (rule that[OF division_ofI]) have *: "{insert (cbox a b) (q K) |K. K \ p} = (\K. insert (cbox a b) (q K)) ` p" by auto show "finite ?D" using "*" pdiv(1) q(1) by auto have "\?D = (\x\p. \(insert (cbox a b) (q x)))" by auto also have "... = \{cbox a b \ t |t. t \ p}" using q(6) by auto also have "... = cbox a b \ \p" using \p \ {}\ by auto finally show "\?D = cbox a b \ \p" . show "K \ cbox a b \ \p" "K \ {}" if "K \ ?D" for K using q that by blast+ show "\a b. K = cbox a b" if "K \ ?D" for K using q(4) that by auto next fix K' K assume K: "K \ ?D" and K': "K' \ ?D" "K \ K'" obtain x where x: "K \ insert (cbox a b) (q x)" "x\p" using K by auto obtain x' where x': "K'\insert (cbox a b) (q x')" "x'\p" using K' by auto show "interior K \ interior K' = {}" proof (cases "x = x' \ K = cbox a b \ K' = cbox a b") case True then show ?thesis using True K' q(5) x' x by auto next case False then have as': "K \ cbox a b" "K' \ cbox a b" by auto obtain c d where K: "K = cbox c d" using q(4) x by blast have "interior K \ interior (cbox a b) = {}" using as' K'(2) q(5) x by blast then have "interior K \ interior x" by (metis \interior (cbox a b) \ {}\ K q(2) x interior_subset_union_intervals) moreover obtain c d where c_d: "K' = cbox c d" using q(4)[OF x'(2,1)] by blast have "interior K' \ interior (cbox a b) = {}" using as'(2) q(5) x' by blast then have "interior K' \ interior x'" by (metis \interior (cbox a b) \ {}\ c_d interior_subset_union_intervals q(2) x'(1) x'(2)) moreover have "interior x \ interior x' = {}" by (meson False assms division_ofD(5) x'(2) x(2)) ultimately show ?thesis using \interior K \ interior x\ \interior K' \ interior x'\ by auto qed qed qed qed lemma elementary_unions_intervals: assumes fin: "finite f" and "\s. s \ f \ \a b. s = cbox a (b::'a::euclidean_space)" obtains p where "p division_of (\f)" proof - have "\p. p division_of (\f)" proof (induct_tac f rule:finite_subset_induct) show "\p. p division_of \{}" using elementary_empty by auto next fix x F assume as: "finite F" "x \ F" "\p. p division_of \F" "x\f" from this(3) obtain p where p: "p division_of \F" .. from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast have *: "\F = \p" using division_ofD[OF p] by auto show "\p. p division_of \(insert x F)" using elementary_union_interval[OF p[unfolded *], of a b] unfolding Union_insert x * by metis qed (insert assms, auto) then show ?thesis using that by auto qed lemma elementary_union: fixes S T :: "'a::euclidean_space set" assumes "ps division_of S" "pt division_of T" obtains p where "p division_of (S \ T)" proof - have *: "S \ T = \ps \ \pt" using assms unfolding division_of_def by auto show ?thesis apply (rule elementary_unions_intervals[of "ps \ pt"]) using assms apply auto by (simp add: * that) qed lemma partial_division_extend: fixes T :: "'a::euclidean_space set" assumes "p division_of S" and "q division_of T" and "S \ T" obtains r where "p \ r" and "r division_of T" proof - note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)] obtain a b where ab: "T \ cbox a b" using elementary_subset_cbox[OF assms(2)] by auto obtain r1 where "p \ r1" "r1 division_of (cbox a b)" using assms by (metis ab dual_order.trans partial_division_extend_interval divp(6)) note r1 = this division_ofD[OF this(2)] obtain p' where "p' division_of \(r1 - p)" apply (rule elementary_unions_intervals[of "r1 - p"]) using r1(3,6) apply auto done then obtain r2 where r2: "r2 division_of (\(r1 - p)) \ (\q)" by (metis assms(2) divq(6) elementary_Int) { fix x assume x: "x \ T" "x \ S" then obtain R where r: "R \ r1" "x \ R" unfolding r1 using ab by (meson division_contains r1(2) subsetCE) moreover have "R \ p" proof assume "R \ p" then have "x \ S" using divp(2) r by auto then show False using x by auto qed ultimately have "x\\(r1 - p)" by auto } then have Teq: "T = \p \ (\(r1 - p) \ \q)" unfolding divp divq using assms(3) by auto have "interior S \ interior (\(r1-p)) = {}" proof (rule Int_interior_Union_intervals) have *: "\S. (\x. x \ S \ False) \ S = {}" by auto show "interior S \ interior m = {}" if "m \ r1 - p" for m proof - have "interior m \ interior (\p) = {}" proof (rule Int_interior_Union_intervals) show "\T. T\p \ interior m \ interior T = {}" by (metis DiffD1 DiffD2 that r1(1) r1(7) rev_subsetD) qed (use divp in auto) then show "interior S \ interior m = {}" unfolding divp by auto qed qed (use r1 in auto) then have "interior S \ interior (\(r1-p) \ (\q)) = {}" using interior_subset by auto then have div: "p \ r2 division_of \p \ \(r1 - p) \ \q" by (simp add: assms(1) division_disjoint_union divp(6) r2) show ?thesis apply (rule that[of "p \ r2"]) apply (auto simp: div Teq) done qed lemma division_split: fixes a :: "'a::euclidean_space" assumes "p division_of (cbox a b)" and k: "k\Basis" shows "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of(cbox a b \ {x. x\k \ c})" (is "?p1 division_of ?I1") and "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of (cbox a b \ {x. x\k \ c})" (is "?p2 division_of ?I2") proof (rule_tac[!] division_ofI) note p = division_ofD[OF assms(1)] show "finite ?p1" "finite ?p2" using p(1) by auto show "\?p1 = ?I1" "\?p2 = ?I2" unfolding p(6)[symmetric] by auto { fix K assume "K \ ?p1" then obtain l where l: "K = l \ {x. x \ k \ c}" "l \ p" "l \ {x. x \ k \ c} \ {}" by blast obtain u v where uv: "l = cbox u v" using assms(1) l(2) by blast show "K \ ?I1" using l p(2) uv by force show "K \ {}" by (simp add: l) show "\a b. K = cbox a b" apply (simp add: l uv p(2-3)[OF l(2)]) apply (subst interval_split[OF k]) apply (auto intro: order.trans) done fix K' assume "K' \ ?p1" then obtain l' where l': "K' = l' \ {x. x \ k \ c}" "l' \ p" "l' \ {x. x \ k \ c} \ {}" by blast assume "K \ K'" then show "interior K \ interior K' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto } { fix K assume "K \ ?p2" then obtain l where l: "K = l \ {x. c \ x \ k}" "l \ p" "l \ {x. c \ x \ k} \ {}" by blast obtain u v where uv: "l = cbox u v" using l(2) p(4) by blast show "K \ ?I2" using l p(2) uv by force show "K \ {}" by (simp add: l) show "\a b. K = cbox a b" apply (simp add: l uv p(2-3)[OF l(2)]) apply (subst interval_split[OF k]) apply (auto intro: order.trans) done fix K' assume "K' \ ?p2" then obtain l' where l': "K' = l' \ {x. c \ x \ k}" "l' \ p" "l' \ {x. c \ x \ k} \ {}" by blast assume "K \ K'" then show "interior K \ interior K' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto } qed subsection \Tagged (partial) divisions\ definition\<^marker>\tag important\ tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) where "s tagged_partial_division_of i \ finite s \ (\x K. (x, K) \ s \ x \ K \ K \ i \ (\a b. K = cbox a b)) \ (\x1 K1 x2 K2. (x1, K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {})" lemma tagged_partial_division_ofD: assumes "s tagged_partial_division_of i" shows "finite s" and "\x K. (x,K) \ s \ x \ K" and "\x K. (x,K) \ s \ K \ i" and "\x K. (x,K) \ s \ \a b. K = cbox a b" and "\x1 K1 x2 K2. (x1,K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" using assms unfolding tagged_partial_division_of_def by blast+ definition\<^marker>\tag important\ tagged_division_of (infixr "tagged'_division'_of" 40) where "s tagged_division_of i \ s tagged_partial_division_of i \ (\{K. \x. (x,K) \ s} = i)" lemma tagged_division_of_finite: "s tagged_division_of i \ finite s" unfolding tagged_division_of_def tagged_partial_division_of_def by auto lemma tagged_division_of: "s tagged_division_of i \ finite s \ (\x K. (x, K) \ s \ x \ K \ K \ i \ (\a b. K = cbox a b)) \ (\x1 K1 x2 K2. (x1, K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}) \ (\{K. \x. (x,K) \ s} = i)" unfolding tagged_division_of_def tagged_partial_division_of_def by auto lemma tagged_division_ofI: assumes "finite s" and "\x K. (x,K) \ s \ x \ K" and "\x K. (x,K) \ s \ K \ i" and "\x K. (x,K) \ s \ \a b. K = cbox a b" and "\x1 K1 x2 K2. (x1,K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" and "(\{K. \x. (x,K) \ s} = i)" shows "s tagged_division_of i" unfolding tagged_division_of using assms by fastforce lemma tagged_division_ofD[dest]: (*FIXME USE A LOCALE*) assumes "s tagged_division_of i" shows "finite s" and "\x K. (x,K) \ s \ x \ K" and "\x K. (x,K) \ s \ K \ i" and "\x K. (x,K) \ s \ \a b. K = cbox a b" and "\x1 K1 x2 K2. (x1, K1) \ s \ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" and "(\{K. \x. (x,K) \ s} = i)" using assms unfolding tagged_division_of by blast+ lemma division_of_tagged_division: assumes "s tagged_division_of i" shows "(snd ` s) division_of i" proof (rule division_ofI) note assm = tagged_division_ofD[OF assms] show "\(snd ` s) = i" "finite (snd ` s)" using assm by auto fix k assume k: "k \ snd ` s" then obtain xk where xk: "(xk, k) \ s" by auto then show "k \ i" "k \ {}" "\a b. k = cbox a b" using assm by fastforce+ fix k' assume k': "k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk': "(xk', k') \ s" by auto then show "interior k \ interior k' = {}" using assm(5) k'(2) xk by blast qed lemma partial_division_of_tagged_division: assumes "s tagged_partial_division_of i" shows "(snd ` s) division_of \(snd ` s)" proof (rule division_ofI) note assm = tagged_partial_division_ofD[OF assms] show "finite (snd ` s)" "\(snd ` s) = \(snd ` s)" using assm by auto fix k assume k: "k \ snd ` s" then obtain xk where xk: "(xk, k) \ s" by auto then show "k \ {}" "\a b. k = cbox a b" "k \ \(snd ` s)" using assm by auto fix k' assume k': "k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk': "(xk', k') \ s" by auto then show "interior k \ interior k' = {}" using assm(5) k'(2) xk by auto qed lemma tagged_partial_division_subset: assumes "s tagged_partial_division_of i" and "t \ s" shows "t tagged_partial_division_of i" using assms finite_subset[OF assms(2)] unfolding tagged_partial_division_of_def by blast lemma tag_in_interval: "p tagged_division_of i \ (x, k) \ p \ x \ i" by auto lemma tagged_division_of_empty: "{} tagged_division_of {}" unfolding tagged_division_of by auto lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \ p = {}" unfolding tagged_partial_division_of_def by auto lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \ p = {}" unfolding tagged_division_of by auto lemma tagged_division_of_self: "x \ cbox a b \ {(x,cbox a b)} tagged_division_of (cbox a b)" by (rule tagged_division_ofI) auto lemma tagged_division_of_self_real: "x \ {a .. b::real} \ {(x,{a .. b})} tagged_division_of {a .. b}" unfolding box_real[symmetric] by (rule tagged_division_of_self) lemma tagged_division_Un: assumes "p1 tagged_division_of s1" and "p2 tagged_division_of s2" and "interior s1 \ interior s2 = {}" shows "(p1 \ p2) tagged_division_of (s1 \ s2)" proof (rule tagged_division_ofI) note p1 = tagged_division_ofD[OF assms(1)] note p2 = tagged_division_ofD[OF assms(2)] show "finite (p1 \ p2)" using p1(1) p2(1) by auto show "\{k. \x. (x, k) \ p1 \ p2} = s1 \ s2" using p1(6) p2(6) by blast fix x k assume xk: "(x, k) \ p1 \ p2" show "x \ k" "\a b. k = cbox a b" using xk p1(2,4) p2(2,4) by auto show "k \ s1 \ s2" using xk p1(3) p2(3) by blast fix x' k' assume xk': "(x', k') \ p1 \ p2" "(x, k) \ (x', k')" have *: "\a b. a \ s1 \ b \ s2 \ interior a \ interior b = {}" using assms(3) interior_mono by blast show "interior k \ interior k' = {}" apply (cases "(x, k) \ p1") apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2)) by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2)) qed lemma tagged_division_Union: assumes "finite I" and tag: "\i. i\I \ pfn i tagged_division_of i" and disj: "\i1 i2. \i1 \ I; i2 \ I; i1 \ i2\ \ interior(i1) \ interior(i2) = {}" shows "\(pfn ` I) tagged_division_of (\I)" proof (rule tagged_division_ofI) note assm = tagged_division_ofD[OF tag] show "finite (\(pfn ` I))" using assms by auto have "\{k. \x. (x, k) \ \(pfn ` I)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` I)" by blast also have "\ = \I" using assm(6) by auto finally show "\{k. \x. (x, k) \ \(pfn ` I)} = \I" . fix x k assume xk: "(x, k) \ \(pfn ` I)" then obtain i where i: "i \ I" "(x, k) \ pfn i" by auto show "x \ k" "\a b. k = cbox a b" "k \ \I" using assm(2-4)[OF i] using i(1) by auto fix x' k' assume xk': "(x', k') \ \(pfn ` I)" "(x, k) \ (x', k')" then obtain i' where i': "i' \ I" "(x', k') \ pfn i'" by auto have *: "\a b. i \ i' \ a \ i \ b \ i' \ interior a \ interior b = {}" using i(1) i'(1) disj interior_mono by blast show "interior k \ interior k' = {}" proof (cases "i = i'") case True then show ?thesis using assm(5) i' i xk'(2) by blast next case False then show ?thesis using "*" assm(3) i' i by auto qed qed lemma tagged_partial_division_of_Union_self: assumes "p tagged_partial_division_of s" shows "p tagged_division_of (\(snd ` p))" apply (rule tagged_division_ofI) using tagged_partial_division_ofD[OF assms] apply auto done lemma tagged_division_of_union_self: assumes "p tagged_division_of s" shows "p tagged_division_of (\(snd ` p))" apply (rule tagged_division_ofI) using tagged_division_ofD[OF assms] apply auto done lemma tagged_division_Un_interval: fixes a :: "'a::euclidean_space" assumes "p1 tagged_division_of (cbox a b \ {x. x\k \ (c::real)})" and "p2 tagged_division_of (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" shows "(p1 \ p2) tagged_division_of (cbox a b)" proof - have *: "cbox a b = (cbox a b \ {x. x\k \ c}) \ (cbox a b \ {x. x\k \ c})" by auto show ?thesis apply (subst *) apply (rule tagged_division_Un[OF assms(1-2)]) unfolding interval_split[OF k] interior_cbox using k apply (auto simp add: box_def elim!: ballE[where x=k]) done qed lemma tagged_division_Un_interval_real: fixes a :: real assumes "p1 tagged_division_of ({a .. b} \ {x. x\k \ (c::real)})" and "p2 tagged_division_of ({a .. b} \ {x. x\k \ c})" and k: "k \ Basis" shows "(p1 \ p2) tagged_division_of {a .. b}" using assms unfolding box_real[symmetric] by (rule tagged_division_Un_interval) lemma tagged_division_split_left_inj: assumes d: "d tagged_division_of i" and tags: "(x1, K1) \ d" "(x2, K2) \ d" and "K1 \ K2" and eq: "K1 \ {x. x \ k \ c} = K2 \ {x. x \ k \ c}" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior (K1 \ K2) = {} \ (x2, K2) = (x1, K1)" using tags d by (metis (no_types) interior_Int tagged_division_ofD(5)) then show ?thesis using eq \K1 \ K2\ by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject) qed lemma tagged_division_split_right_inj: assumes d: "d tagged_division_of i" and tags: "(x1, K1) \ d" "(x2, K2) \ d" and "K1 \ K2" and eq: "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior (K1 \ K2) = {} \ (x2, K2) = (x1, K1)" using tags d by (metis (no_types) interior_Int tagged_division_ofD(5)) then show ?thesis using eq \K1 \ K2\ by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject) qed lemma (in comm_monoid_set) over_tagged_division_lemma: assumes "p tagged_division_of i" and "\u v. box u v = {} \ d (cbox u v) = \<^bold>1" shows "F (\(_, k). d k) p = F d (snd ` p)" proof - have *: "(\(_ ,k). d k) = d \ snd" by (simp add: fun_eq_iff) note assm = tagged_division_ofD[OF assms(1)] show ?thesis unfolding * proof (rule reindex_nontrivial[symmetric]) show "finite p" using assm by auto fix x y assume "x\p" "y\p" "x\y" "snd x = snd y" obtain a b where ab: "snd x = cbox a b" using assm(4)[of "fst x" "snd x"] \x\p\ by auto have "(fst x, snd y) \ p" "(fst x, snd y) \ y" using \x \ p\ \x \ y\ \snd x = snd y\ [symmetric] by auto with \x\p\ \y\p\ have "interior (snd x) \ interior (snd y) = {}" by (intro assm(5)[of "fst x" _ "fst y"]) auto then have "box a b = {}" unfolding \snd x = snd y\[symmetric] ab by auto then have "d (cbox a b) = \<^bold>1" using assm(2)[of "fst x" "snd x"] \x\p\ ab[symmetric] by (intro assms(2)) auto then show "d (snd x) = \<^bold>1" unfolding ab by auto qed qed subsection \Functions closed on boxes: morphisms from boxes to monoids\ text \This auxiliary structure is used to sum up over the elements of a division. Main theorem is \operative_division\. Instances for the monoid are \<^typ>\'a option\, \<^typ>\real\, and \<^typ>\bool\.\ paragraph\<^marker>\tag important\ \Using additivity of lifted function to encode definedness.\ text \%whitespace\ definition\<^marker>\tag important\ lift_option :: "('a \ 'b \ 'c) \ 'a option \ 'b option \ 'c option" where "lift_option f a' b' = Option.bind a' (\a. Option.bind b' (\b. Some (f a b)))" lemma lift_option_simps[simp]: "lift_option f (Some a) (Some b) = Some (f a b)" "lift_option f None b' = None" "lift_option f a' None = None" by (auto simp: lift_option_def) lemma\<^marker>\tag important\ comm_monoid_lift_option: assumes "comm_monoid f z" shows "comm_monoid (lift_option f) (Some z)" proof - from assms interpret comm_monoid f z . show ?thesis by standard (auto simp: lift_option_def ac_simps split: bind_split) qed lemma comm_monoid_and: "comm_monoid HOL.conj True" by standard auto lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True" by (rule comm_monoid_set.intro) (fact comm_monoid_and) paragraph \Misc\ lemma interval_real_split: "{a .. b::real} \ {x. x \ c} = {a .. min b c}" "{a .. b} \ {x. c \ x} = {max a c .. b}" apply (metis Int_atLeastAtMostL1 atMost_def) apply (metis Int_atLeastAtMostL2 atLeast_def) done lemma bgauge_existence_lemma: "(\x\s. \d::real. 0 < d \ q d x) \ (\x. \d>0. x\s \ q d x)" by (meson zero_less_one) paragraph \Division points\ text \%whitespace\ definition\<^marker>\tag important\ "division_points (k::('a::euclidean_space) set) d = {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" lemma division_points_finite: fixes i :: "'a::euclidean_space set" assumes "d division_of i" shows "finite (division_points i d)" proof - note assm = division_ofD[OF assms] let ?M = "\j. {(j,x)|x. (interval_lowerbound i)\j < x \ x < (interval_upperbound i)\j \ (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" have *: "division_points i d = \(?M ` Basis)" unfolding division_points_def by auto show ?thesis unfolding * using assm by auto qed lemma division_points_subset: fixes a :: "'a::euclidean_space" assumes "d division_of (cbox a b)" and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and k: "k \ Basis" shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ division_points (cbox a b) d" (is ?t1) and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ \(l \ {x. x\k \ c} = {})} \ division_points (cbox a b) d" (is ?t2) proof - note assm = division_ofD[OF assms(1)] have *: "\i\Basis. a\i \ b\i" "\i\Basis. a\i \ (\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i" "\i\Basis. (\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i \ b\i" "min (b \ k) c = c" "max (a \ k) c = c" using assms using less_imp_le by auto have "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" if "a \ x < y" "y < (if x = k then c else b \ x)" "interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" "x \ Basis" for i l x y proof - obtain u v where l: "l = cbox u v" using \l \ d\ assms(1) by blast have *: "\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" using that(6) unfolding l interval_split[OF k] box_ne_empty that . have **: "\i\Basis. u\i \ v\i" using l using that(6) unfolding box_ne_empty[symmetric] by auto show ?thesis apply (rule bexI[OF _ \l \ d\]) using that(1-3,5) \x \ Basis\ unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that apply (auto split: if_split_asm) done qed moreover have "\x y. \y < (if x = k then c else b \ x)\ \ y < b \ x" using \c < b\k\ by (auto split: if_split_asm) ultimately show ?t1 unfolding division_points_def interval_split[OF k, of a b] unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * by force have "\x y i l. (if x = k then c else a \ x) < y \ a \ x < y" using \a\k < c\ by (auto split: if_split_asm) moreover have "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" if "(if x = k then c else a \ x) < y" "y < b \ x" "interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" "x \ Basis" for x y i l proof - obtain u v where l: "l = cbox u v" using \l \ d\ assm(4) by blast have *: "\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" using that(6) unfolding l interval_split[OF k] box_ne_empty that . have **: "\i\Basis. u\i \ v\i" using l using that(6) unfolding box_ne_empty[symmetric] by auto show "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" apply (rule bexI[OF _ \l \ d\]) using that(1-3,5) \x \ Basis\ unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that apply (auto split: if_split_asm) done qed ultimately show ?t2 unfolding division_points_def interval_split[OF k, of a b] unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * by force qed lemma division_points_psubset: fixes a :: "'a::euclidean_space" assumes d: "d division_of (cbox a b)" and altb: "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and "l \ d" and disj: "interval_lowerbound l\k = c \ interval_upperbound l\k = c" and k: "k \ Basis" shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ division_points (cbox a b) d" (is "?D1 \ ?D") and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ division_points (cbox a b) d" (is "?D2 \ ?D") proof - have ab: "\i\Basis. a\i \ b\i" using altb by (auto intro!:less_imp_le) obtain u v where l: "l = cbox u v" using d \l \ d\ by blast have uv: "\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" apply (metis assms(5) box_ne_empty(1) cbox_division_memE d l) by (metis assms(5) box_ne_empty(1) cbox_division_memE d l subset_box(1)) have *: "interval_upperbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" "interval_upperbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" unfolding l interval_split[OF k] interval_bounds[OF uv(1)] using uv[rule_format, of k] ab k by auto have "\x. x \ ?D - ?D1" using assms(3-) unfolding division_points_def interval_bounds[OF ab] by (force simp add: *) moreover have "?D1 \ ?D" by (auto simp add: assms division_points_subset) ultimately show "?D1 \ ?D" by blast have *: "interval_lowerbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" "interval_lowerbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" unfolding l interval_split[OF k] interval_bounds[OF uv(1)] using uv[rule_format, of k] ab k by auto have "\x. x \ ?D - ?D2" using assms(3-) unfolding division_points_def interval_bounds[OF ab] by (force simp add: *) moreover have "?D2 \ ?D" by (auto simp add: assms division_points_subset) ultimately show "?D2 \ ?D" by blast qed lemma division_split_left_inj: fixes S :: "'a::euclidean_space set" assumes div: "\ division_of S" and eq: "K1 \ {x::'a. x\k \ c} = K2 \ {x. x\k \ c}" and "K1 \ \" "K2 \ \" "K1 \ K2" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior K2 \ interior {a. a \ k \ c} = interior K1 \ interior {a. a \ k \ c}" by (metis (no_types) eq interior_Int) moreover have "\A. interior A \ interior K2 = {} \ A = K2 \ A \ \" by (meson div \K2 \ \\ division_of_def) ultimately show ?thesis using \K1 \ \\ \K1 \ K2\ by auto qed lemma division_split_right_inj: fixes S :: "'a::euclidean_space set" assumes div: "\ division_of S" and eq: "K1 \ {x::'a. x\k \ c} = K2 \ {x. x\k \ c}" and "K1 \ \" "K2 \ \" "K1 \ K2" shows "interior (K1 \ {x. x\k \ c}) = {}" proof - have "interior K2 \ interior {a. a \ k \ c} = interior K1 \ interior {a. a \ k \ c}" by (metis (no_types) eq interior_Int) moreover have "\A. interior A \ interior K2 = {} \ A = K2 \ A \ \" by (meson div \K2 \ \\ division_of_def) ultimately show ?thesis using \K1 \ \\ \K1 \ K2\ by auto qed lemma interval_doublesplit: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "cbox a b \ {x . \x\k - c\ \ (e::real)} = cbox (\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)" proof - have *: "\x c e::real. \x - c\ \ e \ x \ c - e \ x \ c + e" by auto have **: "\s P Q. s \ {x. P x \ Q x} = (s \ {x. Q x}) \ {x. P x}" by blast show ?thesis unfolding * ** interval_split[OF assms] by (rule refl) qed lemma division_doublesplit: fixes a :: "'a::euclidean_space" assumes "p division_of (cbox a b)" and k: "k \ Basis" shows "(\l. l \ {x. \x\k - c\ \ e}) ` {l\p. l \ {x. \x\k - c\ \ e} \ {}} division_of (cbox a b \ {x. \x\k - c\ \ e})" proof - have *: "\x c. \x - c\ \ e \ x \ c - e \ x \ c + e" by auto have **: "\p q p' q'. p division_of q \ p = p' \ q = q' \ p' division_of q'" by auto note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]] note division_split(2)[OF this, where c="c-e" and k=k,OF k] then show ?thesis apply (rule **) subgoal apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image [symmetric] cong: image_cong_simp) apply (rule equalityI) apply blast apply clarsimp apply (rule_tac x="xa \ {x. c + e \ x \ k}" in exI) apply auto done by (simp add: interval_split k interval_doublesplit) qed paragraph \Operative\ locale operative = comm_monoid_set + fixes g :: "'b::euclidean_space set \ 'a" assumes box_empty_imp: "\a b. box a b = {} \ g (cbox a b) = \<^bold>1" and Basis_imp: "\a b c k. k \ Basis \ g (cbox a b) = g (cbox a b \ {x. x\k \ c}) \<^bold>* g (cbox a b \ {x. x\k \ c})" begin lemma empty [simp]: "g {} = \<^bold>1" proof - have *: "cbox One (-One) = ({}::'b set)" by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv) moreover have "box One (-One) = ({}::'b set)" using box_subset_cbox[of One "-One"] by (auto simp: *) ultimately show ?thesis using box_empty_imp [of One "-One"] by simp qed lemma division: "F g d = g (cbox a b)" if "d division_of (cbox a b)" proof - define C where [abs_def]: "C = card (division_points (cbox a b) d)" then show ?thesis using that proof (induction C arbitrary: a b d rule: less_induct) case (less a b d) show ?case proof cases assume "box a b = {}" { fix k assume "k\d" then obtain a' b' where k: "k = cbox a' b'" using division_ofD(4)[OF less.prems] by blast with \k\d\ division_ofD(2)[OF less.prems] have "cbox a' b' \ cbox a b" by auto then have "box a' b' \ box a b" unfolding subset_box by auto then have "g k = \<^bold>1" using box_empty_imp [of a' b'] k by (simp add: \box a b = {}\) } then show "box a b = {} \ F g d = g (cbox a b)" by (auto intro!: neutral simp: box_empty_imp) next assume "box a b \ {}" then have ab: "\i\Basis. a\i < b\i" and ab': "\i\Basis. a\i \ b\i" by (auto simp: box_ne_empty) show "F g d = g (cbox a b)" proof (cases "division_points (cbox a b) d = {}") case True { fix u v and j :: 'b assume j: "j \ Basis" and as: "cbox u v \ d" then have "cbox u v \ {}" using less.prems by blast then have uv: "\i\Basis. u\i \ v\i" "u\j \ v\j" using j unfolding box_ne_empty by auto have *: "\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ Q (cbox u v)" using as j by auto have "(j, u\j) \ division_points (cbox a b) d" "(j, v\j) \ division_points (cbox a b) d" using True by auto note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps] note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] moreover have "a\j \ u\j" "v\j \ b\j" using division_ofD(2,2,3)[OF \d division_of cbox a b\ as] apply (metis j subset_box(1) uv(1)) by (metis \cbox u v \ cbox a b\ j subset_box(1) uv(1)) ultimately have "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force } then have d': "\i\d. \u v. i = cbox u v \ (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" unfolding forall_in_division[OF less.prems] by blast have "(1/2) *\<^sub>R (a+b) \ cbox a b" unfolding mem_box using ab by (auto simp: inner_simps) note this[unfolded division_ofD(6)[OF \d division_of cbox a b\,symmetric] Union_iff] then obtain i where i: "i \ d" "(1 / 2) *\<^sub>R (a + b) \ i" .. obtain u v where uv: "i = cbox u v" "\j\Basis. u \ j = a \ j \ v \ j = a \ j \ u \ j = b \ j \ v \ j = b \ j \ u \ j = a \ j \ v \ j = b \ j" using d' i(1) by auto have "cbox a b \ d" proof - have "u = a" "v = b" unfolding euclidean_eq_iff[where 'a='b] proof safe fix j :: 'b assume j: "j \ Basis" note i(2)[unfolded uv mem_box,rule_format,of j] then show "u \ j = a \ j" and "v \ j = b \ j" using uv(2)[rule_format,of j] j by (auto simp: inner_simps) qed then have "i = cbox a b" using uv by auto then show ?thesis using i by auto qed then have deq: "d = insert (cbox a b) (d - {cbox a b})" by auto have "F g (d - {cbox a b}) = \<^bold>1" proof (intro neutral ballI) fix x assume x: "x \ d - {cbox a b}" then have "x\d" by auto note d'[rule_format,OF this] then obtain u v where uv: "x = cbox u v" "\j\Basis. u \ j = a \ j \ v \ j = a \ j \ u \ j = b \ j \ v \ j = b \ j \ u \ j = a \ j \ v \ j = b \ j" by blast have "u \ a \ v \ b" using x[unfolded uv] by auto then obtain j where "u\j \ a\j \ v\j \ b\j" and j: "j \ Basis" unfolding euclidean_eq_iff[where 'a='b] by auto then have "u\j = v\j" using uv(2)[rule_format,OF j] by auto then have "box u v = {}" using j unfolding box_eq_empty by (auto intro!: bexI[of _ j]) then show "g x = \<^bold>1" unfolding uv(1) by (rule box_empty_imp) qed then show "F g d = g (cbox a b)" using division_ofD[OF less.prems] apply (subst deq) apply (subst insert) apply auto done next case False then have "\x. x \ division_points (cbox a b) d" by auto then obtain k c where "k \ Basis" "interval_lowerbound (cbox a b) \ k < c" "c < interval_upperbound (cbox a b) \ k" "\i\d. interval_lowerbound i \ k = c \ interval_upperbound i \ k = c" unfolding division_points_def by auto then obtain j where "a \ k < c" "c < b \ k" and "j \ d" and j: "interval_lowerbound j \ k = c \ interval_upperbound j \ k = c" by (metis division_of_trivial empty_iff interval_bounds' less.prems) let ?lec = "{x. x\k \ c}" let ?gec = "{x. x\k \ c}" define d1 where "d1 = {l \ ?lec | l. l \ d \ l \ ?lec \ {}}" define d2 where "d2 = {l \ ?gec | l. l \ d \ l \ ?gec \ {}}" define cb where "cb = (\i\Basis. (if i = k then c else b\i) *\<^sub>R i)" define ca where "ca = (\i\Basis. (if i = k then c else a\i) *\<^sub>R i)" have "division_points (cbox a b \ ?lec) {l \ ?lec |l. l \ d \ l \ ?lec \ {}} \ division_points (cbox a b) d" by (rule division_points_psubset[OF \d division_of cbox a b\ ab \a \ k < c\ \c < b \ k\ \j \ d\ j \k \ Basis\]) with division_points_finite[OF \d division_of cbox a b\] have "card (division_points (cbox a b \ ?lec) {l \ ?lec |l. l \ d \ l \ ?lec \ {}}) < card (division_points (cbox a b) d)" by (rule psubset_card_mono) moreover have "division_points (cbox a b \ {x. c \ x \ k}) {l \ {x. c \ x \ k} |l. l \ d \ l \ {x. c \ x \ k} \ {}} \ division_points (cbox a b) d" by (rule division_points_psubset[OF \d division_of cbox a b\ ab \a \ k < c\ \c < b \ k\ \j \ d\ j \k \ Basis\]) with division_points_finite[OF \d division_of cbox a b\] have "card (division_points (cbox a b \ ?gec) {l \ ?gec |l. l \ d \ l \ ?gec \ {}}) < card (division_points (cbox a b) d)" by (rule psubset_card_mono) ultimately have *: "F g d1 = g (cbox a b \ ?lec)" "F g d2 = g (cbox a b \ ?gec)" unfolding interval_split[OF \k \ Basis\] apply (rule_tac[!] less.hyps) using division_split[OF \d division_of cbox a b\, where k=k and c=c] \k \ Basis\ by (simp_all add: interval_split d1_def d2_def division_points_finite[OF \d division_of cbox a b\]) have fxk_le: "g (l \ ?lec) = \<^bold>1" if "l \ d" "y \ d" "l \ ?lec = y \ ?lec" "l \ y" for l y proof - obtain u v where leq: "l = cbox u v" using \l \ d\ less.prems by auto have "interior (cbox u v \ ?lec) = {}" using that division_split_left_inj leq less.prems by blast then show ?thesis unfolding leq interval_split [OF \k \ Basis\] by (auto intro: box_empty_imp) qed have fxk_ge: "g (l \ {x. x \ k \ c}) = \<^bold>1" if "l \ d" "y \ d" "l \ ?gec = y \ ?gec" "l \ y" for l y proof - obtain u v where leq: "l = cbox u v" using \l \ d\ less.prems by auto have "interior (cbox u v \ ?gec) = {}" using that division_split_right_inj leq less.prems by blast then show ?thesis unfolding leq interval_split[OF \k \ Basis\] by (auto intro: box_empty_imp) qed have d1_alt: "d1 = (\l. l \ ?lec) ` {l \ d. l \ ?lec \ {}}" using d1_def by auto have d2_alt: "d2 = (\l. l \ ?gec) ` {l \ d. l \ ?gec \ {}}" using d2_def by auto have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev") unfolding * using \k \ Basis\ by (auto dest: Basis_imp) also have "F g d1 = F (\l. g (l \ ?lec)) d" unfolding d1_alt using division_of_finite[OF less.prems] fxk_le by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left) also have "F g d2 = F (\l. g (l \ ?gec)) d" unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left) also have *: "\x\d. g x = g (x \ ?lec) \<^bold>* g (x \ ?gec)" unfolding forall_in_division[OF \d division_of cbox a b\] using \k \ Basis\ by (auto dest: Basis_imp) have "F (\l. g (l \ ?lec)) d \<^bold>* F (\l. g (l \ ?gec)) d = F g d" using * by (simp add: distrib) finally show ?thesis by auto qed qed qed qed proposition tagged_division: assumes "d tagged_division_of (cbox a b)" shows "F (\(_, l). g l) d = g (cbox a b)" proof - have "F (\(_, k). g k) d = F g (snd ` d)" using assms box_empty_imp by (rule over_tagged_division_lemma) then show ?thesis unfolding assms [THEN division_of_tagged_division, THEN division] . qed end locale operative_real = comm_monoid_set + fixes g :: "real set \ 'a" assumes neutral: "b \ a \ g {a..b} = \<^bold>1" assumes coalesce_less: "a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" begin sublocale operative where g = g rewrites "box = (greaterThanLessThan :: real \ _)" and "cbox = (atLeastAtMost :: real \ _)" and "\x::real. x \ Basis \ x = 1" proof - show "operative f z g" proof show "g (cbox a b) = \<^bold>1" if "box a b = {}" for a b using that by (simp add: neutral) show "g (cbox a b) = g (cbox a b \ {x. x \ k \ c}) \<^bold>* g (cbox a b \ {x. c \ x \ k})" if "k \ Basis" for a b c k proof - from that have [simp]: "k = 1" by simp from neutral [of 0 1] neutral [of a a for a] coalesce_less have [simp]: "g {} = \<^bold>1" "\a. g {a} = \<^bold>1" "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" by auto have "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" by (auto simp: min_def max_def le_less) then show "g (cbox a b) = g (cbox a b \ {x. x \ k \ c}) \<^bold>* g (cbox a b \ {x. c \ x \ k})" by (simp add: atMost_def [symmetric] atLeast_def [symmetric]) qed qed show "box = (greaterThanLessThan :: real \ _)" and "cbox = (atLeastAtMost :: real \ _)" and "\x::real. x \ Basis \ x = 1" by (simp_all add: fun_eq_iff) qed lemma coalesce_less_eq: "g {a..c} \<^bold>* g {c..b} = g {a..b}" if "a \ c" "c \ b" proof (cases "c = a \ c = b") case False with that have "a < c" "c < b" by auto then show ?thesis by (rule coalesce_less) next case True with that box_empty_imp [of a a] box_empty_imp [of b b] show ?thesis by safe simp_all qed end lemma operative_realI: "operative_real f z g" if "operative f z g" proof - interpret operative f z g using that . show ?thesis proof show "g {a..b} = z" if "b \ a" for a b using that box_empty_imp by simp show "f (g {a..c}) (g {c..b}) = g {a..b}" if "a < c" "c < b" for a b c using that using Basis_imp [of 1 a b c] by (simp_all add: atMost_def [symmetric] atLeast_def [symmetric] max_def min_def) qed qed subsection \Special case of additivity we need for the FTC\ (*fix add explanation here *) lemma additive_tagged_division_1: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and "p tagged_division_of {a..b}" shows "sum (\(x,k). f(Sup k) - f(Inf k)) p = f b - f a" proof - let ?f = "(\k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" interpret operative_real plus 0 ?f rewrites "comm_monoid_set.F (+) 0 = sum" by standard[1] (auto simp add: sum_def) have p_td: "p tagged_division_of cbox a b" using assms(2) box_real(2) by presburger have **: "cbox a b \ {}" using assms(1) by auto then have "f b - f a = (\(x, l)\p. if l = {} then 0 else f (interval_upperbound l) - f (interval_lowerbound l))" proof - have "(if cbox a b = {} then 0 else f (interval_upperbound (cbox a b)) - f (interval_lowerbound (cbox a b))) = f b - f a" using assms by auto then show ?thesis using p_td assms by (simp add: tagged_division) qed then show ?thesis using assms by (auto intro!: sum.cong) qed subsection \Fine-ness of a partition w.r.t. a gauge\ definition\<^marker>\tag important\ fine (infixr "fine" 46) where "d fine s \ (\(x,k) \ s. k \ d x)" lemma fineI: assumes "\x k. (x, k) \ s \ k \ d x" shows "d fine s" using assms unfolding fine_def by auto lemma fineD[dest]: assumes "d fine s" shows "\x k. (x,k) \ s \ k \ d x" using assms unfolding fine_def by auto lemma fine_Int: "(\x. d1 x \ d2 x) fine p \ d1 fine p \ d2 fine p" unfolding fine_def by auto lemma fine_Inter: "(\x. \{f d x | d. d \ s}) fine p \ (\d\s. (f d) fine p)" unfolding fine_def by blast lemma fine_Un: "d fine p1 \ d fine p2 \ d fine (p1 \ p2)" unfolding fine_def by blast lemma fine_Union: "(\p. p \ ps \ d fine p) \ d fine (\ps)" unfolding fine_def by auto lemma fine_subset: "p \ q \ d fine q \ d fine p" unfolding fine_def by blast subsection \Some basic combining lemmas\ lemma tagged_division_Union_exists: assumes "finite I" and "\i\I. \p. p tagged_division_of i \ d fine p" and "\i1\I. \i2\I. i1 \ i2 \ interior i1 \ interior i2 = {}" and "\I = i" obtains p where "p tagged_division_of i" and "d fine p" proof - obtain pfn where pfn: "\x. x \ I \ pfn x tagged_division_of x" "\x. x \ I \ d fine pfn x" using bchoice[OF assms(2)] by auto show thesis apply (rule_tac p="\(pfn ` I)" in that) using assms(1) assms(3) assms(4) pfn(1) tagged_division_Union apply force by (metis (mono_tags, lifting) fine_Union imageE pfn(2)) qed (* FIXME structure here, do not have one lemma in each subsection *) subsection\<^marker>\tag unimportant\ \The set we're concerned with must be closed\ lemma division_of_closed: fixes i :: "'n::euclidean_space set" shows "s division_of i \ closed i" unfolding division_of_def by fastforce (* FIXME structure here, do not have one lemma in each subsection *) subsection \General bisection principle for intervals; might be useful elsewhere\ (* FIXME move? *) lemma interval_bisection_step: fixes type :: "'a::euclidean_space" assumes emp: "P {}" and Un: "\S T. \P S; P T; interior(S) \ interior(T) = {}\ \ P (S \ T)" and non: "\ P (cbox a (b::'a))" obtains c d where "\ P (cbox c d)" and "\i. i \ Basis \ a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" proof - have "cbox a b \ {}" using emp non by metis then have ab: "\i. i\Basis \ a \ i \ b \ i" by (force simp: mem_box) have UN_cases: "\finite \; \S. S\\ \ P S; \S. S\\ \ \a b. S = cbox a b; \S T. S\\ \ T\\ \ S \ T \ interior S \ interior T = {}\ \ P (\\)" for \ proof (induct \ rule: finite_induct) case empty show ?case using emp by auto next case (insert x f) then show ?case unfolding Union_insert by (metis Int_interior_Union_intervals Un insert_iff open_interior) qed let ?ab = "\i. (a\i + b\i) / 2" let ?A = "{cbox c d | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = ?ab i) \ (c\i = ?ab i) \ (d\i = b\i)}" have "P (\?A)" if "\c d. \i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i \ P (cbox c d)" proof (rule UN_cases) let ?B = "(\S. cbox (\i\Basis. (if i \ S then a\i else ?ab i) *\<^sub>R i::'a) (\i\Basis. (if i \ S then ?ab i else b\i) *\<^sub>R i)) ` {s. s \ Basis}" have "?A \ ?B" proof fix x assume "x \ ?A" then obtain c d where x: "x = cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast have "c = (\i\Basis. (if c \ i = a \ i then a \ i else ?ab i) *\<^sub>R i)" "d = (\i\Basis. (if c \ i = a \ i then ?ab i else b \ i) *\<^sub>R i)" using x(2) ab by (fastforce simp add: euclidean_eq_iff[where 'a='a])+ then show "x \ ?B" unfolding x by (rule_tac x="{i. i\Basis \ c\i = a\i}" in image_eqI) auto qed then show "finite ?A" by (rule finite_subset) auto next fix S assume "S \ ?A" then obtain c d where s: "S = cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast show "P S" unfolding s using ab s(2) by (fastforce intro!: that) show "\a b. S = cbox a b" unfolding s by auto fix T assume "T \ ?A" then obtain e f where t: "T = cbox e f" "\i. i \ Basis \ e \ i = a \ i \ f \ i = ?ab i \ e \ i = ?ab i \ f \ i = b \ i" by blast assume "S \ T" then have "\ (c = e \ d = f)" unfolding s t by auto then obtain i where "c\i \ e\i \ d\i \ f\i" and i': "i \ Basis" unfolding euclidean_eq_iff[where 'a='a] by auto then have i: "c\i \ e\i" "d\i \ f\i" using s(2) t(2) apply fastforce using t(2)[OF i'] \c \ i \ e \ i \ d \ i \ f \ i\ i' s(2) t(2) by fastforce have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" by auto show "interior S \ interior T = {}" unfolding s t interior_cbox proof (rule *) fix x assume "x \ box c d" "x \ box e f" then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" unfolding mem_box using i' by force+ show False using s(2)[OF i'] t(2)[OF i'] and i x by auto qed qed also have "\?A = cbox a b" proof (rule set_eqI,rule) fix x assume "x \ \?A" then obtain c d where x: "x \ cbox c d" "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast then show "x\cbox a b" unfolding mem_box by force next fix x assume x: "x \ cbox a b" then have "\i\Basis. \c d. (c = a\i \ d = ?ab i \ c = ?ab i \ d = b\i) \ c\x\i \ x\i \ d" (is "\i\Basis. \c d. ?P i c d") unfolding mem_box by (metis linear) then obtain \ \ where "\i\Basis. (\ \ i = a \ i \ \ \ i = ?ab i \ \ \ i = ?ab i \ \ \ i = b \ i) \ \ \ i \ x \ i \ x \ i \ \ \ i" by (auto simp: choice_Basis_iff) then show "x\\?A" by (force simp add: mem_box) qed finally show thesis by (metis (no_types, lifting) assms(3) that) qed lemma interval_bisection: fixes type :: "'a::euclidean_space" assumes "P {}" and Un: "\S T. \P S; P T; interior(S) \ interior(T) = {}\ \ P (S \ T)" and "\ P (cbox a (b::'a))" obtains x where "x \ cbox a b" and "\e>0. \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ cbox a b \ \ P (cbox c d)" proof - have "\x. \y. \ P (cbox (fst x) (snd x)) \ (\ P (cbox (fst y) (snd y)) \ (\i\Basis. fst x\i \ fst y\i \ fst y\i \ snd y\i \ snd y\i \ snd x\i \ 2 * (snd y\i - fst y\i) \ snd x\i - fst x\i))" (is "\x. ?P x") proof show "?P x" for x proof (cases "P (cbox (fst x) (snd x))") case True then show ?thesis by auto next case False obtain c d where "\ P (cbox c d)" "\i. i \ Basis \ fst x \ i \ c \ i \ c \ i \ d \ i \ d \ i \ snd x \ i \ 2 * (d \ i - c \ i) \ snd x \ i - fst x \ i" by (blast intro: interval_bisection_step[of P, OF assms(1-2) False]) then show ?thesis by (rule_tac x="(c,d)" in exI) auto qed qed then obtain f where f: "\x. \ P (cbox (fst x) (snd x)) \ \ P (cbox (fst (f x)) (snd (f x))) \ (\i\Basis. fst x \ i \ fst (f x) \ i \ fst (f x) \ i \ snd (f x) \ i \ snd (f x) \ i \ snd x \ i \ 2 * (snd (f x) \ i - fst (f x) \ i) \ snd x \ i - fst x \ i)" by metis define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n have [simp]: "A 0 = a" "B 0 = b" and ABRAW: "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") proof - show "A 0 = a" "B 0 = b" unfolding ab_def by auto note S = ab_def funpow.simps o_def id_apply show "?P n" for n proof (induct n) case 0 then show ?case unfolding S using \\ P (cbox a b)\ f by auto next case (Suc n) show ?case unfolding S apply (rule f[rule_format]) using Suc unfolding S apply auto done qed qed then have AB: "A(n)\i \ A(Suc n)\i" "A(Suc n)\i \ B(Suc n)\i" "B(Suc n)\i \ B(n)\i" "2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i" if "i\Basis" for i n using that by blast+ have notPAB: "\ P (cbox (A(Suc n)) (B(Suc n)))" for n using ABRAW by blast have interv: "\n. \x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" if e: "0 < e" for e proof - obtain n where n: "(\i\Basis. b \ i - a \ i) / e < 2 ^ n" using real_arch_pow[of 2 "(sum (\i. b\i - a\i) Basis) / e"] by auto show ?thesis proof (rule exI [where x=n], clarify) fix x y assume xy: "x\cbox (A n) (B n)" "y\cbox (A n) (B n)" have "dist x y \ sum (\i. \(x - y)\i\) Basis" unfolding dist_norm by(rule norm_le_l1) also have "\ \ sum (\i. B n\i - A n\i) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" show "\(x - y) \ i\ \ B n \ i - A n \ i" using xy[unfolded mem_box,THEN bspec, OF i] by (auto simp: inner_diff_left) qed also have "\ \ sum (\i. b\i - a\i) Basis / 2^n" unfolding sum_divide_distrib proof (rule sum_mono) show "B n \ i - A n \ i \ (b \ i - a \ i) / 2 ^ n" if i: "i \ Basis" for i proof (induct n) case 0 then show ?case unfolding AB by auto next case (Suc n) have "B (Suc n) \ i - A (Suc n) \ i \ (B n \ i - A n \ i) / 2" using AB(3) that AB(4)[of i n] using i by auto also have "\ \ (b \ i - a \ i) / 2 ^ Suc n" using Suc by (auto simp add: field_simps) finally show ?case . qed qed also have "\ < e" using n using e by (auto simp add: field_simps) finally show "dist x y < e" . qed qed { fix n m :: nat assume "m \ n" then have "cbox (A n) (B n) \ cbox (A m) (B m)" proof (induction rule: inc_induct) case (step i) show ?case using AB by (intro order_trans[OF step.IH] subset_box_imp) auto qed simp } note ABsubset = this have "\n. cbox (A n) (B n) \ {}" by (meson AB dual_order.trans interval_not_empty) then obtain x0 where x0: "\n. x0 \ cbox (A n) (B n)" using decreasing_closed_nest [OF closed_cbox] ABsubset interv by blast show thesis proof (rule that[rule_format, of x0]) show "x0\cbox a b" using \A 0 = a\ \B 0 = b\ x0 by blast fix e :: real assume "e > 0" from interv[OF this] obtain n where n: "\x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" .. have "\ P (cbox (A n) (B n))" proof (cases "0 < n") case True then show ?thesis by (metis Suc_pred' notPAB) next case False then show ?thesis using \A 0 = a\ \B 0 = b\ \\ P (cbox a b)\ by blast qed moreover have "cbox (A n) (B n) \ ball x0 e" using n using x0[of n] by auto moreover have "cbox (A n) (B n) \ cbox a b" using ABsubset \A 0 = a\ \B 0 = b\ by blast ultimately show "\c d. x0 \ cbox c d \ cbox c d \ ball x0 e \ cbox c d \ cbox a b \ \ P (cbox c d)" apply (rule_tac x="A n" in exI) apply (rule_tac x="B n" in exI) apply (auto simp: x0) done qed qed subsection \Cousin's lemma\ lemma fine_division_exists: (*FIXME rename(?) *) fixes a b :: "'a::euclidean_space" assumes "gauge g" obtains p where "p tagged_division_of (cbox a b)" "g fine p" proof (cases "\p. p tagged_division_of (cbox a b) \ g fine p") case True then show ?thesis using that by auto next case False assume "\ (\p. p tagged_division_of (cbox a b) \ g fine p)" obtain x where x: "x \ (cbox a b)" "\e. 0 < e \ \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ (cbox a b) \ \ (\p. p tagged_division_of cbox c d \ g fine p)" apply (rule interval_bisection[of "\s. \p. p tagged_division_of s \ g fine p", OF _ _ False]) apply (simp add: fine_def) apply (metis tagged_division_Un fine_Un) apply auto done obtain e where e: "e > 0" "ball x e \ g x" using gaugeD[OF assms, of x] unfolding open_contains_ball by auto from x(2)[OF e(1)] obtain c d where c_d: "x \ cbox c d" "cbox c d \ ball x e" "cbox c d \ cbox a b" "\ (\p. p tagged_division_of cbox c d \ g fine p)" by blast have "g fine {(x, cbox c d)}" unfolding fine_def using e using c_d(2) by auto then show ?thesis using tagged_division_of_self[OF c_d(1)] using c_d by auto qed lemma fine_division_exists_real: fixes a b :: real assumes "gauge g" obtains p where "p tagged_division_of {a .. b}" "g fine p" by (metis assms box_real(2) fine_division_exists) subsection \A technical lemma about "refinement" of division\ lemma tagged_division_finer: fixes p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" assumes ptag: "p tagged_division_of (cbox a b)" and "gauge d" obtains q where "q tagged_division_of (cbox a b)" and "d fine q" and "\(x,k) \ p. k \ d(x) \ (x,k) \ q" proof - have p: "finite p" "p tagged_partial_division_of (cbox a b)" using ptag tagged_division_of_def by blast+ have "(\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" if "finite p" "p tagged_partial_division_of (cbox a b)" "gauge d" for p using that proof (induct p) case empty show ?case by (force simp add: fine_def) next case (insert xk p) obtain x k where xk: "xk = (x, k)" using surj_pair[of xk] by metis obtain q1 where q1: "q1 tagged_division_of \{k. \x. (x, k) \ p}" and "d fine q1" and q1I: "\x k. \(x, k)\p; k \ d x\ \ (x, k) \ q1" using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI] by (metis (mono_tags, lifting) insert.hyps(3) insert.prems(2)) have *: "\{l. \y. (y,l) \ insert xk p} = k \ \{l. \y. (y,l) \ p}" unfolding xk by auto note p = tagged_partial_division_ofD[OF insert(4)] obtain u v where uv: "k = cbox u v" using p(4) xk by blast have "finite {k. \x. (x, k) \ p}" apply (rule finite_subset[of _ "snd ` p"]) using image_iff apply fastforce using insert.hyps(1) by blast then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" proof (rule Int_interior_Union_intervals) show "open (interior (cbox u v))" by auto show "\T. T \ {k. \x. (x, k) \ p} \ \a b. T = cbox a b" using p(4) by auto show "\T. T \ {k. \x. (x, k) \ p} \ interior (cbox u v) \ interior T = {}" by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk) qed show ?case proof (cases "cbox u v \ d x") case True have "{(x, cbox u v)} tagged_division_of cbox u v" by (simp add: p(2) uv xk tagged_division_of_self) then have "{(x, cbox u v)} \ q1 tagged_division_of \{k. \x. (x, k) \ insert xk p}" unfolding * uv by (metis (no_types, lifting) int q1 tagged_division_Un) with True show ?thesis apply (rule_tac x="{(x,cbox u v)} \ q1" in exI) using \d fine q1\ fine_def q1I uv xk apply fastforce done next case False obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2" using fine_division_exists[OF assms(2)] by blast show ?thesis apply (rule_tac x="q2 \ q1" in exI) apply (intro conjI) unfolding * uv apply (rule tagged_division_Un q2 q1 int fine_Un)+ apply (auto intro: q1 q2 fine_Un \d fine q1\ simp add: False q1I uv xk) done qed qed with p obtain q where q: "q tagged_division_of \{k. \x. (x, k) \ p}" "d fine q" "\(x, k)\p. k \ d x \ (x, k) \ q" by (meson \gauge d\) with ptag that show ?thesis by auto qed subsubsection\<^marker>\tag important\ \Covering lemma\ text\ Some technical lemmas used in the approximation results that follow. Proof of the covering lemma is an obvious multidimensional generalization of Lemma 3, p65 of Swartz's "Introduction to Gauge Integrals". \ proposition covering_lemma: assumes "S \ cbox a b" "box a b \ {}" "gauge g" obtains \ where "countable \" "\\ \ cbox a b" "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" "pairwise (\A B. interior A \ interior B = {}) \" "\K. K \ \ \ \x \ S \ K. K \ g x" "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" "S \ \\" proof - have aibi: "\i. i \ Basis \ a \ i \ b \ i" and normab: "0 < norm(b - a)" using \box a b \ {}\ box_eq_empty box_sing by fastforce+ let ?K0 = "\(n, f::'a\nat). cbox (\i \ Basis. (a \ i + (f i / 2^n) * (b \ i - a \ i)) *\<^sub>R i) (\i \ Basis. (a \ i + ((f i + 1) / 2^n) * (b \ i - a \ i)) *\<^sub>R i)" let ?D0 = "?K0 ` (SIGMA n:UNIV. Pi\<^sub>E Basis (\i::'a. lessThan (2^n)))" obtain \0 where count: "countable \0" and sub: "\\0 \ cbox a b" and int: "\K. K \ \0 \ (interior K \ {}) \ (\c d. K = cbox c d)" and intdj: "\A B. \A \ \0; B \ \0\ \ A \ B \ B \ A \ interior A \ interior B = {}" and SK: "\x. x \ S \ \K \ \0. x \ K \ K \ g x" and cbox: "\u v. cbox u v \ \0 \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" and fin: "\K. K \ \0 \ finite {L \ \0. K \ L}" proof show "countable ?D0" by (simp add: countable_PiE) next show "\?D0 \ cbox a b" apply (simp add: UN_subset_iff) apply (intro conjI allI ballI subset_box_imp) apply (simp add: field_simps) apply (auto intro: mult_right_mono aibi) apply (force simp: aibi scaling_mono nat_less_real_le dest: PiE_mem intro: mult_right_mono) done next show "\K. K \ ?D0 \ interior K \ {} \ (\c d. K = cbox c d)" using \box a b \ {}\ by (clarsimp simp: box_eq_empty) (fastforce simp add: field_split_simps dest: PiE_mem) next have realff: "(real w) * 2^m < (real v) * 2^n \ w * 2^m < v * 2^n" for m n v w using of_nat_less_iff less_imp_of_nat_less by fastforce have *: "\v w. ?K0(m,v) \ ?K0(n,w) \ ?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" for m n \ \The symmetry argument requires a single HOL formula\ proof (rule linorder_wlog [where a=m and b=n], intro allI impI) fix v w m and n::nat assume "m \ n" \ \WLOG we can assume \<^term>\m \ n\, when the first disjunct becomes impossible\ have "?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" apply (simp add: subset_box disjoint_interval) apply (rule ccontr) apply (clarsimp simp add: aibi mult_le_cancel_right divide_le_cancel not_less not_le) apply (drule_tac x=i in bspec, assumption) using \m\n\ realff [of _ _ "1+_"] realff [of "1+_"_ "1+_"] apply (auto simp: divide_simps add.commute not_le nat_le_iff_add realff) apply (simp_all add: power_add) apply (metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc) apply (metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc) done then show "?K0(m,v) \ ?K0(n,w) \ ?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" by meson qed auto show "\A B. \A \ ?D0; B \ ?D0\ \ A \ B \ B \ A \ interior A \ interior B = {}" apply (erule imageE SigmaE)+ using * by simp next show "\K \ ?D0. x \ K \ K \ g x" if "x \ S" for x proof (simp only: bex_simps split_paired_Bex_Sigma) show "\n. \f \ Basis \\<^sub>E {..<2 ^ n}. x \ ?K0(n,f) \ ?K0(n,f) \ g x" proof - obtain e where "0 < e" and e: "\y. (\i. i \ Basis \ \x \ i - y \ i\ \ e) \ y \ g x" proof - have "x \ g x" "open (g x)" using \gauge g\ by (auto simp: gauge_def) then obtain \ where "0 < \" and \: "ball x \ \ g x" using openE by blast have "norm (x - y) < \" if "(\i. i \ Basis \ \x \ i - y \ i\ \ \ / (2 * real DIM('a)))" for y proof - have "norm (x - y) \ (\i\Basis. \x \ i - y \ i\)" by (metis (no_types, lifting) inner_diff_left norm_le_l1 sum.cong) also have "... \ DIM('a) * (\ / (2 * real DIM('a)))" by (meson sum_bounded_above that) also have "... = \ / 2" by (simp add: field_split_simps) also have "... < \" by (simp add: \0 < \\) finally show ?thesis . qed then show ?thesis by (rule_tac e = "\ / 2 / DIM('a)" in that) (simp_all add: \0 < \\ dist_norm subsetD [OF \]) qed have xab: "x \ cbox a b" using \x \ S\ \S \ cbox a b\ by blast obtain n where n: "norm (b - a) / 2^n < e" using real_arch_pow_inv [of "e / norm(b - a)" "1/2"] normab \0 < e\ by (auto simp: field_split_simps) then have "norm (b - a) < e * 2^n" by (auto simp: field_split_simps) then have bai: "b \ i - a \ i < e * 2 ^ n" if "i \ Basis" for i proof - have "b \ i - a \ i \ norm (b - a)" by (metis abs_of_nonneg dual_order.trans inner_diff_left linear norm_ge_zero Basis_le_norm that) also have "... < e * 2 ^ n" using \norm (b - a) < e * 2 ^ n\ by blast finally show ?thesis . qed have D: "(a + n \ x \ x \ a + m) \ (a + n \ y \ y \ a + m) \ abs(x - y) \ m - n" for a m n x and y::real by auto have "\i\Basis. \k<2 ^ n. (a \ i + real k * (b \ i - a \ i) / 2 ^ n \ x \ i \ x \ i \ a \ i + (real k + 1) * (b \ i - a \ i) / 2 ^ n)" proof fix i::'a assume "i \ Basis" consider "x \ i = b \ i" | "x \ i < b \ i" using \i \ Basis\ mem_box(2) xab by force then show "\k<2 ^ n. (a \ i + real k * (b \ i - a \ i) / 2 ^ n \ x \ i \ x \ i \ a \ i + (real k + 1) * (b \ i - a \ i) / 2 ^ n)" proof cases case 1 then show ?thesis by (rule_tac x = "2^n - 1" in exI) (auto simp: algebra_simps field_split_simps of_nat_diff \i \ Basis\ aibi) next case 2 then have abi_less: "a \ i < b \ i" using \i \ Basis\ xab by (auto simp: mem_box) let ?k = "nat \2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)\" show ?thesis proof (intro exI conjI) show "?k < 2 ^ n" using aibi xab \i \ Basis\ by (force simp: nat_less_iff floor_less_iff field_split_simps 2 mem_box) next have "a \ i + real ?k * (b \ i - a \ i) / 2 ^ n \ a \ i + (2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)) * (b \ i - a \ i) / 2 ^ n" apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor) using aibi [OF \i \ Basis\] xab 2 apply (simp_all add: \i \ Basis\ mem_box field_split_simps) done also have "... = x \ i" using abi_less by (simp add: field_split_simps) finally show "a \ i + real ?k * (b \ i - a \ i) / 2 ^ n \ x \ i" . next have "x \ i \ a \ i + (2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)) * (b \ i - a \ i) / 2 ^ n" using abi_less by (simp add: field_split_simps) also have "... \ a \ i + (real ?k + 1) * (b \ i - a \ i) / 2 ^ n" apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor) using aibi [OF \i \ Basis\] xab apply (auto simp: \i \ Basis\ mem_box divide_simps) done finally show "x \ i \ a \ i + (real ?k + 1) * (b \ i - a \ i) / 2 ^ n" . qed qed qed then have "\f\Basis \\<^sub>E {..<2 ^ n}. x \ ?K0(n,f)" apply (simp add: mem_box Bex_def) apply (clarify dest!: bchoice) apply (rule_tac x="restrict f Basis" in exI, simp) done moreover have "\f. x \ ?K0(n,f) \ ?K0(n,f) \ g x" apply (clarsimp simp add: mem_box) apply (rule e) apply (drule bspec D, assumption)+ apply (erule order_trans) apply (simp add: divide_simps) using bai apply (force simp add: algebra_simps) done ultimately show ?thesis by auto qed qed next show "\u v. cbox u v \ ?D0 \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" by (force simp: eq_cbox box_eq_empty field_simps dest!: aibi) next obtain j::'a where "j \ Basis" using nonempty_Basis by blast have "finite {L \ ?D0. ?K0(n,f) \ L}" if "f \ Basis \\<^sub>E {..<2 ^ n}" for n f proof (rule finite_subset) let ?B = "(\(n, f::'a\nat). cbox (\i\Basis. (a \ i + (f i) / 2^n * (b \ i - a \ i)) *\<^sub>R i) (\i\Basis. (a \ i + ((f i) + 1) / 2^n * (b \ i - a \ i)) *\<^sub>R i)) ` (SIGMA m:atMost n. Pi\<^sub>E Basis (\i::'a. lessThan (2^m)))" have "?K0(m,g) \ ?B" if "g \ Basis \\<^sub>E {..<2 ^ m}" "?K0(n,f) \ ?K0(m,g)" for m g proof - have dd: "w / m \ v / n \ (v+1) / n \ (w+1) / m \ inverse n \ inverse m" for w m v n::real by (auto simp: field_split_simps) have bjaj: "b \ j - a \ j > 0" using \j \ Basis\ \box a b \ {}\ box_eq_empty(1) by fastforce have "((g j) / 2 ^ m) * (b \ j - a \ j) \ ((f j) / 2 ^ n) * (b \ j - a \ j) \ (((f j) + 1) / 2 ^ n) * (b \ j - a \ j) \ (((g j) + 1) / 2 ^ m) * (b \ j - a \ j)" using that \j \ Basis\ by (simp add: subset_box field_split_simps aibi) then have "((g j) / 2 ^ m) \ ((f j) / 2 ^ n) \ - ((real(f j) + 1) / 2 ^ n) \ ((real(g j) + 1) / 2 ^ m)" - by (metis bjaj mult.commute of_nat_1 of_nat_add real_mult_le_cancel_iff2) + ((real(f j) + 1) / 2 ^ n) \ ((real(g j) + 1) / 2 ^ m)" + by (metis bjaj mult.commute of_nat_1 of_nat_add mult_le_cancel_iff2) then have "inverse (2^n) \ (inverse (2^m) :: real)" by (rule dd) then have "m \ n" by auto show ?thesis by (rule imageI) (simp add: \m \ n\ that) qed then show "{L \ ?D0. ?K0(n,f) \ L} \ ?B" by auto show "finite ?B" by (intro finite_imageI finite_SigmaI finite_atMost finite_lessThan finite_PiE finite_Basis) qed then show "finite {L \ ?D0. K \ L}" if "K \ ?D0" for K using that by auto qed let ?D1 = "{K \ \0. \x \ S \ K. K \ g x}" obtain \ where count: "countable \" and sub: "\\ \ cbox a b" "S \ \\" and int: "\K. K \ \ \ (interior K \ {}) \ (\c d. K = cbox c d)" and intdj: "\A B. \A \ \; B \ \\ \ A \ B \ B \ A \ interior A \ interior B = {}" and SK: "\K. K \ \ \ \x. x \ S \ K \ K \ g x" and cbox: "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" and fin: "\K. K \ \ \ finite {L. L \ \ \ K \ L}" proof show "countable ?D1" using count countable_subset by (simp add: count countable_subset) show "\?D1 \ cbox a b" using sub by blast show "S \ \?D1" using SK by (force simp:) show "\K. K \ ?D1 \ (interior K \ {}) \ (\c d. K = cbox c d)" using int by blast show "\A B. \A \ ?D1; B \ ?D1\ \ A \ B \ B \ A \ interior A \ interior B = {}" using intdj by blast show "\K. K \ ?D1 \ \x. x \ S \ K \ K \ g x" by auto show "\u v. cbox u v \ ?D1 \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" using cbox by blast show "\K. K \ ?D1 \ finite {L. L \ ?D1 \ K \ L}" using fin by simp (metis (mono_tags, lifting) Collect_mono rev_finite_subset) qed let ?\ = "{K \ \. \K'. K' \ \ \ K \ K' \ \(K \ K')}" show ?thesis proof (rule that) show "countable ?\" by (blast intro: countable_subset [OF _ count]) show "\?\ \ cbox a b" using sub by blast show "S \ \?\" proof clarsimp fix x assume "x \ S" then obtain X where "x \ X" "X \ \" using \S \ \\\ by blast let ?R = "{(K,L). K \ \ \ L \ \ \ L \ K}" have irrR: "irrefl ?R" by (force simp: irrefl_def) have traR: "trans ?R" by (force simp: trans_def) have finR: "\x. finite {y. (y, x) \ ?R}" by simp (metis (mono_tags, lifting) fin \X \ \\ finite_subset mem_Collect_eq psubset_imp_subset subsetI) have "{X \ \. x \ X} \ {}" using \X \ \\ \x \ X\ by blast then obtain Y where "Y \ {X \ \. x \ X}" "\Y'. (Y', Y) \ ?R \ Y' \ {X \ \. x \ X}" by (rule wfE_min' [OF wf_finite_segments [OF irrR traR finR]]) blast then show "\Y. Y \ \ \ (\K'. K' \ \ \ Y \ K' \ \ Y \ K') \ x \ Y" by blast qed show "\K. K \ ?\ \ interior K \ {} \ (\c d. K = cbox c d)" by (blast intro: dest: int) show "pairwise (\A B. interior A \ interior B = {}) ?\" using intdj by (simp add: pairwise_def) metis show "\K. K \ ?\ \ \x \ S \ K. K \ g x" using SK by force show "\u v. cbox u v \ ?\ \ \n. \i\Basis. v \ i - u \ i = (b \ i - a \ i) / 2^n" using cbox by force qed qed subsection \Division filter\ text \Divisions over all gauges towards finer divisions.\ definition\<^marker>\tag important\ division_filter :: "'a::euclidean_space set \ ('a \ 'a set) set filter" where "division_filter s = (INF g\{g. gauge g}. principal {p. p tagged_division_of s \ g fine p})" proposition eventually_division_filter: "(\\<^sub>F p in division_filter s. P p) \ (\g. gauge g \ (\p. p tagged_division_of s \ g fine p \ P p))" unfolding division_filter_def proof (subst eventually_INF_base; clarsimp) fix g1 g2 :: "'a \ 'a set" show "gauge g1 \ gauge g2 \ \x. gauge x \ {p. p tagged_division_of s \ x fine p} \ {p. p tagged_division_of s \ g1 fine p} \ {p. p tagged_division_of s \ x fine p} \ {p. p tagged_division_of s \ g2 fine p}" by (intro exI[of _ "\x. g1 x \ g2 x"]) (auto simp: fine_Int) qed (auto simp: eventually_principal) lemma division_filter_not_empty: "division_filter (cbox a b) \ bot" unfolding trivial_limit_def eventually_division_filter by (auto elim: fine_division_exists) lemma eventually_division_filter_tagged_division: "eventually (\p. p tagged_division_of s) (division_filter s)" unfolding eventually_division_filter by (intro exI[of _ "\x. ball x 1"]) auto end diff --git a/src/HOL/Analysis/Vitali_Covering_Theorem.thy b/src/HOL/Analysis/Vitali_Covering_Theorem.thy --- a/src/HOL/Analysis/Vitali_Covering_Theorem.thy +++ b/src/HOL/Analysis/Vitali_Covering_Theorem.thy @@ -1,658 +1,658 @@ (* Title: HOL/Analysis/Vitali_Covering_Theorem.thy Authors: LC Paulson, based on material from HOL Light *) section \Vitali Covering Theorem and an Application to Negligibility\ theory Vitali_Covering_Theorem imports Equivalence_Lebesgue_Henstock_Integration "HOL-Library.Permutations" begin lemma stretch_Galois: fixes x :: "real^'n" shows "(\k. m k \ 0) \ ((y = (\ k. m k * x$k)) \ (\ k. y$k / m k) = x)" by auto lemma lambda_swap_Galois: "(x = (\ i. y $ Fun.swap m n id i) \ (\ i. x $ Fun.swap m n id i) = y)" by (auto; simp add: pointfree_idE vec_eq_iff) lemma lambda_add_Galois: fixes x :: "real^'n" shows "m \ n \ (x = (\ i. if i = m then y$m + y$n else y$i) \ (\ i. if i = m then x$m - x$n else x$i) = y)" by (safe; simp add: vec_eq_iff) lemma Vitali_covering_lemma_cballs_balls: fixes a :: "'a \ 'b::euclidean_space" assumes "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" proof (cases "K = {}") case True with that show ?thesis by auto next case False then have "B > 0" using assms less_le_trans by auto have rgt0[simp]: "\i. i \ K \ 0 < r i" using assms by auto let ?djnt = "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j)))" have "\C. \n. (C n \ K \ (\i \ C n. B/2 ^ n \ r i) \ ?djnt (C n) \ (\i \ K. B/2 ^ n < r i \ (\j. j \ C n \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)))) \ (C n \ C(Suc n))" proof (rule dependent_nat_choice, safe) fix C n define D where "D \ {i \ K. B/2 ^ Suc n < r i \ (\j\C. disjnt (cball(a i)(r i)) (cball (a j) (r j)))}" let ?cover_ar = "\i j. \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" assume "C \ K" and Ble: "\i\C. B/2 ^ n \ r i" and djntC: "?djnt C" and cov_n: "\i\K. B/2 ^ n < r i \ (\j. j \ C \ ?cover_ar i j)" have *: "\C\chains {C. C \ D \ ?djnt C}. \C \ {C. C \ D \ ?djnt C}" proof (clarsimp simp: chains_def) fix C assume C: "C \ {C. C \ D \ ?djnt C}" and "chain\<^sub>\ C" show "\C \ D \ ?djnt (\C)" unfolding pairwise_def proof (intro ballI conjI impI) show "\C \ D" using C by blast next fix x y assume "x \ \C" and "y \ \C" and "x \ y" then obtain X Y where XY: "x \ X" "X \ C" "y \ Y" "Y \ C" by blast then consider "X \ Y" | "Y \ X" by (meson \chain\<^sub>\ C\ chain_subset_def) then show "disjnt (cball (a x) (r x)) (cball (a y) (r y))" proof cases case 1 with C XY \x \ y\ show ?thesis unfolding pairwise_def by blast next case 2 with C XY \x \ y\ show ?thesis unfolding pairwise_def by blast qed qed qed obtain E where "E \ D" and djntE: "?djnt E" and maximalE: "\X. \X \ D; ?djnt X; E \ X\ \ X = E" using Zorn_Lemma [OF *] by safe blast show "\L. (L \ K \ (\i\L. B/2 ^ Suc n \ r i) \ ?djnt L \ (\i\K. B/2 ^ Suc n < r i \ (\j. j \ L \ ?cover_ar i j))) \ C \ L" proof (intro exI conjI ballI) show "C \ E \ K" using D_def \C \ K\ \E \ D\ by blast show "B/2 ^ Suc n \ r i" if i: "i \ C \ E" for i using i proof assume "i \ C" have "B/2 ^ Suc n \ B/2 ^ n" using \B > 0\ by (simp add: field_split_simps) also have "\ \ r i" using Ble \i \ C\ by blast finally show ?thesis . qed (use D_def \E \ D\ in auto) show "?djnt (C \ E)" using D_def \C \ K\ \E \ D\ djntC djntE unfolding pairwise_def disjnt_def by blast next fix i assume "i \ K" show "B/2 ^ Suc n < r i \ (\j. j \ C \ E \ ?cover_ar i j)" proof (cases "r i \ B/2^n") case False then show ?thesis using cov_n \i \ K\ by auto next case True have "cball (a i) (r i) \ ball (a j) (5 * r j)" if less: "B/2 ^ Suc n < r i" and j: "j \ C \ E" and nondis: "\ disjnt (cball (a i) (r i)) (cball (a j) (r j))" for j proof - obtain x where x: "dist (a i) x \ r i" "dist (a j) x \ r j" using nondis by (force simp: disjnt_def) have "dist (a i) (a j) \ dist (a i) x + dist x (a j)" by (simp add: dist_triangle) also have "\ \ r i + r j" by (metis add_mono_thms_linordered_semiring(1) dist_commute x) finally have aij: "dist (a i) (a j) + r i < 5 * r j" if "r i < 2 * r j" using that by auto show ?thesis using j proof assume "j \ C" have "B/2^n < 2 * r j" using Ble True \j \ C\ less by auto with aij True show "cball (a i) (r i) \ ball (a j) (5 * r j)" by (simp add: cball_subset_ball_iff) next assume "j \ E" then have "B/2 ^ n < 2 * r j" using D_def \E \ D\ by auto with True have "r i < 2 * r j" by auto with aij show "cball (a i) (r i) \ ball (a j) (5 * r j)" by (simp add: cball_subset_ball_iff) qed qed moreover have "\j. j \ C \ E \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j))" if "B/2 ^ Suc n < r i" proof (rule classical) assume NON: "\ ?thesis" show ?thesis proof (cases "i \ D") case True have "insert i E = E" proof (rule maximalE) show "insert i E \ D" by (simp add: True \E \ D\) show "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (insert i E)" using False NON by (auto simp: pairwise_insert djntE disjnt_sym) qed auto then show ?thesis using \i \ K\ assms by fastforce next case False with that show ?thesis by (auto simp: D_def disjnt_def \i \ K\) qed qed ultimately show "B/2 ^ Suc n < r i \ (\j. j \ C \ E \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j))" by blast qed qed auto qed (use assms in force) then obtain F where FK: "\n. F n \ K" and Fle: "\n i. i \ F n \ B/2 ^ n \ r i" and Fdjnt: "\n. ?djnt (F n)" and FF: "\n i. \i \ K; B/2 ^ n < r i\ \ \j. j \ F n \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" and inc: "\n. F n \ F(Suc n)" by (force simp: all_conj_distrib) show thesis proof have *: "countable I" if "I \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) I" for I proof - show ?thesis proof (rule countable_image_inj_on [of "\i. cball(a i)(r i)"]) show "countable ((\i. cball (a i) (r i)) ` I)" proof (rule countable_disjoint_nonempty_interior_subsets) show "disjoint ((\i. cball (a i) (r i)) ` I)" by (auto simp: dest: pairwiseD [OF pw] intro: pairwise_imageI) show "\S. \S \ (\i. cball (a i) (r i)) ` I; interior S = {}\ \ S = {}" using \I \ K\ by (auto simp: not_less [symmetric]) qed next have "\x y. \x \ I; y \ I; a x = a y; r x = r y\ \ x = y" using pw \I \ K\ assms apply (clarsimp simp: pairwise_def disjnt_def) by (metis assms centre_in_cball subsetD empty_iff inf.idem less_eq_real_def) then show "inj_on (\i. cball (a i) (r i)) I" using \I \ K\ by (fastforce simp: inj_on_def cball_eq_cball_iff dest: assms) qed qed show "(Union(range F)) \ K" using FK by blast moreover show "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (Union(range F))" proof (rule pairwise_chain_Union) show "chain\<^sub>\ (range F)" unfolding chain_subset_def by clarify (meson inc lift_Suc_mono_le linear subsetCE) qed (use Fdjnt in blast) ultimately show "countable (Union(range F))" by (blast intro: *) next fix i assume "i \ K" then obtain n where "(1/2) ^ n < r i / B" using \B > 0\ assms real_arch_pow_inv by fastforce then have B2: "B/2 ^ n < r i" using \B > 0\ by (simp add: field_split_simps) have "0 < r i" "r i \ B" by (auto simp: \i \ K\ assms) show "\j. j \ (Union(range F)) \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" using FF [OF \i \ K\ B2] by auto qed qed subsection\Vitali covering theorem\ lemma Vitali_covering_lemma_cballs: fixes a :: "'a \ 'b::euclidean_space" assumes S: "S \ (\i\K. cball (a i) (r i))" and r: "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "S \ (\i\C. cball (a i) (5 * r i))" proof - obtain C where C: "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ show ?thesis proof have "(\i\K. cball (a i) (r i)) \ (\i\C. cball (a i) (5 * r i))" using cov subset_iff by fastforce with S show "S \ (\i\C. cball (a i) (5 * r i))" by blast qed (use C in auto) qed lemma Vitali_covering_lemma_balls: fixes a :: "'a \ 'b::euclidean_space" assumes S: "S \ (\i\K. ball (a i) (r i))" and r: "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" "S \ (\i\C. ball (a i) (5 * r i))" proof - obtain C where C: "countable C" "C \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ show ?thesis proof have "(\i\K. ball (a i) (r i)) \ (\i\C. ball (a i) (5 * r i))" using cov subset_iff by clarsimp (meson less_imp_le mem_ball mem_cball subset_eq) with S show "S \ (\i\C. ball (a i) (5 * r i))" by blast show "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" using pw by (clarsimp simp: pairwise_def) (meson ball_subset_cball disjnt_subset1 disjnt_subset2) qed (use C in auto) qed theorem Vitali_covering_theorem_cballs: fixes a :: "'a \ 'n::euclidean_space" assumes r: "\i. i \ K \ 0 < r i" and S: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ cball (a i) (r i) \ r i < d" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "negligible(S - (\i \ C. cball (a i) (r i)))" proof - let ?\ = "measure lebesgue" have *: "\C. countable C \ C \ K \ pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C \ negligible(S - (\i \ C. cball (a i) (r i)))" if r01: "\i. i \ K \ 0 < r i \ r i \ 1" and Sd: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ cball (a i) (r i) \ r i < d" for K r and a :: "'a \ 'n" proof - obtain C where C: "countable C" "C \ K" and pwC: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [of K r 1 a]) (auto simp: r01) have ar_injective: "\x y. \x \ C; y \ C; a x = a y; r x = r y\ \ x = y" using \C \ K\ pwC cov by (force simp: pairwise_def disjnt_def) show ?thesis proof (intro exI conjI) show "negligible (S - (\i\C. cball (a i) (r i)))" proof (clarsimp simp: negligible_on_intervals [of "S-T" for T]) fix l u show "negligible ((S - (\i\C. cball (a i) (r i))) \ cbox l u)" unfolding negligible_outer_le proof (intro allI impI) fix e::real assume "e > 0" define D where "D \ {i \ C. \ disjnt (ball(a i) (5 * r i)) (cbox l u)}" then have "D \ C" by auto have "countable D" unfolding D_def using \countable C\ by simp have UD: "(\i\D. cball (a i) (r i)) \ lmeasurable" proof (rule fmeasurableI2) show "cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One) \ lmeasurable" by blast have "y \ cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" if "i \ C" and x: "x \ cbox l u" and ai: "dist (a i) y \ r i" "dist (a i) x < 5 * r i" for i x y proof - have d6: "dist y x < 6 * r i" using dist_triangle3 [of y x "a i"] that by linarith show ?thesis proof (clarsimp simp: mem_box algebra_simps) fix j::'n assume j: "j \ Basis" then have xyj: "\x \ j - y \ j\ \ dist y x" by (metis Basis_le_norm dist_commute dist_norm inner_diff_left) have "l \ j \ x \ j" using \j \ Basis\ mem_box \x \ cbox l u\ by blast also have "\ \ y \ j + 6 * r i" using d6 xyj by (auto simp: algebra_simps) also have "\ \ y \ j + 6" using r01 [of i] \C \ K\ \i \ C\ by auto finally have l: "l \ j \ y \ j + 6" . have "y \ j \ x \ j + 6 * r i" using d6 xyj by (auto simp: algebra_simps) also have "\ \ u \ j + 6 * r i" using j x by (auto simp: mem_box) also have "\ \ u \ j + 6" using r01 [of i] \C \ K\ \i \ C\ by auto finally have u: "y \ j \ u \ j + 6" . show "l \ j \ y \ j + 6 \ y \ j \ u \ j + 6" using l u by blast qed qed then show "(\i\D. cball (a i) (r i)) \ cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" by (force simp: D_def disjnt_def) show "(\i\D. cball (a i) (r i)) \ sets lebesgue" using \countable D\ by auto qed obtain D1 where "D1 \ D" "finite D1" and measD1: "?\ (\i\D. cball (a i) (r i)) - e / 5 ^ DIM('n) < ?\ (\i\D1. cball (a i) (r i))" proof (rule measure_countable_Union_approachable [where e = "e / 5 ^ (DIM('n))"]) show "countable ((\i. cball (a i) (r i)) ` D)" using \countable D\ by auto show "\d. d \ (\i. cball (a i) (r i)) ` D \ d \ lmeasurable" by auto show "\D'. \D' \ (\i. cball (a i) (r i)) ` D; finite D'\ \ ?\ (\D') \ ?\ (\i\D. cball (a i) (r i))" by (fastforce simp add: intro!: measure_mono_fmeasurable UD) qed (use \e > 0\ in \auto dest: finite_subset_image\) show "\T. (S - (\i\C. cball (a i) (r i))) \ cbox l u \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "(S - (\i\C. cball (a i) (r i))) \ cbox l u \ (\i\D - D1. ball (a i) (5 * r i))" proof clarify fix x assume x: "x \ cbox l u" "x \ S" "x \ (\i\C. cball (a i) (r i))" have "closed (\i\D1. cball (a i) (r i))" using \finite D1\ by blast moreover have "x \ (\j\D1. cball (a j) (r j))" using x \D1 \ D\ unfolding D_def by blast ultimately obtain q where "q > 0" and q: "ball x q \ - (\i\D1. cball (a i) (r i))" by (metis (no_types, lifting) ComplI open_contains_ball closed_def) obtain i where "i \ K" and xi: "x \ cball (a i) (r i)" and ri: "r i < q/2" using Sd [OF \x \ S\] \q > 0\ half_gt_zero by blast then obtain j where "j \ C" and nondisj: "\ disjnt (cball (a i) (r i)) (cball (a j) (r j))" and sub5j: "cball (a i) (r i) \ ball (a j) (5 * r j)" using cov [OF \i \ K\] by metis show "x \ (\i\D - D1. ball (a i) (5 * r i))" proof show "j \ D - D1" proof show "j \ D" using \j \ C\ sub5j \x \ cbox l u\ xi by (auto simp: D_def disjnt_def) obtain y where yi: "dist (a i) y \ r i" and yj: "dist (a j) y \ r j" using disjnt_def nondisj by fastforce have "dist x y \ r i + r i" by (metis add_mono dist_commute dist_triangle_le mem_cball xi yi) also have "\ < q" using ri by linarith finally have "y \ ball x q" by simp with yj q show "j \ D1" by (auto simp: disjoint_UN_iff) qed show "x \ ball (a j) (5 * r j)" using xi sub5j by blast qed qed have 3: "?\ (\i\D2. ball (a i) (5 * r i)) \ e" if D2: "D2 \ D - D1" and "finite D2" for D2 proof - have rgt0: "0 < r i" if "i \ D2" for i using \C \ K\ D_def \i \ D2\ D2 r01 by (simp add: subset_iff) then have inj: "inj_on (\i. ball (a i) (5 * r i)) D2" using \C \ K\ D2 by (fastforce simp: inj_on_def D_def ball_eq_ball_iff intro: ar_injective) have "?\ (\i\D2. ball (a i) (5 * r i)) \ sum (?\) ((\i. ball (a i) (5 * r i)) ` D2)" using that by (force intro: measure_Union_le) also have "\ = (\i\D2. ?\ (ball (a i) (5 * r i)))" by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) also have "\ = (\i\D2. 5 ^ DIM('n) * ?\ (ball (a i) (r i)))" proof (rule sum.cong [OF refl]) fix i assume "i \ D2" thus "?\ (ball (a i) (5 * r i)) = 5 ^ DIM('n) * ?\ (ball (a i) (r i))" using content_ball_conv_unit_ball[of "5 * r i" "a i"] content_ball_conv_unit_ball[of "r i" "a i"] rgt0[of i] by auto qed also have "\ = (\i\D2. ?\ (ball (a i) (r i))) * 5 ^ DIM('n)" by (simp add: sum_distrib_left mult.commute) finally have "?\ (\i\D2. ball (a i) (5 * r i)) \ (\i\D2. ?\ (ball (a i) (r i))) * 5 ^ DIM('n)" . moreover have "(\i\D2. ?\ (ball (a i) (r i))) \ e / 5 ^ DIM('n)" proof - have D12_dis: "((\x\D1. cball (a x) (r x)) \ (\x\D2. cball (a x) (r x))) \ {}" proof clarify fix w d1 d2 assume "d1 \ D1" "w d1 d2 \ cball (a d1) (r d1)" "d2 \ D2" "w d1 d2 \ cball (a d2) (r d2)" then show "w d1 d2 \ {}" by (metis DiffE disjnt_iff subsetCE D2 \D1 \ D\ \D \ C\ pairwiseD [OF pwC, of d1 d2]) qed have inj: "inj_on (\i. cball (a i) (r i)) D2" using rgt0 D2 \D \ C\ by (force simp: inj_on_def cball_eq_cball_iff intro!: ar_injective) have ds: "disjoint ((\i. cball (a i) (r i)) ` D2)" using D2 \D \ C\ by (auto intro: pairwiseI pairwiseD [OF pwC]) have "(\i\D2. ?\ (ball (a i) (r i))) = (\i\D2. ?\ (cball (a i) (r i)))" by (simp add: content_cball_conv_ball) also have "\ = sum ?\ ((\i. cball (a i) (r i)) ` D2)" by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) also have "\ = ?\ (\i\D2. cball (a i) (r i))" by (auto intro: measure_Union' [symmetric] ds simp add: \finite D2\) finally have "?\ (\i\D1. cball (a i) (r i)) + (\i\D2. ?\ (ball (a i) (r i))) = ?\ (\i\D1. cball (a i) (r i)) + ?\ (\i\D2. cball (a i) (r i))" by simp also have "\ = ?\ (\i \ D1 \ D2. cball (a i) (r i))" using D12_dis by (simp add: measure_Un3 \finite D1\ \finite D2\ fmeasurable.finite_UN) also have "\ \ ?\ (\i\D. cball (a i) (r i))" using D2 \D1 \ D\ by (fastforce intro!: measure_mono_fmeasurable [OF _ _ UD] \finite D1\ \finite D2\) finally have "?\ (\i\D1. cball (a i) (r i)) + (\i\D2. ?\ (ball (a i) (r i))) \ ?\ (\i\D. cball (a i) (r i))" . with measD1 show ?thesis by simp qed ultimately show ?thesis by (simp add: field_split_simps) qed have co: "countable (D - D1)" by (simp add: \countable D\) show "(\i\D - D1. ball (a i) (5 * r i)) \ lmeasurable" using \e > 0\ by (auto simp: fmeasurable_UN_bound [OF co _ 3]) show "?\ (\i\D - D1. ball (a i) (5 * r i)) \ e" using \e > 0\ by (auto simp: measure_UN_bound [OF co _ 3]) qed qed qed qed (use C pwC in auto) qed define K' where "K' \ {i \ K. r i \ 1}" have 1: "\i. i \ K' \ 0 < r i \ r i \ 1" using K'_def r by auto have 2: "\i. i \ K' \ x \ cball (a i) (r i) \ r i < d" if "x \ S \ 0 < d" for x d using that by (auto simp: K'_def dest!: S [where d = "min d 1"]) have "K' \ K" using K'_def by auto then show thesis using * [OF 1 2] that by fastforce qed theorem Vitali_covering_theorem_balls: fixes a :: "'a \ 'b::euclidean_space" assumes S: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ ball (a i) (r i) \ r i < d" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" "negligible(S - (\i \ C. ball (a i) (r i)))" proof - have 1: "\i. i \ {i \ K. 0 < r i} \ x \ cball (a i) (r i) \ r i < d" if xd: "x \ S" "d > 0" for x d by (metis (mono_tags, lifting) assms ball_eq_empty less_eq_real_def mem_Collect_eq mem_ball mem_cball not_le xd(1) xd(2)) obtain C where C: "countable C" "C \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and neg: "negligible(S - (\i \ C. cball (a i) (r i)))" by (rule Vitali_covering_theorem_cballs [of "{i \ K. 0 < r i}" r S a, OF _ 1]) auto show thesis proof show "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" apply (rule pairwise_mono [OF pw]) apply (auto simp: disjnt_def) by (meson disjoint_iff_not_equal less_imp_le mem_cball) have "negligible (\i\C. sphere (a i) (r i))" by (auto intro: negligible_sphere \countable C\) then have "negligible (S - (\i \ C. cball(a i)(r i)) \ (\i \ C. sphere (a i) (r i)))" by (rule negligible_Un [OF neg]) then show "negligible (S - (\i\C. ball (a i) (r i)))" by (rule negligible_subset) force qed (use C in auto) qed lemma negligible_eq_zero_density_alt: "negligible S \ (\x \ S. \e > 0. \d U. 0 < d \ d \ e \ S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d))" (is "_ = (\x \ S. \e > 0. ?Q x e)") proof (intro iffI ballI allI impI) fix x and e :: real assume "negligible S" and "x \ S" and "e > 0" then show "\d U. 0 < d \ d \ e \ S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d)" apply (rule_tac x=e in exI) apply (rule_tac x="S \ ball x e" in exI) apply (auto simp: negligible_imp_measurable negligible_Int negligible_imp_measure0 zero_less_measure_iff intro: mult_pos_pos content_ball_pos) done next assume R [rule_format]: "\x \ S. \e > 0. ?Q x e" let ?\ = "measure lebesgue" have "\U. openin (top_of_set S) U \ z \ U \ negligible U" if "z \ S" for z proof (intro exI conjI) show "openin (top_of_set S) (S \ ball z 1)" by (simp add: openin_open_Int) show "z \ S \ ball z 1" using \z \ S\ by auto show "negligible (S \ ball z 1)" proof (clarsimp simp: negligible_outer_le) fix e :: "real" assume "e > 0" let ?K = "{(x,d). x \ S \ 0 < d \ ball x d \ ball z 1 \ (\U. S \ ball x d \ U \ U \ lmeasurable \ ?\ U < e / ?\ (ball z 1) * ?\ (ball x d))}" obtain C where "countable C" and Csub: "C \ ?K" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible((S \ ball z 1) - (\i \ C. ball (fst i) (snd i)))" proof (rule Vitali_covering_theorem_balls [of "S \ ball z 1" ?K fst snd]) fix x and d :: "real" assume x: "x \ S \ ball z 1" and "d > 0" obtain k where "k > 0" and k: "ball x k \ ball z 1" by (meson Int_iff open_ball openE x) let ?\ = "min (e / ?\ (ball z 1) / 2) (min (d / 2) k)" obtain r U where r: "r > 0" "r \ ?\" and U: "S \ ball x r \ U" "U \ lmeasurable" and mU: "?\ U < ?\ * ?\ (ball x r)" using R [of x ?\] \d > 0\ \e > 0\ \k > 0\ x by (auto simp: content_ball_pos) show "\i. i \ ?K \ x \ ball (fst i) (snd i) \ snd i < d" proof (rule exI [of _ "(x,r)"], simp, intro conjI exI) have "ball x r \ ball x k" using r by (simp add: ball_subset_ball_iff) also have "\ \ ball z 1" using ball_subset_ball_iff k by auto finally show "ball x r \ ball z 1" . have "?\ * ?\ (ball x r) \ e * content (ball x r) / content (ball z 1)" using r \e > 0\ by (simp add: ord_class.min_def field_split_simps content_ball_pos) with mU show "?\ U < e * content (ball x r) / content (ball z 1)" by auto qed (use r U x in auto) qed have "\U. case p of (x,d) \ S \ ball x d \ U \ U \ lmeasurable \ ?\ U < e / ?\ (ball z 1) * ?\ (ball x d)" if "p \ C" for p using that Csub unfolding case_prod_unfold by blast then obtain U where U: "\p. p \ C \ case p of (x,d) \ S \ ball x d \ U p \ U p \ lmeasurable \ ?\ (U p) < e / ?\ (ball z 1) * ?\ (ball x d)" by (rule that [OF someI_ex]) let ?T = "((S \ ball z 1) - (\(x,d)\C. ball x d)) \ \(U ` C)" show "\T. S \ ball z 1 \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "S \ ball z 1 \ ?T" using U by fastforce { have Um: "U i \ lmeasurable" if "i \ C" for i using that U by blast have lee: "?\ (\i\I. U i) \ e" if "I \ C" "finite I" for I proof - have "?\ (\(x,d)\I. ball x d) \ ?\ (ball z 1)" apply (rule measure_mono_fmeasurable) using \I \ C\ \finite I\ Csub by (force simp: prod.case_eq_if sets.finite_UN)+ then have le1: "(?\ (\(x,d)\I. ball x d) / ?\ (ball z 1)) \ 1" by (simp add: content_ball_pos) have "?\ (\i\I. U i) \ (\i\I. ?\ (U i))" using that U by (blast intro: measure_UNION_le) also have "\ \ (\(x,r)\I. e / ?\ (ball z 1) * ?\ (ball x r))" by (rule sum_mono) (use \I \ C\ U in force) also have "\ = (e / ?\ (ball z 1)) * (\(x,r)\I. ?\ (ball x r))" by (simp add: case_prod_app prod.case_distrib sum_distrib_left) also have "\ = e * (?\ (\(x,r)\I. ball x r) / ?\ (ball z 1))" apply (subst measure_UNION') using that pwC by (auto simp: case_prod_unfold elim: pairwise_mono) also have "\ \ e" - by (metis mult.commute mult.left_neutral real_mult_le_cancel_iff1 \e > 0\ le1) + by (metis mult.commute mult.left_neutral mult_le_cancel_iff1 \e > 0\ le1) finally show ?thesis . qed have "\(U ` C) \ lmeasurable" "?\ (\(U ` C)) \ e" using \e > 0\ Um lee by(auto intro!: fmeasurable_UN_bound [OF \countable C\] measure_UN_bound [OF \countable C\]) } moreover have "?\ ?T = ?\ (\(U ` C))" proof (rule measure_negligible_symdiff [OF \\(U ` C) \ lmeasurable\]) show "negligible((\(U ` C) - ?T) \ (?T - \(U ` C)))" by (force intro!: negligible_subset [OF negC]) qed ultimately show "?T \ lmeasurable" "?\ ?T \ e" by (simp_all add: fmeasurable.Un negC negligible_imp_measurable split_def) qed qed qed with locally_negligible_alt show "negligible S" by metis qed proposition negligible_eq_zero_density: "negligible S \ (\x\S. \r>0. \e>0. \d. 0 < d \ d \ r \ (\U. S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d)))" proof - let ?Q = "\x d e. \U. S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d)" have "(\e>0. \d>0. d \ e \ ?Q x d e) = (\r>0. \e>0. \d>0. d \ r \ ?Q x d e)" if "x \ S" for x proof (intro iffI allI impI) fix r :: "real" and e :: "real" assume L [rule_format]: "\e>0. \d>0. d \ e \ ?Q x d e" and "r > 0" "e > 0" show "\d>0. d \ r \ ?Q x d e" using L [of "min r e"] apply (rule ex_forward) using \r > 0\ \e > 0\ by (auto intro: less_le_trans elim!: ex_forward simp: content_ball_pos) qed auto then show ?thesis by (force simp: negligible_eq_zero_density_alt) qed end diff --git a/src/HOL/Factorial.thy b/src/HOL/Factorial.thy --- a/src/HOL/Factorial.thy +++ b/src/HOL/Factorial.thy @@ -1,438 +1,451 @@ (* Title: HOL/Factorial.thy Author: Jacques D. Fleuriot Author: Lawrence C Paulson Author: Jeremy Avigad Author: Chaitanya Mangla Author: Manuel Eberl *) section \Factorial Function, Rising Factorials\ theory Factorial imports Groups_List begin subsection \Factorial Function\ context semiring_char_0 begin definition fact :: "nat \ 'a" where fact_prod: "fact n = of_nat (\{1..n})" lemma fact_prod_Suc: "fact n = of_nat (prod Suc {0..i = 0..{1..n}" by (simp add: atLeast0LessThan prod.atLeast1_atMost_eq) then have "prod Suc {0..i. i" 1 n] by presburger then show ?thesis unfolding fact_prod_Suc by (simp add: atLeast0LessThan prod.atLeast1_atMost_eq) qed lemma fact_0 [simp]: "fact 0 = 1" by (simp add: fact_prod) lemma fact_1 [simp]: "fact 1 = 1" by (simp add: fact_prod) lemma fact_Suc_0 [simp]: "fact (Suc 0) = 1" by (simp add: fact_prod) lemma fact_Suc [simp]: "fact (Suc n) = of_nat (Suc n) * fact n" by (simp add: fact_prod atLeastAtMostSuc_conv algebra_simps) lemma fact_2 [simp]: "fact 2 = 2" by (simp add: numeral_2_eq_2) lemma fact_split: "k \ n \ fact n = of_nat (prod Suc {n - k.. 0 \ fact n = of_nat n * fact (n - 1)" by (cases n) auto lemma fact_nonzero [simp]: "fact n \ (0::'a::{semiring_char_0,semiring_no_zero_divisors})" apply (induct n) apply auto using of_nat_eq_0_iff apply fastforce done lemma fact_mono_nat: "m \ n \ fact m \ (fact n :: nat)" by (induct n) (auto simp: le_Suc_eq) lemma fact_in_Nats: "fact n \ \" by (induct n) auto lemma fact_in_Ints: "fact n \ \" by (induct n) auto context assumes "SORT_CONSTRAINT('a::linordered_semidom)" begin lemma fact_mono: "m \ n \ fact m \ (fact n :: 'a)" by (metis of_nat_fact of_nat_le_iff fact_mono_nat) lemma fact_ge_1 [simp]: "fact n \ (1 :: 'a)" by (metis le0 fact_0 fact_mono) lemma fact_gt_zero [simp]: "fact n > (0 :: 'a)" using fact_ge_1 less_le_trans zero_less_one by blast lemma fact_ge_zero [simp]: "fact n \ (0 :: 'a)" by (simp add: less_imp_le) lemma fact_not_neg [simp]: "\ fact n < (0 :: 'a)" by (simp add: not_less_iff_gr_or_eq) lemma fact_le_power: "fact n \ (of_nat (n^n) :: 'a)" proof (induct n) case 0 then show ?case by simp next case (Suc n) then have *: "fact n \ (of_nat (Suc n ^ n) ::'a)" by (rule order_trans) (simp add: power_mono del: of_nat_power) have "fact (Suc n) = (of_nat (Suc n) * fact n ::'a)" by (simp add: algebra_simps) also have "\ \ of_nat (Suc n) * of_nat (Suc n ^ n)" by (simp add: * ordered_comm_semiring_class.comm_mult_left_mono del: of_nat_power) also have "\ \ of_nat (Suc n ^ Suc n)" by (metis of_nat_mult order_refl power_Suc) finally show ?case . qed end lemma fact_less_mono_nat: "0 < m \ m < n \ fact m < (fact n :: nat)" by (induct n) (auto simp: less_Suc_eq) lemma fact_less_mono: "0 < m \ m < n \ fact m < (fact n :: 'a::linordered_semidom)" by (metis of_nat_fact of_nat_less_iff fact_less_mono_nat) lemma fact_ge_Suc_0_nat [simp]: "fact n \ Suc 0" by (metis One_nat_def fact_ge_1) lemma dvd_fact: "1 \ m \ m \ n \ m dvd fact n" by (induct n) (auto simp: dvdI le_Suc_eq) lemma fact_ge_self: "fact n \ n" by (cases "n = 0") (simp_all add: dvd_imp_le dvd_fact) lemma fact_dvd: "n \ m \ fact n dvd (fact m :: 'a::linordered_semidom)" by (induct m) (auto simp: le_Suc_eq) lemma fact_mod: "m \ n \ fact n mod (fact m :: 'a::{semidom_modulo, linordered_semidom}) = 0" by (simp add: mod_eq_0_iff_dvd fact_dvd) lemma fact_div_fact: assumes "m \ n" shows "fact m div fact n = \{n + 1..m}" proof - obtain d where "d = m - n" by auto with assms have "m = n + d" by auto have "fact (n + d) div (fact n) = \{n + 1..n + d}" proof (induct d) case 0 show ?case by simp next case (Suc d') have "fact (n + Suc d') div fact n = Suc (n + d') * fact (n + d') div fact n" by simp also from Suc.hyps have "\ = Suc (n + d') * \{n + 1..n + d'}" unfolding div_mult1_eq[of _ "fact (n + d')"] by (simp add: fact_mod) also have "\ = \{n + 1..n + Suc d'}" by (simp add: atLeastAtMostSuc_conv) finally show ?case . qed with \m = n + d\ show ?thesis by simp qed lemma fact_num_eq_if: "fact m = (if m = 0 then 1 else of_nat m * fact (m - 1))" by (cases m) auto lemma fact_div_fact_le_pow: assumes "r \ n" shows "fact n div fact (n - r) \ n ^ r" proof - have "r \ n \ \{n - r..n} = (n - r) * \{Suc (n - r)..n}" for r by (subst prod.insert[symmetric]) (auto simp: atLeastAtMost_insertL) with assms show ?thesis by (induct r rule: nat.induct) (auto simp add: fact_div_fact Suc_diff_Suc mult_le_mono) qed lemma prod_Suc_fact: "prod Suc {0.. = prod Suc (insert 0 {Suc 0.. \Evaluation for specific numerals\ by (metis fact_Suc numeral_eq_Suc of_nat_numeral) subsection \Pochhammer's symbol: generalized rising factorial\ text \See \<^url>\https://en.wikipedia.org/wiki/Pochhammer_symbol\.\ context comm_semiring_1 begin definition pochhammer :: "'a \ nat \ 'a" where pochhammer_prod: "pochhammer a n = prod (\i. a + of_nat i) {0..i. a + of_nat (n - i)) {1..n}" using prod.atLeastLessThan_rev_at_least_Suc_atMost [of "\i. a + of_nat i" 0 n] by (simp add: pochhammer_prod) lemma pochhammer_Suc_prod: "pochhammer a (Suc n) = prod (\i. a + of_nat i) {0..n}" by (simp add: pochhammer_prod atLeastLessThanSuc_atLeastAtMost) lemma pochhammer_Suc_prod_rev: "pochhammer a (Suc n) = prod (\i. a + of_nat (n - i)) {0..n}" using prod.atLeast_Suc_atMost_Suc_shift by (simp add: pochhammer_prod_rev prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) lemma pochhammer_0 [simp]: "pochhammer a 0 = 1" by (simp add: pochhammer_prod) lemma pochhammer_1 [simp]: "pochhammer a 1 = a" by (simp add: pochhammer_prod lessThan_Suc) lemma pochhammer_Suc0 [simp]: "pochhammer a (Suc 0) = a" by (simp add: pochhammer_prod lessThan_Suc) lemma pochhammer_Suc: "pochhammer a (Suc n) = pochhammer a n * (a + of_nat n)" by (simp add: pochhammer_prod atLeast0_lessThan_Suc ac_simps) end lemma pochhammer_nonneg: fixes x :: "'a :: linordered_semidom" shows "x > 0 \ pochhammer x n \ 0" by (induction n) (auto simp: pochhammer_Suc intro!: mult_nonneg_nonneg add_nonneg_nonneg) lemma pochhammer_pos: fixes x :: "'a :: linordered_semidom" shows "x > 0 \ pochhammer x n > 0" by (induction n) (auto simp: pochhammer_Suc intro!: mult_pos_pos add_pos_nonneg) context comm_semiring_1 begin lemma pochhammer_of_nat: "pochhammer (of_nat x) n = of_nat (pochhammer x n)" by (simp add: pochhammer_prod Factorial.pochhammer_prod) end context comm_ring_1 begin lemma pochhammer_of_int: "pochhammer (of_int x) n = of_int (pochhammer x n)" by (simp add: pochhammer_prod Factorial.pochhammer_prod) end lemma pochhammer_rec: "pochhammer a (Suc n) = a * pochhammer (a + 1) n" by (simp add: pochhammer_prod prod.atLeast0_lessThan_Suc_shift ac_simps del: prod.op_ivl_Suc) lemma pochhammer_rec': "pochhammer z (Suc n) = (z + of_nat n) * pochhammer z n" by (simp add: pochhammer_prod prod.atLeast0_lessThan_Suc ac_simps) lemma pochhammer_fact: "fact n = pochhammer 1 n" by (simp add: pochhammer_prod fact_prod_Suc) lemma pochhammer_of_nat_eq_0_lemma: "k > n \ pochhammer (- (of_nat n :: 'a:: idom)) k = 0" by (auto simp add: pochhammer_prod) lemma pochhammer_of_nat_eq_0_lemma': assumes kn: "k \ n" shows "pochhammer (- (of_nat n :: 'a::{idom,ring_char_0})) k \ 0" proof (cases k) case 0 then show ?thesis by simp next case (Suc h) then show ?thesis apply (simp add: pochhammer_Suc_prod) using Suc kn apply (auto simp add: algebra_simps) done qed lemma pochhammer_of_nat_eq_0_iff: "pochhammer (- (of_nat n :: 'a::{idom,ring_char_0})) k = 0 \ k > n" (is "?l = ?r") using pochhammer_of_nat_eq_0_lemma[of n k, where ?'a='a] pochhammer_of_nat_eq_0_lemma'[of k n, where ?'a = 'a] by (auto simp add: not_le[symmetric]) lemma pochhammer_0_left: "pochhammer 0 n = (if n = 0 then 1 else 0)" by (induction n) (simp_all add: pochhammer_rec) lemma pochhammer_eq_0_iff: "pochhammer a n = (0::'a::field_char_0) \ (\k < n. a = - of_nat k)" by (auto simp add: pochhammer_prod eq_neg_iff_add_eq_0) lemma pochhammer_eq_0_mono: "pochhammer a n = (0::'a::field_char_0) \ m \ n \ pochhammer a m = 0" unfolding pochhammer_eq_0_iff by auto lemma pochhammer_neq_0_mono: "pochhammer a m \ (0::'a::field_char_0) \ m \ n \ pochhammer a n \ 0" unfolding pochhammer_eq_0_iff by auto lemma pochhammer_minus: "pochhammer (- b) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (b - of_nat k + 1) k" proof (cases k) case 0 then show ?thesis by simp next case (Suc h) have eq: "((- 1) ^ Suc h :: 'a) = (\i = 0..h. - 1)" using prod_constant [where A="{0.. h}" and y="- 1 :: 'a"] by auto with Suc show ?thesis using pochhammer_Suc_prod_rev [of "b - of_nat k + 1"] by (auto simp add: pochhammer_Suc_prod prod.distrib [symmetric] eq of_nat_diff simp del: prod_constant) qed lemma pochhammer_minus': "pochhammer (b - of_nat k + 1) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (- b) k" by (simp add: pochhammer_minus) lemma pochhammer_same: "pochhammer (- of_nat n) n = ((- 1) ^ n :: 'a::{semiring_char_0,comm_ring_1,semiring_no_zero_divisors}) * fact n" unfolding pochhammer_minus by (simp add: of_nat_diff pochhammer_fact) lemma pochhammer_product': "pochhammer z (n + m) = pochhammer z n * pochhammer (z + of_nat n) m" proof (induct n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have "pochhammer z (Suc n) * pochhammer (z + of_nat (Suc n)) m = z * (pochhammer (z + 1) n * pochhammer (z + 1 + of_nat n) m)" by (simp add: pochhammer_rec ac_simps) also note Suc[symmetric] also have "z * pochhammer (z + 1) (n + m) = pochhammer z (Suc (n + m))" by (subst pochhammer_rec) simp finally show ?case by simp qed lemma pochhammer_product: "m \ n \ pochhammer z n = pochhammer z m * pochhammer (z + of_nat m) (n - m)" using pochhammer_product'[of z m "n - m"] by simp lemma pochhammer_times_pochhammer_half: fixes z :: "'a::field_char_0" shows "pochhammer z (Suc n) * pochhammer (z + 1/2) (Suc n) = (\k=0..2*n+1. z + of_nat k / 2)" proof (induct n) case 0 then show ?case by (simp add: atLeast0_atMost_Suc) next case (Suc n) define n' where "n' = Suc n" have "pochhammer z (Suc n') * pochhammer (z + 1 / 2) (Suc n') = (pochhammer z n' * pochhammer (z + 1 / 2) n') * ((z + of_nat n') * (z + 1/2 + of_nat n'))" (is "_ = _ * ?A") by (simp_all add: pochhammer_rec' mult_ac) also have "?A = (z + of_nat (Suc (2 * n + 1)) / 2) * (z + of_nat (Suc (Suc (2 * n + 1))) / 2)" (is "_ = ?B") by (simp add: field_simps n'_def) also note Suc[folded n'_def] also have "(\k=0..2 * n + 1. z + of_nat k / 2) * ?B = (\k=0..2 * Suc n + 1. z + of_nat k / 2)" by (simp add: atLeast0_atMost_Suc) finally show ?case by (simp add: n'_def) qed lemma pochhammer_double: fixes z :: "'a::field_char_0" shows "pochhammer (2 * z) (2 * n) = of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "pochhammer (2 * z) (2 * (Suc n)) = pochhammer (2 * z) (2 * n) * (2 * (z + of_nat n)) * (2 * (z + of_nat n) + 1)" by (simp add: pochhammer_rec' ac_simps) also note Suc also have "of_nat (2 ^ (2 * n)) * pochhammer z n * pochhammer (z + 1/2) n * (2 * (z + of_nat n)) * (2 * (z + of_nat n) + 1) = of_nat (2 ^ (2 * (Suc n))) * pochhammer z (Suc n) * pochhammer (z + 1/2) (Suc n)" by (simp add: field_simps pochhammer_rec') finally show ?case . qed lemma fact_double: "fact (2 * n) = (2 ^ (2 * n) * pochhammer (1 / 2) n * fact n :: 'a::field_char_0)" using pochhammer_double[of "1/2::'a" n] by (simp add: pochhammer_fact) lemma pochhammer_absorb_comp: "(r - of_nat k) * pochhammer (- r) k = r * pochhammer (-r + 1) k" (is "?lhs = ?rhs") for r :: "'a::comm_ring_1" proof - have "?lhs = - pochhammer (- r) (Suc k)" by (subst pochhammer_rec') (simp add: algebra_simps) also have "\ = ?rhs" by (subst pochhammer_rec) simp finally show ?thesis . qed subsection \Misc\ lemma fact_code [code]: "fact n = (of_nat (fold_atLeastAtMost_nat ((*)) 2 n 1) :: 'a::semiring_char_0)" proof - have "fact n = (of_nat (\{1..n}) :: 'a)" by (simp add: fact_prod) also have "\{1..n} = \{2..n}" by (intro prod.mono_neutral_right) auto also have "\ = fold_atLeastAtMost_nat ((*)) 2 n 1" by (simp add: prod_atLeastAtMost_code) finally show ?thesis . qed lemma pochhammer_code [code]: "pochhammer a n = (if n = 0 then 1 else fold_atLeastAtMost_nat (\n acc. (a + of_nat n) * acc) 0 (n - 1) 1)" by (cases n) (simp_all add: pochhammer_prod prod_atLeastAtMost_code [symmetric] atLeastLessThanSuc_atLeastAtMost) + +lemma mult_less_iff1: "0 < z \ x * z < y * z \ x < y" + for x y z :: "'a::linordered_idom" + by simp + +lemma mult_le_cancel_iff1: "0 < z \ x * z \ y * z \ x \ y" + for x y z :: "'a::linordered_idom" + by simp + +lemma mult_le_cancel_iff2: "0 < z \ z * x \ z * y \ x \ y" + for x y z :: "'a::linordered_idom" + by simp + end diff --git a/src/HOL/Real.thy b/src/HOL/Real.thy --- a/src/HOL/Real.thy +++ b/src/HOL/Real.thy @@ -1,1784 +1,1766 @@ (* Title: HOL/Real.thy Author: Jacques D. Fleuriot, University of Edinburgh, 1998 Author: Larry Paulson, University of Cambridge Author: Jeremy Avigad, Carnegie Mellon University Author: Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 Construction of Cauchy Reals by Brian Huffman, 2010 *) section \Development of the Reals using Cauchy Sequences\ theory Real imports Rat begin text \ This theory contains a formalization of the real numbers as equivalence classes of Cauchy sequences of rationals. See \<^file>\~~/src/HOL/ex/Dedekind_Real.thy\ for an alternative construction using Dedekind cuts. \ subsection \Preliminary lemmas\ text\Useful in convergence arguments\ lemma inverse_of_nat_le: fixes n::nat shows "\n \ m; n\0\ \ 1 / of_nat m \ (1::'a::linordered_field) / of_nat n" by (simp add: frac_le) lemma add_diff_add: "(a + c) - (b + d) = (a - b) + (c - d)" for a b c d :: "'a::ab_group_add" by simp lemma minus_diff_minus: "- a - - b = - (a - b)" for a b :: "'a::ab_group_add" by simp lemma mult_diff_mult: "(x * y - a * b) = x * (y - b) + (x - a) * b" for x y a b :: "'a::ring" by (simp add: algebra_simps) lemma inverse_diff_inverse: fixes a b :: "'a::division_ring" assumes "a \ 0" and "b \ 0" shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)" using assms by (simp add: algebra_simps) lemma obtain_pos_sum: fixes r :: rat assumes r: "0 < r" obtains s t where "0 < s" and "0 < t" and "r = s + t" proof from r show "0 < r/2" by simp from r show "0 < r/2" by simp show "r = r/2 + r/2" by simp qed subsection \Sequences that converge to zero\ definition vanishes :: "(nat \ rat) \ bool" where "vanishes X \ (\r>0. \k. \n\k. \X n\ < r)" lemma vanishesI: "(\r. 0 < r \ \k. \n\k. \X n\ < r) \ vanishes X" unfolding vanishes_def by simp lemma vanishesD: "vanishes X \ 0 < r \ \k. \n\k. \X n\ < r" unfolding vanishes_def by simp lemma vanishes_const [simp]: "vanishes (\n. c) \ c = 0" proof (cases "c = 0") case True then show ?thesis by (simp add: vanishesI) next case False then show ?thesis unfolding vanishes_def using zero_less_abs_iff by blast qed lemma vanishes_minus: "vanishes X \ vanishes (\n. - X n)" unfolding vanishes_def by simp lemma vanishes_add: assumes X: "vanishes X" and Y: "vanishes Y" shows "vanishes (\n. X n + Y n)" proof (rule vanishesI) fix r :: rat assume "0 < r" then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" by (rule obtain_pos_sum) obtain i where i: "\n\i. \X n\ < s" using vanishesD [OF X s] .. obtain j where j: "\n\j. \Y n\ < t" using vanishesD [OF Y t] .. have "\n\max i j. \X n + Y n\ < r" proof clarsimp fix n assume n: "i \ n" "j \ n" have "\X n + Y n\ \ \X n\ + \Y n\" by (rule abs_triangle_ineq) also have "\ < s + t" by (simp add: add_strict_mono i j n) finally show "\X n + Y n\ < r" by (simp only: r) qed then show "\k. \n\k. \X n + Y n\ < r" .. qed lemma vanishes_diff: assumes "vanishes X" "vanishes Y" shows "vanishes (\n. X n - Y n)" unfolding diff_conv_add_uminus by (intro vanishes_add vanishes_minus assms) lemma vanishes_mult_bounded: assumes X: "\a>0. \n. \X n\ < a" assumes Y: "vanishes (\n. Y n)" shows "vanishes (\n. X n * Y n)" proof (rule vanishesI) fix r :: rat assume r: "0 < r" obtain a where a: "0 < a" "\n. \X n\ < a" using X by blast obtain b where b: "0 < b" "r = a * b" proof show "0 < r / a" using r a by simp show "r = a * (r / a)" using a by simp qed obtain k where k: "\n\k. \Y n\ < b" using vanishesD [OF Y b(1)] .. have "\n\k. \X n * Y n\ < r" by (simp add: b(2) abs_mult mult_strict_mono' a k) then show "\k. \n\k. \X n * Y n\ < r" .. qed subsection \Cauchy sequences\ definition cauchy :: "(nat \ rat) \ bool" where "cauchy X \ (\r>0. \k. \m\k. \n\k. \X m - X n\ < r)" lemma cauchyI: "(\r. 0 < r \ \k. \m\k. \n\k. \X m - X n\ < r) \ cauchy X" unfolding cauchy_def by simp lemma cauchyD: "cauchy X \ 0 < r \ \k. \m\k. \n\k. \X m - X n\ < r" unfolding cauchy_def by simp lemma cauchy_const [simp]: "cauchy (\n. x)" unfolding cauchy_def by simp lemma cauchy_add [simp]: assumes X: "cauchy X" and Y: "cauchy Y" shows "cauchy (\n. X n + Y n)" proof (rule cauchyI) fix r :: rat assume "0 < r" then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" by (rule obtain_pos_sum) obtain i where i: "\m\i. \n\i. \X m - X n\ < s" using cauchyD [OF X s] .. obtain j where j: "\m\j. \n\j. \Y m - Y n\ < t" using cauchyD [OF Y t] .. have "\m\max i j. \n\max i j. \(X m + Y m) - (X n + Y n)\ < r" proof clarsimp fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" have "\(X m + Y m) - (X n + Y n)\ \ \X m - X n\ + \Y m - Y n\" unfolding add_diff_add by (rule abs_triangle_ineq) also have "\ < s + t" by (rule add_strict_mono) (simp_all add: i j *) finally show "\(X m + Y m) - (X n + Y n)\ < r" by (simp only: r) qed then show "\k. \m\k. \n\k. \(X m + Y m) - (X n + Y n)\ < r" .. qed lemma cauchy_minus [simp]: assumes X: "cauchy X" shows "cauchy (\n. - X n)" using assms unfolding cauchy_def unfolding minus_diff_minus abs_minus_cancel . lemma cauchy_diff [simp]: assumes "cauchy X" "cauchy Y" shows "cauchy (\n. X n - Y n)" using assms unfolding diff_conv_add_uminus by (simp del: add_uminus_conv_diff) lemma cauchy_imp_bounded: assumes "cauchy X" shows "\b>0. \n. \X n\ < b" proof - obtain k where k: "\m\k. \n\k. \X m - X n\ < 1" using cauchyD [OF assms zero_less_one] .. show "\b>0. \n. \X n\ < b" proof (intro exI conjI allI) have "0 \ \X 0\" by simp also have "\X 0\ \ Max (abs ` X ` {..k})" by simp finally have "0 \ Max (abs ` X ` {..k})" . then show "0 < Max (abs ` X ` {..k}) + 1" by simp next fix n :: nat show "\X n\ < Max (abs ` X ` {..k}) + 1" proof (rule linorder_le_cases) assume "n \ k" then have "\X n\ \ Max (abs ` X ` {..k})" by simp then show "\X n\ < Max (abs ` X ` {..k}) + 1" by simp next assume "k \ n" have "\X n\ = \X k + (X n - X k)\" by simp also have "\X k + (X n - X k)\ \ \X k\ + \X n - X k\" by (rule abs_triangle_ineq) also have "\ < Max (abs ` X ` {..k}) + 1" by (rule add_le_less_mono) (simp_all add: k \k \ n\) finally show "\X n\ < Max (abs ` X ` {..k}) + 1" . qed qed qed lemma cauchy_mult [simp]: assumes X: "cauchy X" and Y: "cauchy Y" shows "cauchy (\n. X n * Y n)" proof (rule cauchyI) fix r :: rat assume "0 < r" then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v" by (rule obtain_pos_sum) obtain a where a: "0 < a" "\n. \X n\ < a" using cauchy_imp_bounded [OF X] by blast obtain b where b: "0 < b" "\n. \Y n\ < b" using cauchy_imp_bounded [OF Y] by blast obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b" proof show "0 < v/b" using v b(1) by simp show "0 < u/a" using u a(1) by simp show "r = a * (u/a) + (v/b) * b" using a(1) b(1) \r = u + v\ by simp qed obtain i where i: "\m\i. \n\i. \X m - X n\ < s" using cauchyD [OF X s] .. obtain j where j: "\m\j. \n\j. \Y m - Y n\ < t" using cauchyD [OF Y t] .. have "\m\max i j. \n\max i j. \X m * Y m - X n * Y n\ < r" proof clarsimp fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" have "\X m * Y m - X n * Y n\ = \X m * (Y m - Y n) + (X m - X n) * Y n\" unfolding mult_diff_mult .. also have "\ \ \X m * (Y m - Y n)\ + \(X m - X n) * Y n\" by (rule abs_triangle_ineq) also have "\ = \X m\ * \Y m - Y n\ + \X m - X n\ * \Y n\" unfolding abs_mult .. also have "\ < a * t + s * b" by (simp_all add: add_strict_mono mult_strict_mono' a b i j *) finally show "\X m * Y m - X n * Y n\ < r" by (simp only: r) qed then show "\k. \m\k. \n\k. \X m * Y m - X n * Y n\ < r" .. qed lemma cauchy_not_vanishes_cases: assumes X: "cauchy X" assumes nz: "\ vanishes X" shows "\b>0. \k. (\n\k. b < - X n) \ (\n\k. b < X n)" proof - obtain r where "0 < r" and r: "\k. \n\k. r \ \X n\" using nz unfolding vanishes_def by (auto simp add: not_less) obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t" using \0 < r\ by (rule obtain_pos_sum) obtain i where i: "\m\i. \n\i. \X m - X n\ < s" using cauchyD [OF X s] .. obtain k where "i \ k" and "r \ \X k\" using r by blast have k: "\n\k. \X n - X k\ < s" using i \i \ k\ by auto have "X k \ - r \ r \ X k" using \r \ \X k\\ by auto then have "(\n\k. t < - X n) \ (\n\k. t < X n)" unfolding \r = s + t\ using k by auto then have "\k. (\n\k. t < - X n) \ (\n\k. t < X n)" .. then show "\t>0. \k. (\n\k. t < - X n) \ (\n\k. t < X n)" using t by auto qed lemma cauchy_not_vanishes: assumes X: "cauchy X" and nz: "\ vanishes X" shows "\b>0. \k. \n\k. b < \X n\" using cauchy_not_vanishes_cases [OF assms] by (elim ex_forward conj_forward asm_rl) auto lemma cauchy_inverse [simp]: assumes X: "cauchy X" and nz: "\ vanishes X" shows "cauchy (\n. inverse (X n))" proof (rule cauchyI) fix r :: rat assume "0 < r" obtain b i where b: "0 < b" and i: "\n\i. b < \X n\" using cauchy_not_vanishes [OF X nz] by blast from b i have nz: "\n\i. X n \ 0" by auto obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b" proof show "0 < b * r * b" by (simp add: \0 < r\ b) show "r = inverse b * (b * r * b) * inverse b" using b by simp qed obtain j where j: "\m\j. \n\j. \X m - X n\ < s" using cauchyD [OF X s] .. have "\m\max i j. \n\max i j. \inverse (X m) - inverse (X n)\ < r" proof clarsimp fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" have "\inverse (X m) - inverse (X n)\ = inverse \X m\ * \X m - X n\ * inverse \X n\" by (simp add: inverse_diff_inverse nz * abs_mult) also have "\ < inverse b * s * inverse b" by (simp add: mult_strict_mono less_imp_inverse_less i j b * s) finally show "\inverse (X m) - inverse (X n)\ < r" by (simp only: r) qed then show "\k. \m\k. \n\k. \inverse (X m) - inverse (X n)\ < r" .. qed lemma vanishes_diff_inverse: assumes X: "cauchy X" "\ vanishes X" and Y: "cauchy Y" "\ vanishes Y" and XY: "vanishes (\n. X n - Y n)" shows "vanishes (\n. inverse (X n) - inverse (Y n))" proof (rule vanishesI) fix r :: rat assume r: "0 < r" obtain a i where a: "0 < a" and i: "\n\i. a < \X n\" using cauchy_not_vanishes [OF X] by blast obtain b j where b: "0 < b" and j: "\n\j. b < \Y n\" using cauchy_not_vanishes [OF Y] by blast obtain s where s: "0 < s" and "inverse a * s * inverse b = r" proof show "0 < a * r * b" using a r b by simp show "inverse a * (a * r * b) * inverse b = r" using a r b by simp qed obtain k where k: "\n\k. \X n - Y n\ < s" using vanishesD [OF XY s] .. have "\n\max (max i j) k. \inverse (X n) - inverse (Y n)\ < r" proof clarsimp fix n assume n: "i \ n" "j \ n" "k \ n" with i j a b have "X n \ 0" and "Y n \ 0" by auto then have "\inverse (X n) - inverse (Y n)\ = inverse \X n\ * \X n - Y n\ * inverse \Y n\" by (simp add: inverse_diff_inverse abs_mult) also have "\ < inverse a * s * inverse b" by (intro mult_strict_mono' less_imp_inverse_less) (simp_all add: a b i j k n) also note \inverse a * s * inverse b = r\ finally show "\inverse (X n) - inverse (Y n)\ < r" . qed then show "\k. \n\k. \inverse (X n) - inverse (Y n)\ < r" .. qed subsection \Equivalence relation on Cauchy sequences\ definition realrel :: "(nat \ rat) \ (nat \ rat) \ bool" where "realrel = (\X Y. cauchy X \ cauchy Y \ vanishes (\n. X n - Y n))" lemma realrelI [intro?]: "cauchy X \ cauchy Y \ vanishes (\n. X n - Y n) \ realrel X Y" by (simp add: realrel_def) lemma realrel_refl: "cauchy X \ realrel X X" by (simp add: realrel_def) lemma symp_realrel: "symp realrel" by (simp add: abs_minus_commute realrel_def symp_def vanishes_def) lemma transp_realrel: "transp realrel" unfolding realrel_def by (rule transpI) (force simp add: dest: vanishes_add) lemma part_equivp_realrel: "part_equivp realrel" by (blast intro: part_equivpI symp_realrel transp_realrel realrel_refl cauchy_const) subsection \The field of real numbers\ quotient_type real = "nat \ rat" / partial: realrel morphisms rep_real Real by (rule part_equivp_realrel) lemma cr_real_eq: "pcr_real = (\x y. cauchy x \ Real x = y)" unfolding real.pcr_cr_eq cr_real_def realrel_def by auto lemma Real_induct [induct type: real]: (* TODO: generate automatically *) assumes "\X. cauchy X \ P (Real X)" shows "P x" proof (induct x) case (1 X) then have "cauchy X" by (simp add: realrel_def) then show "P (Real X)" by (rule assms) qed lemma eq_Real: "cauchy X \ cauchy Y \ Real X = Real Y \ vanishes (\n. X n - Y n)" using real.rel_eq_transfer unfolding real.pcr_cr_eq cr_real_def rel_fun_def realrel_def by simp lemma Domainp_pcr_real [transfer_domain_rule]: "Domainp pcr_real = cauchy" by (simp add: real.domain_eq realrel_def) instantiation real :: field begin lift_definition zero_real :: "real" is "\n. 0" by (simp add: realrel_refl) lift_definition one_real :: "real" is "\n. 1" by (simp add: realrel_refl) lift_definition plus_real :: "real \ real \ real" is "\X Y n. X n + Y n" unfolding realrel_def add_diff_add by (simp only: cauchy_add vanishes_add simp_thms) lift_definition uminus_real :: "real \ real" is "\X n. - X n" unfolding realrel_def minus_diff_minus by (simp only: cauchy_minus vanishes_minus simp_thms) lift_definition times_real :: "real \ real \ real" is "\X Y n. X n * Y n" proof - fix f1 f2 f3 f4 have "\cauchy f1; cauchy f4; vanishes (\n. f1 n - f2 n); vanishes (\n. f3 n - f4 n)\ \ vanishes (\n. f1 n * (f3 n - f4 n) + f4 n * (f1 n - f2 n))" by (simp add: vanishes_add vanishes_mult_bounded cauchy_imp_bounded) then show "\realrel f1 f2; realrel f3 f4\ \ realrel (\n. f1 n * f3 n) (\n. f2 n * f4 n)" by (simp add: mult.commute realrel_def mult_diff_mult) qed lift_definition inverse_real :: "real \ real" is "\X. if vanishes X then (\n. 0) else (\n. inverse (X n))" proof - fix X Y assume "realrel X Y" then have X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\n. X n - Y n)" by (simp_all add: realrel_def) have "vanishes X \ vanishes Y" proof assume "vanishes X" from vanishes_diff [OF this XY] show "vanishes Y" by simp next assume "vanishes Y" from vanishes_add [OF this XY] show "vanishes X" by simp qed then show "?thesis X Y" by (simp add: vanishes_diff_inverse X Y XY realrel_def) qed definition "x - y = x + - y" for x y :: real definition "x div y = x * inverse y" for x y :: real lemma add_Real: "cauchy X \ cauchy Y \ Real X + Real Y = Real (\n. X n + Y n)" using plus_real.transfer by (simp add: cr_real_eq rel_fun_def) lemma minus_Real: "cauchy X \ - Real X = Real (\n. - X n)" using uminus_real.transfer by (simp add: cr_real_eq rel_fun_def) lemma diff_Real: "cauchy X \ cauchy Y \ Real X - Real Y = Real (\n. X n - Y n)" by (simp add: minus_Real add_Real minus_real_def) lemma mult_Real: "cauchy X \ cauchy Y \ Real X * Real Y = Real (\n. X n * Y n)" using times_real.transfer by (simp add: cr_real_eq rel_fun_def) lemma inverse_Real: "cauchy X \ inverse (Real X) = (if vanishes X then 0 else Real (\n. inverse (X n)))" using inverse_real.transfer zero_real.transfer unfolding cr_real_eq rel_fun_def by (simp split: if_split_asm, metis) instance proof fix a b c :: real show "a + b = b + a" by transfer (simp add: ac_simps realrel_def) show "(a + b) + c = a + (b + c)" by transfer (simp add: ac_simps realrel_def) show "0 + a = a" by transfer (simp add: realrel_def) show "- a + a = 0" by transfer (simp add: realrel_def) show "a - b = a + - b" by (rule minus_real_def) show "(a * b) * c = a * (b * c)" by transfer (simp add: ac_simps realrel_def) show "a * b = b * a" by transfer (simp add: ac_simps realrel_def) show "1 * a = a" by transfer (simp add: ac_simps realrel_def) show "(a + b) * c = a * c + b * c" by transfer (simp add: distrib_right realrel_def) show "(0::real) \ (1::real)" by transfer (simp add: realrel_def) have "vanishes (\n. inverse (X n) * X n - 1)" if X: "cauchy X" "\ vanishes X" for X proof (rule vanishesI) fix r::rat assume "0 < r" obtain b k where "b>0" "\n\k. b < \X n\" using X cauchy_not_vanishes by blast then show "\k. \n\k. \inverse (X n) * X n - 1\ < r" using \0 < r\ by force qed then show "a \ 0 \ inverse a * a = 1" by transfer (simp add: realrel_def) show "a div b = a * inverse b" by (rule divide_real_def) show "inverse (0::real) = 0" by transfer (simp add: realrel_def) qed end subsection \Positive reals\ lift_definition positive :: "real \ bool" is "\X. \r>0. \k. \n\k. r < X n" proof - have 1: "\r>0. \k. \n\k. r < Y n" if *: "realrel X Y" and **: "\r>0. \k. \n\k. r < X n" for X Y proof - from * have XY: "vanishes (\n. X n - Y n)" by (simp_all add: realrel_def) from ** obtain r i where "0 < r" and i: "\n\i. r < X n" by blast obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" using \0 < r\ by (rule obtain_pos_sum) obtain j where j: "\n\j. \X n - Y n\ < s" using vanishesD [OF XY s] .. have "\n\max i j. t < Y n" proof clarsimp fix n assume n: "i \ n" "j \ n" have "\X n - Y n\ < s" and "r < X n" using i j n by simp_all then show "t < Y n" by (simp add: r) qed then show ?thesis using t by blast qed fix X Y assume "realrel X Y" then have "realrel X Y" and "realrel Y X" using symp_realrel by (auto simp: symp_def) then show "?thesis X Y" by (safe elim!: 1) qed lemma positive_Real: "cauchy X \ positive (Real X) \ (\r>0. \k. \n\k. r < X n)" using positive.transfer by (simp add: cr_real_eq rel_fun_def) lemma positive_zero: "\ positive 0" by transfer auto lemma positive_add: assumes "positive x" "positive y" shows "positive (x + y)" proof - have *: "\\n\i. a < x n; \n\j. b < y n; 0 < a; 0 < b; n \ max i j\ \ a+b < x n + y n" for x y and a b::rat and i j n::nat by (simp add: add_strict_mono) show ?thesis using assms by transfer (blast intro: * pos_add_strict) qed lemma positive_mult: assumes "positive x" "positive y" shows "positive (x * y)" proof - have *: "\\n\i. a < x n; \n\j. b < y n; 0 < a; 0 < b; n \ max i j\ \ a*b < x n * y n" for x y and a b::rat and i j n::nat by (simp add: mult_strict_mono') show ?thesis using assms by transfer (blast intro: * mult_pos_pos) qed lemma positive_minus: "\ positive x \ x \ 0 \ positive (- x)" apply transfer apply (simp add: realrel_def) apply (blast dest: cauchy_not_vanishes_cases) done instantiation real :: linordered_field begin definition "x < y \ positive (y - x)" definition "x \ y \ x < y \ x = y" for x y :: real definition "\a\ = (if a < 0 then - a else a)" for a :: real definition "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" for a :: real instance proof fix a b c :: real show "\a\ = (if a < 0 then - a else a)" by (rule abs_real_def) show "a < b \ a \ b \ \ b \ a" "a \ b \ b \ c \ a \ c" "a \ a" "a \ b \ b \ a \ a = b" "a \ b \ c + a \ c + b" unfolding less_eq_real_def less_real_def by (force simp add: positive_zero dest: positive_add)+ show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" by (rule sgn_real_def) show "a \ b \ b \ a" by (auto dest!: positive_minus simp: less_eq_real_def less_real_def) show "a < b \ 0 < c \ c * a < c * b" unfolding less_real_def by (force simp add: algebra_simps dest: positive_mult) qed end instantiation real :: distrib_lattice begin definition "(inf :: real \ real \ real) = min" definition "(sup :: real \ real \ real) = max" instance by standard (auto simp add: inf_real_def sup_real_def max_min_distrib2) end lemma of_nat_Real: "of_nat x = Real (\n. of_nat x)" by (induct x) (simp_all add: zero_real_def one_real_def add_Real) lemma of_int_Real: "of_int x = Real (\n. of_int x)" by (cases x rule: int_diff_cases) (simp add: of_nat_Real diff_Real) lemma of_rat_Real: "of_rat x = Real (\n. x)" proof (induct x) case (Fract a b) then show ?case apply (simp add: Fract_of_int_quotient of_rat_divide) apply (simp add: of_int_Real divide_inverse inverse_Real mult_Real) done qed instance real :: archimedean_field proof show "\z. x \ of_int z" for x :: real proof (induct x) case (1 X) then obtain b where "0 < b" and b: "\n. \X n\ < b" by (blast dest: cauchy_imp_bounded) then have "Real X < of_int (\b\ + 1)" using 1 apply (simp add: of_int_Real less_real_def diff_Real positive_Real) apply (rule_tac x=1 in exI) apply (simp add: algebra_simps) by (metis abs_ge_self le_less_trans le_of_int_ceiling less_le) then show ?case using less_eq_real_def by blast qed qed instantiation real :: floor_ceiling begin definition [code del]: "\x::real\ = (THE z. of_int z \ x \ x < of_int (z + 1))" instance proof show "of_int \x\ \ x \ x < of_int (\x\ + 1)" for x :: real unfolding floor_real_def using floor_exists1 by (rule theI') qed end subsection \Completeness\ lemma not_positive_Real: assumes "cauchy X" shows "\ positive (Real X) \ (\r>0. \k. \n\k. X n \ r)" (is "?lhs = ?rhs") unfolding positive_Real [OF assms] proof (intro iffI allI notI impI) show "\k. \n\k. X n \ r" if r: "\ (\r>0. \k. \n\k. r < X n)" and "0 < r" for r proof - obtain s t where "s > 0" "t > 0" "r = s+t" using \r > 0\ obtain_pos_sum by blast obtain k where k: "\m n. \m\k; n\k\ \ \X m - X n\ < t" using cauchyD [OF assms \t > 0\] by blast obtain n where "n \ k" "X n \ s" by (meson r \0 < s\ not_less) then have "X l \ r" if "l \ n" for l using k [OF \n \ k\, of l] that \r = s+t\ by linarith then show ?thesis by blast qed qed (meson le_cases not_le) lemma le_Real: assumes "cauchy X" "cauchy Y" shows "Real X \ Real Y = (\r>0. \k. \n\k. X n \ Y n + r)" unfolding not_less [symmetric, where 'a=real] less_real_def apply (simp add: diff_Real not_positive_Real assms) apply (simp add: diff_le_eq ac_simps) done lemma le_RealI: assumes Y: "cauchy Y" shows "\n. x \ of_rat (Y n) \ x \ Real Y" proof (induct x) fix X assume X: "cauchy X" and "\n. Real X \ of_rat (Y n)" then have le: "\m r. 0 < r \ \k. \n\k. X n \ Y m + r" by (simp add: of_rat_Real le_Real) then have "\k. \n\k. X n \ Y n + r" if "0 < r" for r :: rat proof - from that obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" by (rule obtain_pos_sum) obtain i where i: "\m\i. \n\i. \Y m - Y n\ < s" using cauchyD [OF Y s] .. obtain j where j: "\n\j. X n \ Y i + t" using le [OF t] .. have "\n\max i j. X n \ Y n + r" proof clarsimp fix n assume n: "i \ n" "j \ n" have "X n \ Y i + t" using n j by simp moreover have "\Y i - Y n\ < s" using n i by simp ultimately show "X n \ Y n + r" unfolding r by simp qed then show ?thesis .. qed then show "Real X \ Real Y" by (simp add: of_rat_Real le_Real X Y) qed lemma Real_leI: assumes X: "cauchy X" assumes le: "\n. of_rat (X n) \ y" shows "Real X \ y" proof - have "- y \ - Real X" by (simp add: minus_Real X le_RealI of_rat_minus le) then show ?thesis by simp qed lemma less_RealD: assumes "cauchy Y" shows "x < Real Y \ \n. x < of_rat (Y n)" apply (erule contrapos_pp) apply (simp add: not_less) apply (erule Real_leI [OF assms]) done lemma of_nat_less_two_power [simp]: "of_nat n < (2::'a::linordered_idom) ^ n" apply (induct n) apply simp apply (metis add_le_less_mono mult_2 of_nat_Suc one_le_numeral one_le_power power_Suc) done lemma complete_real: fixes S :: "real set" assumes "\x. x \ S" and "\z. \x\S. x \ z" shows "\y. (\x\S. x \ y) \ (\z. (\x\S. x \ z) \ y \ z)" proof - obtain x where x: "x \ S" using assms(1) .. obtain z where z: "\x\S. x \ z" using assms(2) .. define P where "P x \ (\y\S. y \ of_rat x)" for x obtain a where a: "\ P a" proof have "of_int \x - 1\ \ x - 1" by (rule of_int_floor_le) also have "x - 1 < x" by simp finally have "of_int \x - 1\ < x" . then have "\ x \ of_int \x - 1\" by (simp only: not_le) then show "\ P (of_int \x - 1\)" unfolding P_def of_rat_of_int_eq using x by blast qed obtain b where b: "P b" proof show "P (of_int \z\)" unfolding P_def of_rat_of_int_eq proof fix y assume "y \ S" then have "y \ z" using z by simp also have "z \ of_int \z\" by (rule le_of_int_ceiling) finally show "y \ of_int \z\" . qed qed define avg where "avg x y = x/2 + y/2" for x y :: rat define bisect where "bisect = (\(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y))" define A where "A n = fst ((bisect ^^ n) (a, b))" for n define B where "B n = snd ((bisect ^^ n) (a, b))" for n define C where "C n = avg (A n) (B n)" for n have A_0 [simp]: "A 0 = a" unfolding A_def by simp have B_0 [simp]: "B 0 = b" unfolding B_def by simp have A_Suc [simp]: "\n. A (Suc n) = (if P (C n) then A n else C n)" unfolding A_def B_def C_def bisect_def split_def by simp have B_Suc [simp]: "\n. B (Suc n) = (if P (C n) then C n else B n)" unfolding A_def B_def C_def bisect_def split_def by simp have width: "B n - A n = (b - a) / 2^n" for n proof (induct n) case (Suc n) then show ?case by (simp add: C_def eq_divide_eq avg_def algebra_simps) qed simp have twos: "\n. y / 2 ^ n < r" if "0 < r" for y r :: rat proof - obtain n where "y / r < rat_of_nat n" using \0 < r\ reals_Archimedean2 by blast then have "\n. y < r * 2 ^ n" by (metis divide_less_eq less_trans mult.commute of_nat_less_two_power that) then show ?thesis by (simp add: field_split_simps) qed have PA: "\ P (A n)" for n by (induct n) (simp_all add: a) have PB: "P (B n)" for n by (induct n) (simp_all add: b) have ab: "a < b" using a b unfolding P_def by (meson leI less_le_trans of_rat_less) have AB: "A n < B n" for n by (induct n) (simp_all add: ab C_def avg_def) have "A i \ A j \ B j \ B i" if "i < j" for i j using that proof (induction rule: less_Suc_induct) case (1 i) then show ?case apply (clarsimp simp add: C_def avg_def add_divide_distrib [symmetric]) apply (rule AB [THEN less_imp_le]) done qed simp then have A_mono: "A i \ A j" and B_mono: "B j \ B i" if "i \ j" for i j by (metis eq_refl le_neq_implies_less that)+ have cauchy_lemma: "cauchy X" if *: "\n i. i\n \ A n \ X i \ X i \ B n" for X proof (rule cauchyI) fix r::rat assume "0 < r" then obtain k where k: "(b - a) / 2 ^ k < r" using twos by blast have "\X m - X n\ < r" if "m\k" "n\k" for m n proof - have "\X m - X n\ \ B k - A k" by (simp add: * abs_rat_def diff_mono that) also have "... < r" by (simp add: k width) finally show ?thesis . qed then show "\k. \m\k. \n\k. \X m - X n\ < r" by blast qed have "cauchy A" by (rule cauchy_lemma) (meson AB A_mono B_mono dual_order.strict_implies_order less_le_trans) have "cauchy B" by (rule cauchy_lemma) (meson AB A_mono B_mono dual_order.strict_implies_order le_less_trans) have "\x\S. x \ Real B" proof fix x assume "x \ S" then show "x \ Real B" using PB [unfolded P_def] \cauchy B\ by (simp add: le_RealI) qed moreover have "\z. (\x\S. x \ z) \ Real A \ z" by (meson PA Real_leI P_def \cauchy A\ le_cases order.trans) moreover have "vanishes (\n. (b - a) / 2 ^ n)" proof (rule vanishesI) fix r :: rat assume "0 < r" then obtain k where k: "\b - a\ / 2 ^ k < r" using twos by blast have "\n\k. \(b - a) / 2 ^ n\ < r" proof clarify fix n assume n: "k \ n" have "\(b - a) / 2 ^ n\ = \b - a\ / 2 ^ n" by simp also have "\ \ \b - a\ / 2 ^ k" using n by (simp add: divide_left_mono) also note k finally show "\(b - a) / 2 ^ n\ < r" . qed then show "\k. \n\k. \(b - a) / 2 ^ n\ < r" .. qed then have "Real B = Real A" by (simp add: eq_Real \cauchy A\ \cauchy B\ width) ultimately show "\y. (\x\S. x \ y) \ (\z. (\x\S. x \ z) \ y \ z)" by force qed instantiation real :: linear_continuum begin subsection \Supremum of a set of reals\ definition "Sup X = (LEAST z::real. \x\X. x \ z)" definition "Inf X = - Sup (uminus ` X)" for X :: "real set" instance proof show Sup_upper: "x \ Sup X" if "x \ X" "bdd_above X" for x :: real and X :: "real set" proof - from that obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" using complete_real[of X] unfolding bdd_above_def by blast then show ?thesis unfolding Sup_real_def by (rule LeastI2_order) (auto simp: that) qed show Sup_least: "Sup X \ z" if "X \ {}" and z: "\x. x \ X \ x \ z" for z :: real and X :: "real set" proof - from that obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" using complete_real [of X] by blast then have "Sup X = s" unfolding Sup_real_def by (best intro: Least_equality) also from s z have "\ \ z" by blast finally show ?thesis . qed show "Inf X \ x" if "x \ X" "bdd_below X" for x :: real and X :: "real set" using Sup_upper [of "-x" "uminus ` X"] by (auto simp: Inf_real_def that) show "z \ Inf X" if "X \ {}" "\x. x \ X \ z \ x" for z :: real and X :: "real set" using Sup_least [of "uminus ` X" "- z"] by (force simp: Inf_real_def that) show "\a b::real. a \ b" using zero_neq_one by blast qed end subsection \Hiding implementation details\ hide_const (open) vanishes cauchy positive Real declare Real_induct [induct del] declare Abs_real_induct [induct del] declare Abs_real_cases [cases del] lifting_update real.lifting lifting_forget real.lifting -subsection \More Lemmas\ - -text \BH: These lemmas should not be necessary; they should be - covered by existing simp rules and simplification procedures.\ - -lemma real_mult_less_iff1 [simp]: "0 < z \ x * z < y * z \ x < y" - for x y z :: real - by simp (* solved by linordered_ring_less_cancel_factor simproc *) - -lemma real_mult_le_cancel_iff1 [simp]: "0 < z \ x * z \ y * z \ x \ y" - for x y z :: real - by simp (* solved by linordered_ring_le_cancel_factor simproc *) - -lemma real_mult_le_cancel_iff2 [simp]: "0 < z \ z * x \ z * y \ x \ y" - for x y z :: real - by simp (* solved by linordered_ring_le_cancel_factor simproc *) - - subsection \Embedding numbers into the Reals\ abbreviation real_of_nat :: "nat \ real" where "real_of_nat \ of_nat" abbreviation real :: "nat \ real" where "real \ of_nat" abbreviation real_of_int :: "int \ real" where "real_of_int \ of_int" abbreviation real_of_rat :: "rat \ real" where "real_of_rat \ of_rat" declare [[coercion_enabled]] declare [[coercion "of_nat :: nat \ int"]] declare [[coercion "of_nat :: nat \ real"]] declare [[coercion "of_int :: int \ real"]] (* We do not add rat to the coerced types, this has often unpleasant side effects when writing inverse (Suc n) which sometimes gets two coercions: of_rat (inverse (of_nat (Suc n))) *) declare [[coercion_map map]] declare [[coercion_map "\f g h x. g (h (f x))"]] declare [[coercion_map "\f g (x,y). (f x, g y)"]] declare of_int_eq_0_iff [algebra, presburger] declare of_int_eq_1_iff [algebra, presburger] declare of_int_eq_iff [algebra, presburger] declare of_int_less_0_iff [algebra, presburger] declare of_int_less_1_iff [algebra, presburger] declare of_int_less_iff [algebra, presburger] declare of_int_le_0_iff [algebra, presburger] declare of_int_le_1_iff [algebra, presburger] declare of_int_le_iff [algebra, presburger] declare of_int_0_less_iff [algebra, presburger] declare of_int_0_le_iff [algebra, presburger] declare of_int_1_less_iff [algebra, presburger] declare of_int_1_le_iff [algebra, presburger] lemma int_less_real_le: "n < m \ real_of_int n + 1 \ real_of_int m" proof - have "(0::real) \ 1" by (metis less_eq_real_def zero_less_one) then show ?thesis by (metis floor_of_int less_floor_iff) qed lemma int_le_real_less: "n \ m \ real_of_int n < real_of_int m + 1" by (meson int_less_real_le not_le) lemma real_of_int_div_aux: "(real_of_int x) / (real_of_int d) = real_of_int (x div d) + (real_of_int (x mod d)) / (real_of_int d)" proof - have "x = (x div d) * d + x mod d" by auto then have "real_of_int x = real_of_int (x div d) * real_of_int d + real_of_int(x mod d)" by (metis of_int_add of_int_mult) then have "real_of_int x / real_of_int d = \ / real_of_int d" by simp then show ?thesis by (auto simp add: add_divide_distrib algebra_simps) qed lemma real_of_int_div: "d dvd n \ real_of_int (n div d) = real_of_int n / real_of_int d" for d n :: int by (simp add: real_of_int_div_aux) lemma real_of_int_div2: "0 \ real_of_int n / real_of_int x - real_of_int (n div x)" proof (cases "x = 0") case False then show ?thesis by (metis diff_ge_0_iff_ge floor_divide_of_int_eq of_int_floor_le) qed simp lemma real_of_int_div3: "real_of_int n / real_of_int x - real_of_int (n div x) \ 1" apply (simp add: algebra_simps) by (metis add.commute floor_correct floor_divide_of_int_eq less_eq_real_def of_int_1 of_int_add) lemma real_of_int_div4: "real_of_int (n div x) \ real_of_int n / real_of_int x" using real_of_int_div2 [of n x] by simp subsection \Embedding the Naturals into the Reals\ lemma real_of_card: "real (card A) = sum (\x. 1) A" by simp lemma nat_less_real_le: "n < m \ real n + 1 \ real m" by (metis discrete of_nat_1 of_nat_add of_nat_le_iff) lemma nat_le_real_less: "n \ m \ real n < real m + 1" for m n :: nat by (meson nat_less_real_le not_le) lemma real_of_nat_div_aux: "real x / real d = real (x div d) + real (x mod d) / real d" proof - have "x = (x div d) * d + x mod d" by auto then have "real x = real (x div d) * real d + real(x mod d)" by (metis of_nat_add of_nat_mult) then have "real x / real d = \ / real d" by simp then show ?thesis by (auto simp add: add_divide_distrib algebra_simps) qed lemma real_of_nat_div: "d dvd n \ real(n div d) = real n / real d" by (subst real_of_nat_div_aux) (auto simp add: dvd_eq_mod_eq_0 [symmetric]) lemma real_of_nat_div2: "0 \ real n / real x - real (n div x)" for n x :: nat apply (simp add: algebra_simps) by (metis floor_divide_of_nat_eq of_int_floor_le of_int_of_nat_eq) lemma real_of_nat_div3: "real n / real x - real (n div x) \ 1" for n x :: nat proof (cases "x = 0") case False then show ?thesis by (metis of_int_of_nat_eq real_of_int_div3 zdiv_int) qed auto lemma real_of_nat_div4: "real (n div x) \ real n / real x" for n x :: nat using real_of_nat_div2 [of n x] by simp subsection \The Archimedean Property of the Reals\ lemma real_arch_inverse: "0 < e \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" using reals_Archimedean[of e] less_trans[of 0 "1 / real n" e for n::nat] by (auto simp add: field_simps cong: conj_cong simp del: of_nat_Suc) lemma reals_Archimedean3: "0 < x \ \y. \n. y < real n * x" by (auto intro: ex_less_of_nat_mult) lemma real_archimedian_rdiv_eq_0: assumes x0: "x \ 0" and c: "c \ 0" and xc: "\m::nat. m > 0 \ real m * x \ c" shows "x = 0" by (metis reals_Archimedean3 dual_order.order_iff_strict le0 le_less_trans not_le x0 xc) subsection \Rationals\ lemma Rats_abs_iff[simp]: "\(x::real)\ \ \ \ x \ \" by(simp add: abs_real_def split: if_splits) lemma Rats_eq_int_div_int: "\ = {real_of_int i / real_of_int j | i j. j \ 0}" (is "_ = ?S") proof show "\ \ ?S" proof fix x :: real assume "x \ \" then obtain r where "x = of_rat r" unfolding Rats_def .. have "of_rat r \ ?S" by (cases r) (auto simp add: of_rat_rat) then show "x \ ?S" using \x = of_rat r\ by simp qed next show "?S \ \" proof (auto simp: Rats_def) fix i j :: int assume "j \ 0" then have "real_of_int i / real_of_int j = of_rat (Fract i j)" by (simp add: of_rat_rat) then show "real_of_int i / real_of_int j \ range of_rat" by blast qed qed lemma Rats_eq_int_div_nat: "\ = { real_of_int i / real n | i n. n \ 0}" proof (auto simp: Rats_eq_int_div_int) fix i j :: int assume "j \ 0" show "\(i'::int) (n::nat). real_of_int i / real_of_int j = real_of_int i' / real n \ 0 < n" proof (cases "j > 0") case True then have "real_of_int i / real_of_int j = real_of_int i / real (nat j) \ 0 < nat j" by simp then show ?thesis by blast next case False with \j \ 0\ have "real_of_int i / real_of_int j = real_of_int (- i) / real (nat (- j)) \ 0 < nat (- j)" by simp then show ?thesis by blast qed next fix i :: int and n :: nat assume "0 < n" then have "real_of_int i / real n = real_of_int i / real_of_int(int n) \ int n \ 0" by simp then show "\i' j. real_of_int i / real n = real_of_int i' / real_of_int j \ j \ 0" by blast qed lemma Rats_abs_nat_div_natE: assumes "x \ \" obtains m n :: nat where "n \ 0" and "\x\ = real m / real n" and "coprime m n" proof - from \x \ \\ obtain i :: int and n :: nat where "n \ 0" and "x = real_of_int i / real n" by (auto simp add: Rats_eq_int_div_nat) then have "\x\ = real (nat \i\) / real n" by simp then obtain m :: nat where x_rat: "\x\ = real m / real n" by blast let ?gcd = "gcd m n" from \n \ 0\ have gcd: "?gcd \ 0" by simp let ?k = "m div ?gcd" let ?l = "n div ?gcd" let ?gcd' = "gcd ?k ?l" have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m" by (rule dvd_mult_div_cancel) have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n" by (rule dvd_mult_div_cancel) from \n \ 0\ and gcd_l have "?gcd * ?l \ 0" by simp then have "?l \ 0" by (blast dest!: mult_not_zero) moreover have "\x\ = real ?k / real ?l" proof - from gcd have "real ?k / real ?l = real (?gcd * ?k) / real (?gcd * ?l)" by (simp add: real_of_nat_div) also from gcd_k and gcd_l have "\ = real m / real n" by simp also from x_rat have "\ = \x\" .. finally show ?thesis .. qed moreover have "?gcd' = 1" proof - have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)" by (rule gcd_mult_distrib_nat) with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp with gcd show ?thesis by auto qed then have "coprime ?k ?l" by (simp only: coprime_iff_gcd_eq_1) ultimately show ?thesis .. qed subsection \Density of the Rational Reals in the Reals\ text \ This density proof is due to Stefan Richter and was ported by TN. The original source is \<^emph>\Real Analysis\ by H.L. Royden. It employs the Archimedean property of the reals.\ lemma Rats_dense_in_real: fixes x :: real assumes "x < y" shows "\r\\. x < r \ r < y" proof - from \x < y\ have "0 < y - x" by simp with reals_Archimedean obtain q :: nat where q: "inverse (real q) < y - x" and "0 < q" by blast define p where "p = \y * real q\ - 1" define r where "r = of_int p / real q" from q have "x < y - inverse (real q)" by simp also from \0 < q\ have "y - inverse (real q) \ r" by (simp add: r_def p_def le_divide_eq left_diff_distrib) finally have "x < r" . moreover from \0 < q\ have "r < y" by (simp add: r_def p_def divide_less_eq diff_less_eq less_ceiling_iff [symmetric]) moreover have "r \ \" by (simp add: r_def) ultimately show ?thesis by blast qed lemma of_rat_dense: fixes x y :: real assumes "x < y" shows "\q :: rat. x < of_rat q \ of_rat q < y" using Rats_dense_in_real [OF \x < y\] by (auto elim: Rats_cases) subsection \Numerals and Arithmetic\ declaration \ K (Lin_Arith.add_inj_const (\<^const_name>\of_nat\, \<^typ>\nat \ real\) #> Lin_Arith.add_inj_const (\<^const_name>\of_int\, \<^typ>\int \ real\)) \ subsection \Simprules combining \x + y\ and \0\\ (* FIXME ARE THEY NEEDED? *) lemma real_add_minus_iff [simp]: "x + - a = 0 \ x = a" for x a :: real by arith lemma real_add_less_0_iff: "x + y < 0 \ y < - x" for x y :: real by auto lemma real_0_less_add_iff: "0 < x + y \ - x < y" for x y :: real by auto lemma real_add_le_0_iff: "x + y \ 0 \ y \ - x" for x y :: real by auto lemma real_0_le_add_iff: "0 \ x + y \ - x \ y" for x y :: real by auto subsection \Lemmas about powers\ lemma two_realpow_ge_one: "(1::real) \ 2 ^ n" by simp (* FIXME: declare this [simp] for all types, or not at all *) declare sum_squares_eq_zero_iff [simp] sum_power2_eq_zero_iff [simp] lemma real_minus_mult_self_le [simp]: "- (u * u) \ x * x" for u x :: real by (rule order_trans [where y = 0]) auto lemma realpow_square_minus_le [simp]: "- u\<^sup>2 \ x\<^sup>2" for u x :: real by (auto simp add: power2_eq_square) subsection \Density of the Reals\ lemma field_lbound_gt_zero: "0 < d1 \ 0 < d2 \ \e. 0 < e \ e < d1 \ e < d2" for d1 d2 :: "'a::linordered_field" by (rule exI [where x = "min d1 d2 / 2"]) (simp add: min_def) lemma field_less_half_sum: "x < y \ x < (x + y) / 2" for x y :: "'a::linordered_field" by auto lemma field_sum_of_halves: "x / 2 + x / 2 = x" for x :: "'a::linordered_field" by simp subsection \Archimedean properties and useful consequences\ text\Bernoulli's inequality\ proposition Bernoulli_inequality: fixes x :: real assumes "-1 \ x" shows "1 + n * x \ (1 + x) ^ n" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "1 + Suc n * x \ 1 + (Suc n)*x + n * x^2" by (simp add: algebra_simps) also have "... = (1 + x) * (1 + n*x)" by (auto simp: power2_eq_square algebra_simps) also have "... \ (1 + x) ^ Suc n" using Suc.hyps assms mult_left_mono by fastforce finally show ?case . qed corollary Bernoulli_inequality_even: fixes x :: real assumes "even n" shows "1 + n * x \ (1 + x) ^ n" proof (cases "-1 \ x \ n=0") case True then show ?thesis by (auto simp: Bernoulli_inequality) next case False then have "real n \ 1" by simp with False have "n * x \ -1" by (metis linear minus_zero mult.commute mult.left_neutral mult_left_mono_neg neg_le_iff_le order_trans zero_le_one) then have "1 + n * x \ 0" by auto also have "... \ (1 + x) ^ n" using assms using zero_le_even_power by blast finally show ?thesis . qed corollary real_arch_pow: fixes x :: real assumes x: "1 < x" shows "\n. y < x^n" proof - from x have x0: "x - 1 > 0" by arith from reals_Archimedean3[OF x0, rule_format, of y] obtain n :: nat where n: "y < real n * (x - 1)" by metis from x0 have x00: "x- 1 \ -1" by arith from Bernoulli_inequality[OF x00, of n] n have "y < x^n" by auto then show ?thesis by metis qed corollary real_arch_pow_inv: fixes x y :: real assumes y: "y > 0" and x1: "x < 1" shows "\n. x^n < y" proof (cases "x > 0") case True with x1 have ix: "1 < 1/x" by (simp add: field_simps) from real_arch_pow[OF ix, of "1/y"] obtain n where n: "1/y < (1/x)^n" by blast then show ?thesis using y \x > 0\ by (auto simp add: field_simps) next case False with y x1 show ?thesis by (metis less_le_trans not_less power_one_right) qed lemma forall_pos_mono: "(\d e::real. d < e \ P d \ P e) \ (\n::nat. n \ 0 \ P (inverse (real n))) \ (\e. 0 < e \ P e)" by (metis real_arch_inverse) lemma forall_pos_mono_1: "(\d e::real. d < e \ P d \ P e) \ (\n. P (inverse (real (Suc n)))) \ 0 < e \ P e" apply (rule forall_pos_mono) apply auto apply (metis Suc_pred of_nat_Suc) done subsection \Floor and Ceiling Functions from the Reals to the Integers\ (* FIXME: theorems for negative numerals. Many duplicates, e.g. from Archimedean_Field.thy. *) lemma real_of_nat_less_numeral_iff [simp]: "real n < numeral w \ n < numeral w" for n :: nat by (metis of_nat_less_iff of_nat_numeral) lemma numeral_less_real_of_nat_iff [simp]: "numeral w < real n \ numeral w < n" for n :: nat by (metis of_nat_less_iff of_nat_numeral) lemma numeral_le_real_of_nat_iff [simp]: "numeral n \ real m \ numeral n \ m" for m :: nat by (metis not_le real_of_nat_less_numeral_iff) lemma of_int_floor_cancel [simp]: "of_int \x\ = x \ (\n::int. x = of_int n)" by (metis floor_of_int) lemma floor_eq: "real_of_int n < x \ x < real_of_int n + 1 \ \x\ = n" by linarith lemma floor_eq2: "real_of_int n \ x \ x < real_of_int n + 1 \ \x\ = n" by (fact floor_unique) lemma floor_eq3: "real n < x \ x < real (Suc n) \ nat \x\ = n" by linarith lemma floor_eq4: "real n \ x \ x < real (Suc n) \ nat \x\ = n" by linarith lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \ real_of_int \r\" by linarith lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real_of_int \r\" by linarith lemma real_of_int_floor_add_one_ge [simp]: "r \ real_of_int \r\ + 1" by linarith lemma real_of_int_floor_add_one_gt [simp]: "r < real_of_int \r\ + 1" by linarith lemma floor_divide_real_eq_div: assumes "0 \ b" shows "\a / real_of_int b\ = \a\ div b" proof (cases "b = 0") case True then show ?thesis by simp next case False with assms have b: "b > 0" by simp have "j = i div b" if "real_of_int i \ a" "a < 1 + real_of_int i" "real_of_int j * real_of_int b \ a" "a < real_of_int b + real_of_int j * real_of_int b" for i j :: int proof - from that have "i < b + j * b" by (metis le_less_trans of_int_add of_int_less_iff of_int_mult) moreover have "j * b < 1 + i" proof - have "real_of_int (j * b) < real_of_int i + 1" using \a < 1 + real_of_int i\ \real_of_int j * real_of_int b \ a\ by force then show "j * b < 1 + i" by linarith qed ultimately have "(j - i div b) * b \ i mod b" "i mod b < ((j - i div b) + 1) * b" by (auto simp: field_simps) then have "(j - i div b) * b < 1 * b" "0 * b < ((j - i div b) + 1) * b" using pos_mod_bound [OF b, of i] pos_mod_sign [OF b, of i] by linarith+ then show ?thesis using b unfolding mult_less_cancel_right by auto qed with b show ?thesis by (auto split: floor_split simp: field_simps) qed lemma floor_one_divide_eq_div_numeral [simp]: "\1 / numeral b::real\ = 1 div numeral b" by (metis floor_divide_of_int_eq of_int_1 of_int_numeral) lemma floor_minus_one_divide_eq_div_numeral [simp]: "\- (1 / numeral b)::real\ = - 1 div numeral b" by (metis (mono_tags, hide_lams) div_minus_right minus_divide_right floor_divide_of_int_eq of_int_neg_numeral of_int_1) lemma floor_divide_eq_div_numeral [simp]: "\numeral a / numeral b::real\ = numeral a div numeral b" by (metis floor_divide_of_int_eq of_int_numeral) lemma floor_minus_divide_eq_div_numeral [simp]: "\- (numeral a / numeral b)::real\ = - numeral a div numeral b" by (metis divide_minus_left floor_divide_of_int_eq of_int_neg_numeral of_int_numeral) lemma of_int_ceiling_cancel [simp]: "of_int \x\ = x \ (\n::int. x = of_int n)" using ceiling_of_int by metis lemma ceiling_eq: "of_int n < x \ x \ of_int n + 1 \ \x\ = n + 1" by (simp add: ceiling_unique) lemma of_int_ceiling_diff_one_le [simp]: "of_int \r\ - 1 \ r" by linarith lemma of_int_ceiling_le_add_one [simp]: "of_int \r\ \ r + 1" by linarith lemma ceiling_le: "x \ of_int a \ \x\ \ a" by (simp add: ceiling_le_iff) lemma ceiling_divide_eq_div: "\of_int a / of_int b\ = - (- a div b)" by (metis ceiling_def floor_divide_of_int_eq minus_divide_left of_int_minus) lemma ceiling_divide_eq_div_numeral [simp]: "\numeral a / numeral b :: real\ = - (- numeral a div numeral b)" using ceiling_divide_eq_div[of "numeral a" "numeral b"] by simp lemma ceiling_minus_divide_eq_div_numeral [simp]: "\- (numeral a / numeral b :: real)\ = - (numeral a div numeral b)" using ceiling_divide_eq_div[of "- numeral a" "numeral b"] by simp text \ The following lemmas are remnants of the erstwhile functions natfloor and natceiling. \ lemma nat_floor_neg: "x \ 0 \ nat \x\ = 0" for x :: real by linarith lemma le_nat_floor: "real x \ a \ x \ nat \a\" by linarith lemma le_mult_nat_floor: "nat \a\ * nat \b\ \ nat \a * b\" by (cases "0 \ a \ 0 \ b") (auto simp add: nat_mult_distrib[symmetric] nat_mono le_mult_floor) lemma nat_ceiling_le_eq [simp]: "nat \x\ \ a \ x \ real a" by linarith lemma real_nat_ceiling_ge: "x \ real (nat \x\)" by linarith lemma Rats_no_top_le: "\q \ \. x \ q" for x :: real by (auto intro!: bexI[of _ "of_nat (nat \x\)"]) linarith lemma Rats_no_bot_less: "\q \ \. q < x" for x :: real by (auto intro!: bexI[of _ "of_int (\x\ - 1)"]) linarith subsection \Exponentiation with floor\ lemma floor_power: assumes "x = of_int \x\" shows "\x ^ n\ = \x\ ^ n" proof - have "x ^ n = of_int (\x\ ^ n)" using assms by (induct n arbitrary: x) simp_all then show ?thesis by (metis floor_of_int) qed lemma floor_numeral_power [simp]: "\numeral x ^ n\ = numeral x ^ n" by (metis floor_of_int of_int_numeral of_int_power) lemma ceiling_numeral_power [simp]: "\numeral x ^ n\ = numeral x ^ n" by (metis ceiling_of_int of_int_numeral of_int_power) subsection \Implementation of rational real numbers\ text \Formal constructor\ definition Ratreal :: "rat \ real" where [code_abbrev, simp]: "Ratreal = real_of_rat" code_datatype Ratreal text \Quasi-Numerals\ lemma [code_abbrev]: "real_of_rat (numeral k) = numeral k" "real_of_rat (- numeral k) = - numeral k" "real_of_rat (rat_of_int a) = real_of_int a" by simp_all lemma [code_post]: "real_of_rat 0 = 0" "real_of_rat 1 = 1" "real_of_rat (- 1) = - 1" "real_of_rat (1 / numeral k) = 1 / numeral k" "real_of_rat (numeral k / numeral l) = numeral k / numeral l" "real_of_rat (- (1 / numeral k)) = - (1 / numeral k)" "real_of_rat (- (numeral k / numeral l)) = - (numeral k / numeral l)" by (simp_all add: of_rat_divide of_rat_minus) text \Operations\ lemma zero_real_code [code]: "0 = Ratreal 0" by simp lemma one_real_code [code]: "1 = Ratreal 1" by simp instantiation real :: equal begin definition "HOL.equal x y \ x - y = 0" for x :: real instance by standard (simp add: equal_real_def) lemma real_equal_code [code]: "HOL.equal (Ratreal x) (Ratreal y) \ HOL.equal x y" by (simp add: equal_real_def equal) lemma [code nbe]: "HOL.equal x x \ True" for x :: real by (rule equal_refl) end lemma real_less_eq_code [code]: "Ratreal x \ Ratreal y \ x \ y" by (simp add: of_rat_less_eq) lemma real_less_code [code]: "Ratreal x < Ratreal y \ x < y" by (simp add: of_rat_less) lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)" by (simp add: of_rat_add) lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)" by (simp add: of_rat_mult) lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)" by (simp add: of_rat_minus) lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)" by (simp add: of_rat_diff) lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)" by (simp add: of_rat_inverse) lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)" by (simp add: of_rat_divide) lemma real_floor_code [code]: "\Ratreal x\ = \x\" by (metis Ratreal_def floor_le_iff floor_unique le_floor_iff of_int_floor_le of_rat_of_int_eq real_less_eq_code) text \Quickcheck\ definition (in term_syntax) valterm_ratreal :: "rat \ (unit \ Code_Evaluation.term) \ real \ (unit \ Code_Evaluation.term)" where [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\} k" notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) instantiation real :: random begin definition "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\r. Pair (valterm_ratreal r))" instance .. end no_notation fcomp (infixl "\>" 60) no_notation scomp (infixl "\\" 60) instantiation real :: exhaustive begin definition "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (\r. f (Ratreal r)) d" instance .. end instantiation real :: full_exhaustive begin definition "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (\r. f (valterm_ratreal r)) d" instance .. end instantiation real :: narrowing begin definition "narrowing_real = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing" instance .. end subsection \Setup for Nitpick\ declaration \ Nitpick_HOL.register_frac_type \<^type_name>\real\ [(\<^const_name>\zero_real_inst.zero_real\, \<^const_name>\Nitpick.zero_frac\), (\<^const_name>\one_real_inst.one_real\, \<^const_name>\Nitpick.one_frac\), (\<^const_name>\plus_real_inst.plus_real\, \<^const_name>\Nitpick.plus_frac\), (\<^const_name>\times_real_inst.times_real\, \<^const_name>\Nitpick.times_frac\), (\<^const_name>\uminus_real_inst.uminus_real\, \<^const_name>\Nitpick.uminus_frac\), (\<^const_name>\inverse_real_inst.inverse_real\, \<^const_name>\Nitpick.inverse_frac\), (\<^const_name>\ord_real_inst.less_real\, \<^const_name>\Nitpick.less_frac\), (\<^const_name>\ord_real_inst.less_eq_real\, \<^const_name>\Nitpick.less_eq_frac\)] \ lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real times_real_inst.times_real uminus_real_inst.uminus_real zero_real_inst.zero_real subsection \Setup for SMT\ ML_file \Tools/SMT/smt_real.ML\ ML_file \Tools/SMT/z3_real.ML\ lemma [z3_rule]: "0 + x = x" "x + 0 = x" "0 * x = 0" "1 * x = x" "-x = -1 * x" "x + y = y + x" for x y :: real by auto lemma [smt_arith_multiplication]: fixes A B :: real and p n :: real assumes "A \ B" "0 < n" "p > 0" shows "(A / n) * p \ (B / n) * p" using assms by (auto simp: field_simps) lemma [smt_arith_multiplication]: fixes A B :: real and p n :: real assumes "A < B" "0 < n" "p > 0" shows "(A / n) * p < (B / n) * p" using assms by (auto simp: field_simps) lemma [smt_arith_multiplication]: fixes A B :: real and p n :: int assumes "A \ B" "0 < n" "p > 0" shows "(A / n) * p \ (B / n) * p" using assms by (auto simp: field_simps) lemma [smt_arith_multiplication]: fixes A B :: real and p n :: int assumes "A < B" "0 < n" "p > 0" shows "(A / n) * p < (B / n) * p" using assms by (auto simp: field_simps) lemmas [smt_arith_multiplication] = verit_le_mono_div[THEN mult_left_mono, unfolded int_distrib, of _ _ \nat (floor (_ :: real))\ \nat (floor (_ :: real))\] div_le_mono[THEN mult_left_mono, unfolded int_distrib, of _ _ \nat (floor (_ :: real))\ \nat (floor (_ :: real))\] verit_le_mono_div_int[THEN mult_left_mono, unfolded int_distrib, of _ _ \floor (_ :: real)\ \floor (_ :: real)\] zdiv_mono1[THEN mult_left_mono, unfolded int_distrib, of _ _ \floor (_ :: real)\ \floor (_ :: real)\] arg_cong[of _ _ \\a :: real. a / real (n::nat) * real (p::nat)\ for n p :: nat, THEN sym] arg_cong[of _ _ \\a :: real. a / real_of_int n * real_of_int p\ for n p :: int, THEN sym] arg_cong[of _ _ \\a :: real. a / n * p\ for n p :: real, THEN sym] lemmas [smt_arith_simplify] = floor_one floor_numeral div_by_1 times_divide_eq_right nonzero_mult_div_cancel_left division_ring_divide_zero div_0 divide_minus_left zero_less_divide_iff subsection \Setup for Argo\ ML_file \Tools/Argo/argo_real.ML\ end