author | hoelzl |
Mon, 08 Aug 2016 14:13:14 +0200 | |
changeset 63627 | 6ddb43c6b711 |
parent 63626 | 44ce6b524ff3 |
child 63631 | 2edc8da89edc |
child 63633 | 2accfb71e33b |
--- a/NEWS Fri Aug 05 18:34:57 2016 +0200 +++ b/NEWS Mon Aug 08 14:13:14 2016 +0200 @@ -190,6 +190,13 @@ *** HOL *** +* Renamed session HOL-Multivariate_Analysis to HOL-Analysis. + +* Moved measure theory from HOL-Probability to HOL-Analysis. When importing +HOL-Analysis some theorems need additional name spaces prefixes due to name +clashes. +INCOMPATIBILITY. + * Number_Theory: algebraic foundation for primes: Introduction of predicates "is_prime", "irreducible", a "prime_factorization" function, the "factorial_ring" typeclass with instance proofs for
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Analysis.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,20 @@ +theory Analysis +imports + Regularity + Lebesgue_Integral_Substitution + Embed_Measure + Complete_Measure + Radon_Nikodym + Fashoda_Theorem + Determinants + Homeomorphism + Bounded_Continuous_Function + Weierstrass_Theorems + Polytope + Poly_Roots + Conformal_Mappings + Generalised_Binomial_Theorem + Gamma_Function +begin + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Binary_Product_Measure.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,1110 @@ +(* Title: HOL/Analysis/Binary_Product_Measure.thy + Author: Johannes Hölzl, TU München +*) + +section \<open>Binary product measures\<close> + +theory Binary_Product_Measure +imports Nonnegative_Lebesgue_Integration +begin + +lemma Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})" + by auto + +lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})" + by auto + +subsection "Binary products" + +definition pair_measure (infixr "\<Otimes>\<^sub>M" 80) where + "A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B) + {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} + (\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)" + +lemma pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)" + using sets.space_closed[of A] sets.space_closed[of B] by auto + +lemma space_pair_measure: + "space (A \<Otimes>\<^sub>M B) = space A \<times> space B" + unfolding pair_measure_def using pair_measure_closed[of A B] + by (rule space_measure_of) + +lemma SIGMA_Collect_eq: "(SIGMA x:space M. {y\<in>space N. P x y}) = {x\<in>space (M \<Otimes>\<^sub>M N). P (fst x) (snd x)}" + by (auto simp: space_pair_measure) + +lemma sets_pair_measure: + "sets (A \<Otimes>\<^sub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}" + unfolding pair_measure_def using pair_measure_closed[of A B] + by (rule sets_measure_of) + +lemma sets_pair_measure_cong[measurable_cong, cong]: + "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')" + unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq) + +lemma pair_measureI[intro, simp, measurable]: + "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)" + by (auto simp: sets_pair_measure) + +lemma sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)" + using pair_measureI[of "{x}" M1 "{y}" M2] by simp + +lemma measurable_pair_measureI: + assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2" + assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M" + shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" + unfolding pair_measure_def using 1 2 + by (intro measurable_measure_of) (auto dest: sets.sets_into_space) + +lemma measurable_split_replace[measurable (raw)]: + "(\<lambda>x. f x (fst (g x)) (snd (g x))) \<in> measurable M N \<Longrightarrow> (\<lambda>x. case_prod (f x) (g x)) \<in> measurable M N" + unfolding split_beta' . + +lemma measurable_Pair[measurable (raw)]: + assumes f: "f \<in> measurable M M1" and g: "g \<in> measurable M M2" + shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" +proof (rule measurable_pair_measureI) + show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2" + using f g by (auto simp: measurable_def) + fix A B assume *: "A \<in> sets M1" "B \<in> sets M2" + have "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)" + by auto + also have "\<dots> \<in> sets M" + by (rule sets.Int) (auto intro!: measurable_sets * f g) + finally show "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M \<in> sets M" . +qed + +lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1" + by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times + measurable_def) + +lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2" + by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times + measurable_def) + +lemma measurable_Pair_compose_split[measurable_dest]: + assumes f: "case_prod f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N" + assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2" + shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N" + using measurable_compose[OF measurable_Pair f, OF g h] by simp + +lemma measurable_Pair1_compose[measurable_dest]: + assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" + assumes [measurable]: "h \<in> measurable N M" + shows "(\<lambda>x. f (h x)) \<in> measurable N M1" + using measurable_compose[OF f measurable_fst] by simp + +lemma measurable_Pair2_compose[measurable_dest]: + assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" + assumes [measurable]: "h \<in> measurable N M" + shows "(\<lambda>x. g (h x)) \<in> measurable N M2" + using measurable_compose[OF f measurable_snd] by simp + +lemma measurable_pair: + assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2" + shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)" + using measurable_Pair[OF assms] by simp + +lemma + assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)" + shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N" + and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P" + by simp_all + +lemma + assumes f[measurable]: "f \<in> measurable M N" + shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^sub>M P) N" + and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N" + by simp_all + +lemma sets_pair_in_sets: + assumes "\<And>a b. a \<in> sets A \<Longrightarrow> b \<in> sets B \<Longrightarrow> a \<times> b \<in> sets N" + shows "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets N" + unfolding sets_pair_measure + by (intro sets.sigma_sets_subset') (auto intro!: assms) + +lemma sets_pair_eq_sets_fst_snd: + "sets (A \<Otimes>\<^sub>M B) = sets (Sup {vimage_algebra (space A \<times> space B) fst A, vimage_algebra (space A \<times> space B) snd B})" + (is "?P = sets (Sup {?fst, ?snd})") +proof - + { fix a b assume ab: "a \<in> sets A" "b \<in> sets B" + then have "a \<times> b = (fst -` a \<inter> (space A \<times> space B)) \<inter> (snd -` b \<inter> (space A \<times> space B))" + by (auto dest: sets.sets_into_space) + also have "\<dots> \<in> sets (Sup {?fst, ?snd})" + apply (rule sets.Int) + apply (rule in_sets_Sup) + apply auto [] + apply (rule insertI1) + apply (auto intro: ab in_vimage_algebra) [] + apply (rule in_sets_Sup) + apply auto [] + apply (rule insertI2) + apply (auto intro: ab in_vimage_algebra) + done + finally have "a \<times> b \<in> sets (Sup {?fst, ?snd})" . } + moreover have "sets ?fst \<subseteq> sets (A \<Otimes>\<^sub>M B)" + by (rule sets_image_in_sets) (auto simp: space_pair_measure[symmetric]) + moreover have "sets ?snd \<subseteq> sets (A \<Otimes>\<^sub>M B)" + by (rule sets_image_in_sets) (auto simp: space_pair_measure) + ultimately show ?thesis + apply (intro antisym[of "sets A" for A] sets_Sup_in_sets sets_pair_in_sets) + apply simp + apply simp + apply simp + apply (elim disjE) + apply (simp add: space_pair_measure) + apply (simp add: space_pair_measure) + apply (auto simp add: space_pair_measure) + done +qed + +lemma measurable_pair_iff: + "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2" + by (auto intro: measurable_pair[of f M M1 M2]) + +lemma measurable_split_conv: + "(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B" + by (intro arg_cong2[where f="op \<in>"]) auto + +lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)" + by (auto intro!: measurable_Pair simp: measurable_split_conv) + +lemma measurable_pair_swap: + assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^sub>M M1) M" + using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def) + +lemma measurable_pair_swap_iff: + "f \<in> measurable (M2 \<Otimes>\<^sub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" + by (auto dest: measurable_pair_swap) + +lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)" + by simp + +lemma sets_Pair1[measurable (raw)]: + assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2" +proof - + have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})" + using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure) + also have "\<dots> \<in> sets M2" + using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: if_split_asm) + finally show ?thesis . +qed + +lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)" + by (auto intro!: measurable_Pair) + +lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1" +proof - + have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})" + using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure) + also have "\<dots> \<in> sets M1" + using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: if_split_asm) + finally show ?thesis . +qed + +lemma measurable_Pair2: + assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1" + shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M" + using measurable_comp[OF measurable_Pair1' f, OF x] + by (simp add: comp_def) + +lemma measurable_Pair1: + assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2" + shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M" + using measurable_comp[OF measurable_Pair2' f, OF y] + by (simp add: comp_def) + +lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}" + unfolding Int_stable_def + by safe (auto simp add: times_Int_times) + +lemma (in finite_measure) finite_measure_cut_measurable: + assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)" + shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" + (is "?s Q \<in> _") + using Int_stable_pair_measure_generator pair_measure_closed assms + unfolding sets_pair_measure +proof (induct rule: sigma_sets_induct_disjoint) + case (compl A) + with sets.sets_into_space have "\<And>x. emeasure M (Pair x -` ((space N \<times> space M) - A)) = + (if x \<in> space N then emeasure M (space M) - ?s A x else 0)" + unfolding sets_pair_measure[symmetric] + by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1) + with compl sets.top show ?case + by (auto intro!: measurable_If simp: space_pair_measure) +next + case (union F) + then have "\<And>x. emeasure M (Pair x -` (\<Union>i. F i)) = (\<Sum>i. ?s (F i) x)" + by (simp add: suminf_emeasure disjoint_family_on_vimageI subset_eq vimage_UN sets_pair_measure[symmetric]) + with union show ?case + unfolding sets_pair_measure[symmetric] by simp +qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If) + +lemma (in sigma_finite_measure) measurable_emeasure_Pair: + assumes Q: "Q \<in> sets (N \<Otimes>\<^sub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _") +proof - + from sigma_finite_disjoint guess F . note F = this + then have F_sets: "\<And>i. F i \<in> sets M" by auto + let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q" + { fix i + have [simp]: "space N \<times> F i \<inter> space N \<times> space M = space N \<times> F i" + using F sets.sets_into_space by auto + let ?R = "density M (indicator (F i))" + have "finite_measure ?R" + using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq) + then have "(\<lambda>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))) \<in> borel_measurable N" + by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q) + moreover have "\<And>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q)) + = emeasure M (F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q))" + using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1) + moreover have "\<And>x. F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q) = ?C x i" + using sets.sets_into_space[OF Q] by (auto simp: space_pair_measure) + ultimately have "(\<lambda>x. emeasure M (?C x i)) \<in> borel_measurable N" + by simp } + moreover + { fix x + have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)" + proof (intro suminf_emeasure) + show "range (?C x) \<subseteq> sets M" + using F \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> by (auto intro!: sets_Pair1) + have "disjoint_family F" using F by auto + show "disjoint_family (?C x)" + by (rule disjoint_family_on_bisimulation[OF \<open>disjoint_family F\<close>]) auto + qed + also have "(\<Union>i. ?C x i) = Pair x -` Q" + using F sets.sets_into_space[OF \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close>] + by (auto simp: space_pair_measure) + finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))" + by simp } + ultimately show ?thesis using \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> F_sets + by auto +qed + +lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]: + assumes space: "\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M" + assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)" + shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N" +proof - + from space have "\<And>x. x \<in> space N \<Longrightarrow> Pair x -` {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} = A x" + by (auto simp: space_pair_measure) + with measurable_emeasure_Pair[OF A] show ?thesis + by (auto cong: measurable_cong) +qed + +lemma (in sigma_finite_measure) emeasure_pair_measure: + assumes "X \<in> sets (N \<Otimes>\<^sub>M M)" + shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X") +proof (rule emeasure_measure_of[OF pair_measure_def]) + show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>" + by (auto simp: positive_def) + have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y" + by (auto simp: indicator_def) + show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>" + proof (rule countably_additiveI) + fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^sub>M M)" "disjoint_family F" + from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^sub>M M)" by auto + moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)" + by (intro disjoint_family_on_bisimulation[OF F(2)]) auto + moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M" + using F by (auto simp: sets_Pair1) + ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)" + by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure + intro!: nn_integral_cong nn_integral_indicator[symmetric]) + qed + show "{a \<times> b |a b. a \<in> sets N \<and> b \<in> sets M} \<subseteq> Pow (space N \<times> space M)" + using sets.space_closed[of N] sets.space_closed[of M] by auto +qed fact + +lemma (in sigma_finite_measure) emeasure_pair_measure_alt: + assumes X: "X \<in> sets (N \<Otimes>\<^sub>M M)" + shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)" +proof - + have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y" + by (auto simp: indicator_def) + show ?thesis + using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure sets_Pair1) +qed + +lemma (in sigma_finite_measure) emeasure_pair_measure_Times: + assumes A: "A \<in> sets N" and B: "B \<in> sets M" + shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B" +proof - + have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)" + using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt) + also have "\<dots> = emeasure M B * emeasure N A" + using A by (simp add: nn_integral_cmult_indicator) + finally show ?thesis + by (simp add: ac_simps) +qed + +subsection \<open>Binary products of $\sigma$-finite emeasure spaces\<close> + +locale pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2 + for M1 :: "'a measure" and M2 :: "'b measure" + +lemma (in pair_sigma_finite) measurable_emeasure_Pair1: + "Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1" + using M2.measurable_emeasure_Pair . + +lemma (in pair_sigma_finite) measurable_emeasure_Pair2: + assumes Q: "Q \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2" +proof - + have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)" + using Q measurable_pair_swap' by (auto intro: measurable_sets) + note M1.measurable_emeasure_Pair[OF this] + moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1)) = (\<lambda>x. (x, y)) -` Q" + using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure) + ultimately show ?thesis by simp +qed + +lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: + defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}" + shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and> + (\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)" +proof - + from M1.sigma_finite_incseq guess F1 . note F1 = this + from M2.sigma_finite_incseq guess F2 . note F2 = this + from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto + let ?F = "\<lambda>i. F1 i \<times> F2 i" + show ?thesis + proof (intro exI[of _ ?F] conjI allI) + show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD) + next + have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)" + proof (intro subsetI) + fix x assume "x \<in> space M1 \<times> space M2" + then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j" + by (auto simp: space) + then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)" + using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_def + by (force split: split_max)+ + then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)" + by (intro SigmaI) (auto simp add: max.commute) + then show "x \<in> (\<Union>i. ?F i)" by auto + qed + then show "(\<Union>i. ?F i) = space M1 \<times> space M2" + using space by (auto simp: space) + next + fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)" + using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_Suc_iff by auto + next + fix i + from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto + with F1 F2 show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>" + by (auto simp add: emeasure_pair_measure_Times ennreal_mult_eq_top_iff) + qed +qed + +sublocale pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2" +proof + from M1.sigma_finite_countable guess F1 .. + moreover from M2.sigma_finite_countable guess F2 .. + ultimately show + "\<exists>A. countable A \<and> A \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> \<Union>A = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>a\<in>A. emeasure (M1 \<Otimes>\<^sub>M M2) a \<noteq> \<infinity>)" + by (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (F1 \<times> F2)"] conjI) + (auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq ennreal_mult_eq_top_iff) +qed + +lemma sigma_finite_pair_measure: + assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B" + shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)" +proof - + interpret A: sigma_finite_measure A by fact + interpret B: sigma_finite_measure B by fact + interpret AB: pair_sigma_finite A B .. + show ?thesis .. +qed + +lemma sets_pair_swap: + assumes "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" + shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)" + using measurable_pair_swap' assms by (rule measurable_sets) + +lemma (in pair_sigma_finite) distr_pair_swap: + "M1 \<Otimes>\<^sub>M M2 = distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D") +proof - + from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this + let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}" + show ?thesis + proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) + show "?E \<subseteq> Pow (space ?P)" + using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure) + show "sets ?P = sigma_sets (space ?P) ?E" + by (simp add: sets_pair_measure space_pair_measure) + then show "sets ?D = sigma_sets (space ?P) ?E" + by simp + next + show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>" + using F by (auto simp: space_pair_measure) + next + fix X assume "X \<in> ?E" + then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto + have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^sub>M M1) = B \<times> A" + using sets.sets_into_space[OF A] sets.sets_into_space[OF B] by (auto simp: space_pair_measure) + with A B show "emeasure (M1 \<Otimes>\<^sub>M M2) X = emeasure ?D X" + by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr + measurable_pair_swap' ac_simps) + qed +qed + +lemma (in pair_sigma_finite) emeasure_pair_measure_alt2: + assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" + shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)" + (is "_ = ?\<nu> A") +proof - + have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1))) = (\<lambda>x. (x, y)) -` A" + using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) + show ?thesis using A + by (subst distr_pair_swap) + (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap'] + M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A]) +qed + +lemma (in pair_sigma_finite) AE_pair: + assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x" + shows "AE x in M1. (AE y in M2. Q (x, y))" +proof - + obtain N where N: "N \<in> sets (M1 \<Otimes>\<^sub>M M2)" "emeasure (M1 \<Otimes>\<^sub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> Q x} \<subseteq> N" + using assms unfolding eventually_ae_filter by auto + show ?thesis + proof (rule AE_I) + from N measurable_emeasure_Pair1[OF \<open>N \<in> sets (M1 \<Otimes>\<^sub>M M2)\<close>] + show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0" + by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff) + show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1" + by (intro borel_measurable_eq measurable_emeasure_Pair1 N sets.sets_Collect_neg N) simp + { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0" + have "AE y in M2. Q (x, y)" + proof (rule AE_I) + show "emeasure M2 (Pair x -` N) = 0" by fact + show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1) + show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N" + using N \<open>x \<in> space M1\<close> unfolding space_pair_measure by auto + qed } + then show "{x \<in> space M1. \<not> (AE y in M2. Q (x, y))} \<subseteq> {x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0}" + by auto + qed +qed + +lemma (in pair_sigma_finite) AE_pair_measure: + assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)" + assumes ae: "AE x in M1. AE y in M2. P (x, y)" + shows "AE x in M1 \<Otimes>\<^sub>M M2. P x" +proof (subst AE_iff_measurable[OF _ refl]) + show "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)" + by (rule sets.sets_Collect) fact + then have "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = + (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)" + by (simp add: M2.emeasure_pair_measure) + also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. 0 \<partial>M2 \<partial>M1)" + using ae + apply (safe intro!: nn_integral_cong_AE) + apply (intro AE_I2) + apply (safe intro!: nn_integral_cong_AE) + apply auto + done + finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp +qed + +lemma (in pair_sigma_finite) AE_pair_iff: + "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> + (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x))" + using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto + +lemma (in pair_sigma_finite) AE_commute: + assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)" + shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)" +proof - + interpret Q: pair_sigma_finite M2 M1 .. + have [simp]: "\<And>x. (fst (case x of (x, y) \<Rightarrow> (y, x))) = snd x" "\<And>x. (snd (case x of (x, y) \<Rightarrow> (y, x))) = fst x" + by auto + have "{x \<in> space (M2 \<Otimes>\<^sub>M M1). P (snd x) (fst x)} = + (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^sub>M M1)" + by (auto simp: space_pair_measure) + also have "\<dots> \<in> sets (M2 \<Otimes>\<^sub>M M1)" + by (intro sets_pair_swap P) + finally show ?thesis + apply (subst AE_pair_iff[OF P]) + apply (subst distr_pair_swap) + apply (subst AE_distr_iff[OF measurable_pair_swap' P]) + apply (subst Q.AE_pair_iff) + apply simp_all + done +qed + +subsection "Fubinis theorem" + +lemma measurable_compose_Pair1: + "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L" + by simp + +lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst: + assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" + shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1" +using f proof induct + case (cong u v) + then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)" + by (auto simp: space_pair_measure) + show ?case + apply (subst measurable_cong) + apply (rule nn_integral_cong) + apply fact+ + done +next + case (set Q) + have [simp]: "\<And>x y. indicator Q (x, y) = indicator (Pair x -` Q) y" + by (auto simp: indicator_def) + have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^sup>+ y. indicator Q (x, y) \<partial>M" + by (simp add: sets_Pair1[OF set]) + from this measurable_emeasure_Pair[OF set] show ?case + by (rule measurable_cong[THEN iffD1]) +qed (simp_all add: nn_integral_add nn_integral_cmult measurable_compose_Pair1 + nn_integral_monotone_convergence_SUP incseq_def le_fun_def + cong: measurable_cong) + +lemma (in sigma_finite_measure) nn_integral_fst: + assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" + shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _") +using f proof induct + case (cong u v) + then have "?I u = ?I v" + by (intro nn_integral_cong) (auto simp: space_pair_measure) + with cong show ?case + by (simp cong: nn_integral_cong) +qed (simp_all add: emeasure_pair_measure nn_integral_cmult nn_integral_add + nn_integral_monotone_convergence_SUP measurable_compose_Pair1 + borel_measurable_nn_integral_fst nn_integral_mono incseq_def le_fun_def + cong: nn_integral_cong) + +lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]: + "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N" + using borel_measurable_nn_integral_fst[of "case_prod f" N] by simp + +lemma (in pair_sigma_finite) nn_integral_snd: + assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" + shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f" +proof - + note measurable_pair_swap[OF f] + from M1.nn_integral_fst[OF this] + have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))" + by simp + also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f" + by (subst distr_pair_swap) (auto simp add: nn_integral_distr intro!: nn_integral_cong) + finally show ?thesis . +qed + +lemma (in pair_sigma_finite) Fubini: + assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" + shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)" + unfolding nn_integral_snd[OF assms] M2.nn_integral_fst[OF assms] .. + +lemma (in pair_sigma_finite) Fubini': + assumes f: "case_prod f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" + shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f x y \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f x y \<partial>M2) \<partial>M1)" + using Fubini[OF f] by simp + +subsection \<open>Products on counting spaces, densities and distributions\<close> + +lemma sigma_prod: + assumes X_cover: "\<exists>E\<subseteq>A. countable E \<and> X = \<Union>E" and A: "A \<subseteq> Pow X" + assumes Y_cover: "\<exists>E\<subseteq>B. countable E \<and> Y = \<Union>E" and B: "B \<subseteq> Pow Y" + shows "sigma X A \<Otimes>\<^sub>M sigma Y B = sigma (X \<times> Y) {a \<times> b | a b. a \<in> A \<and> b \<in> B}" + (is "?P = ?S") +proof (rule measure_eqI) + have [simp]: "snd \<in> X \<times> Y \<rightarrow> Y" "fst \<in> X \<times> Y \<rightarrow> X" + by auto + let ?XY = "{{fst -` a \<inter> X \<times> Y | a. a \<in> A}, {snd -` b \<inter> X \<times> Y | b. b \<in> B}}" + have "sets ?P = sets (SUP xy:?XY. sigma (X \<times> Y) xy)" + by (simp add: vimage_algebra_sigma sets_pair_eq_sets_fst_snd A B) + also have "\<dots> = sets (sigma (X \<times> Y) (\<Union>?XY))" + by (intro Sup_sigma arg_cong[where f=sets]) auto + also have "\<dots> = sets ?S" + proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) + show "\<Union>?XY \<subseteq> Pow (X \<times> Y)" "{a \<times> b |a b. a \<in> A \<and> b \<in> B} \<subseteq> Pow (X \<times> Y)" + using A B by auto + next + interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}" + using A B by (intro sigma_algebra_sigma_sets) auto + fix Z assume "Z \<in> \<Union>?XY" + then show "Z \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}" + proof safe + fix a assume "a \<in> A" + from Y_cover obtain E where E: "E \<subseteq> B" "countable E" and "Y = \<Union>E" + by auto + with \<open>a \<in> A\<close> A have eq: "fst -` a \<inter> X \<times> Y = (\<Union>e\<in>E. a \<times> e)" + by auto + show "fst -` a \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}" + using \<open>a \<in> A\<close> E unfolding eq by (auto intro!: XY.countable_UN') + next + fix b assume "b \<in> B" + from X_cover obtain E where E: "E \<subseteq> A" "countable E" and "X = \<Union>E" + by auto + with \<open>b \<in> B\<close> B have eq: "snd -` b \<inter> X \<times> Y = (\<Union>e\<in>E. e \<times> b)" + by auto + show "snd -` b \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}" + using \<open>b \<in> B\<close> E unfolding eq by (auto intro!: XY.countable_UN') + qed + next + fix Z assume "Z \<in> {a \<times> b |a b. a \<in> A \<and> b \<in> B}" + then obtain a b where "Z = a \<times> b" and ab: "a \<in> A" "b \<in> B" + by auto + then have Z: "Z = (fst -` a \<inter> X \<times> Y) \<inter> (snd -` b \<inter> X \<times> Y)" + using A B by auto + interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) (\<Union>?XY)" + by (intro sigma_algebra_sigma_sets) auto + show "Z \<in> sigma_sets (X \<times> Y) (\<Union>?XY)" + unfolding Z by (rule XY.Int) (blast intro: ab)+ + qed + finally show "sets ?P = sets ?S" . +next + interpret finite_measure "sigma X A" for X A + proof qed (simp add: emeasure_sigma) + fix A assume "A \<in> sets ?P" then show "emeasure ?P A = emeasure ?S A" + by (simp add: emeasure_pair_measure_alt emeasure_sigma) +qed + +lemma sigma_sets_pair_measure_generator_finite: + assumes "finite A" and "finite B" + shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)" + (is "sigma_sets ?prod ?sets = _") +proof safe + have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product) + fix x assume subset: "x \<subseteq> A \<times> B" + hence "finite x" using fin by (rule finite_subset) + from this subset show "x \<in> sigma_sets ?prod ?sets" + proof (induct x) + case empty show ?case by (rule sigma_sets.Empty) + next + case (insert a x) + hence "{a} \<in> sigma_sets ?prod ?sets" by auto + moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto + ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) + qed +next + fix x a b + assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x" + from sigma_sets_into_sp[OF _ this(1)] this(2) + show "a \<in> A" and "b \<in> B" by auto +qed + +lemma borel_prod: + "(borel \<Otimes>\<^sub>M borel) = (borel :: ('a::second_countable_topology \<times> 'b::second_countable_topology) measure)" + (is "?P = ?B") +proof - + have "?B = sigma UNIV {A \<times> B | A B. open A \<and> open B}" + by (rule second_countable_borel_measurable[OF open_prod_generated]) + also have "\<dots> = ?P" + unfolding borel_def + by (subst sigma_prod) (auto intro!: exI[of _ "{UNIV}"]) + finally show ?thesis .. +qed + +lemma pair_measure_count_space: + assumes A: "finite A" and B: "finite B" + shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C") +proof (rule measure_eqI) + interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact + interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact + interpret P: pair_sigma_finite "count_space A" "count_space B" .. + show eq: "sets ?P = sets ?C" + by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B) + fix X assume X: "X \<in> sets ?P" + with eq have X_subset: "X \<subseteq> A \<times> B" by simp + with A B have fin_Pair: "\<And>x. finite (Pair x -` X)" + by (intro finite_subset[OF _ B]) auto + have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B) + have pos_card: "(0::ennreal) < of_nat (card (Pair x -` X)) \<longleftrightarrow> Pair x -` X \<noteq> {}" for x + by (auto simp: card_eq_0_iff fin_Pair) blast + + show "emeasure ?P X = emeasure ?C X" + using X_subset A fin_Pair fin_X + apply (subst B.emeasure_pair_measure_alt[OF X]) + apply (subst emeasure_count_space) + apply (auto simp add: emeasure_count_space nn_integral_count_space + pos_card of_nat_setsum[symmetric] card_SigmaI[symmetric] + simp del: of_nat_setsum card_SigmaI + intro!: arg_cong[where f=card]) + done +qed + + +lemma emeasure_prod_count_space: + assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M M)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)") + shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator A (x, y) \<partial>?B \<partial>?A)" + by (rule emeasure_measure_of[OF pair_measure_def]) + (auto simp: countably_additive_def positive_def suminf_indicator A + nn_integral_suminf[symmetric] dest: sets.sets_into_space) + +lemma emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1" +proof - + have [simp]: "\<And>a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ennreal)" + by (auto split: split_indicator) + show ?thesis + by (cases x) (auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair) +qed + +lemma emeasure_count_space_prod_eq: + fixes A :: "('a \<times> 'b) set" + assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M count_space UNIV)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)") + shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A" +proof - + { fix A :: "('a \<times> 'b) set" assume "countable A" + then have "emeasure (?A \<Otimes>\<^sub>M ?B) (\<Union>a\<in>A. {a}) = (\<integral>\<^sup>+a. emeasure (?A \<Otimes>\<^sub>M ?B) {a} \<partial>count_space A)" + by (intro emeasure_UN_countable) (auto simp: sets_Pair disjoint_family_on_def) + also have "\<dots> = (\<integral>\<^sup>+a. indicator A a \<partial>count_space UNIV)" + by (subst nn_integral_count_space_indicator) auto + finally have "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A" + by simp } + note * = this + + show ?thesis + proof cases + assume "finite A" then show ?thesis + by (intro * countable_finite) + next + assume "infinite A" + then obtain C where "countable C" and "infinite C" and "C \<subseteq> A" + by (auto dest: infinite_countable_subset') + with A have "emeasure (?A \<Otimes>\<^sub>M ?B) C \<le> emeasure (?A \<Otimes>\<^sub>M ?B) A" + by (intro emeasure_mono) auto + also have "emeasure (?A \<Otimes>\<^sub>M ?B) C = emeasure (count_space UNIV) C" + using \<open>countable C\<close> by (rule *) + finally show ?thesis + using \<open>infinite C\<close> \<open>infinite A\<close> by (simp add: top_unique) + qed +qed + +lemma nn_integral_count_space_prod_eq: + "nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f" + (is "nn_integral ?P f = _") +proof cases + assume cntbl: "countable {x. f x \<noteq> 0}" + have [simp]: "\<And>x. card ({x} \<inter> {x. f x \<noteq> 0}) = (indicator {x. f x \<noteq> 0} x::ennreal)" + by (auto split: split_indicator) + have [measurable]: "\<And>y. (\<lambda>x. indicator {y} x) \<in> borel_measurable ?P" + by (rule measurable_discrete_difference[of "\<lambda>x. 0" _ borel "{y}" "\<lambda>x. indicator {y} x" for y]) + (auto intro: sets_Pair) + + have "(\<integral>\<^sup>+x. f x \<partial>?P) = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x * indicator {x} x' \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)" + by (auto simp add: nn_integral_cmult nn_integral_indicator' intro!: nn_integral_cong split: split_indicator) + also have "\<dots> = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x' * indicator {x'} x \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)" + by (auto intro!: nn_integral_cong split: split_indicator) + also have "\<dots> = (\<integral>\<^sup>+x'. \<integral>\<^sup>+x. f x' * indicator {x'} x \<partial>?P \<partial>count_space {x. f x \<noteq> 0})" + by (intro nn_integral_count_space_nn_integral cntbl) auto + also have "\<dots> = (\<integral>\<^sup>+x'. f x' \<partial>count_space {x. f x \<noteq> 0})" + by (intro nn_integral_cong) (auto simp: nn_integral_cmult sets_Pair) + finally show ?thesis + by (auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) +next + { fix x assume "f x \<noteq> 0" + then have "(\<exists>r\<ge>0. 0 < r \<and> f x = ennreal r) \<or> f x = \<infinity>" + by (cases "f x" rule: ennreal_cases) (auto simp: less_le) + then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f x" + by (auto elim!: nat_approx_posE intro!: less_imp_le) } + note * = this + + assume cntbl: "uncountable {x. f x \<noteq> 0}" + also have "{x. f x \<noteq> 0} = (\<Union>n. {x. 1/Suc n \<le> f x})" + using * by auto + finally obtain n where "infinite {x. 1/Suc n \<le> f x}" + by (meson countableI_type countable_UN uncountable_infinite) + then obtain C where C: "C \<subseteq> {x. 1/Suc n \<le> f x}" and "countable C" "infinite C" + by (metis infinite_countable_subset') + + have [measurable]: "C \<in> sets ?P" + using sets.countable[OF _ \<open>countable C\<close>, of ?P] by (auto simp: sets_Pair) + + have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) \<le> nn_integral ?P f" + using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) + moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) = \<infinity>" + using \<open>infinite C\<close> by (simp add: nn_integral_cmult emeasure_count_space_prod_eq ennreal_mult_top) + moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) \<le> nn_integral (count_space UNIV) f" + using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) + moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) = \<infinity>" + using \<open>infinite C\<close> by (simp add: nn_integral_cmult ennreal_mult_top) + ultimately show ?thesis + by (simp add: top_unique) +qed + +lemma pair_measure_density: + assumes f: "f \<in> borel_measurable M1" + assumes g: "g \<in> borel_measurable M2" + assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)" + shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R") +proof (rule measure_eqI) + interpret M2: sigma_finite_measure M2 by fact + interpret D2: sigma_finite_measure "density M2 g" by fact + + fix A assume A: "A \<in> sets ?L" + with f g have "(\<integral>\<^sup>+ x. f x * \<integral>\<^sup>+ y. g y * indicator A (x, y) \<partial>M2 \<partial>M1) = + (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)" + by (intro nn_integral_cong_AE) + (auto simp add: nn_integral_cmult[symmetric] ac_simps) + with A f g show "emeasure ?L A = emeasure ?R A" + by (simp add: D2.emeasure_pair_measure emeasure_density nn_integral_density + M2.nn_integral_fst[symmetric] + cong: nn_integral_cong) +qed simp + +lemma sigma_finite_measure_distr: + assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N" + shows "sigma_finite_measure M" +proof - + interpret sigma_finite_measure "distr M N f" by fact + from sigma_finite_countable guess A .. note A = this + show ?thesis + proof + show "\<exists>A. countable A \<and> A \<subseteq> sets M \<and> \<Union>A = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)" + using A f + by (intro exI[of _ "(\<lambda>a. f -` a \<inter> space M) ` A"]) + (auto simp: emeasure_distr set_eq_iff subset_eq intro: measurable_space) + qed +qed + +lemma pair_measure_distr: + assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T" + assumes "sigma_finite_measure (distr N T g)" + shows "distr M S f \<Otimes>\<^sub>M distr N T g = distr (M \<Otimes>\<^sub>M N) (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D") +proof (rule measure_eqI) + interpret T: sigma_finite_measure "distr N T g" by fact + interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+ + + fix A assume A: "A \<in> sets ?P" + with f g show "emeasure ?P A = emeasure ?D A" + by (auto simp add: N.emeasure_pair_measure_alt space_pair_measure emeasure_distr + T.emeasure_pair_measure_alt nn_integral_distr + intro!: nn_integral_cong arg_cong[where f="emeasure N"]) +qed simp + +lemma pair_measure_eqI: + assumes "sigma_finite_measure M1" "sigma_finite_measure M2" + assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M" + assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)" + shows "M1 \<Otimes>\<^sub>M M2 = M" +proof - + interpret M1: sigma_finite_measure M1 by fact + interpret M2: sigma_finite_measure M2 by fact + interpret pair_sigma_finite M1 M2 .. + from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this + let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}" + let ?P = "M1 \<Otimes>\<^sub>M M2" + show ?thesis + proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) + show "?E \<subseteq> Pow (space ?P)" + using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure) + show "sets ?P = sigma_sets (space ?P) ?E" + by (simp add: sets_pair_measure space_pair_measure) + then show "sets M = sigma_sets (space ?P) ?E" + using sets[symmetric] by simp + next + show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>" + using F by (auto simp: space_pair_measure) + next + fix X assume "X \<in> ?E" + then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto + then have "emeasure ?P X = emeasure M1 A * emeasure M2 B" + by (simp add: M2.emeasure_pair_measure_Times) + also have "\<dots> = emeasure M (A \<times> B)" + using A B emeasure by auto + finally show "emeasure ?P X = emeasure M X" + by simp + qed +qed + +lemma sets_pair_countable: + assumes "countable S1" "countable S2" + assumes M: "sets M = Pow S1" and N: "sets N = Pow S2" + shows "sets (M \<Otimes>\<^sub>M N) = Pow (S1 \<times> S2)" +proof auto + fix x a b assume x: "x \<in> sets (M \<Otimes>\<^sub>M N)" "(a, b) \<in> x" + from sets.sets_into_space[OF x(1)] x(2) + sets_eq_imp_space_eq[of N "count_space S2"] sets_eq_imp_space_eq[of M "count_space S1"] M N + show "a \<in> S1" "b \<in> S2" + by (auto simp: space_pair_measure) +next + fix X assume X: "X \<subseteq> S1 \<times> S2" + then have "countable X" + by (metis countable_subset \<open>countable S1\<close> \<open>countable S2\<close> countable_SIGMA) + have "X = (\<Union>(a, b)\<in>X. {a} \<times> {b})" by auto + also have "\<dots> \<in> sets (M \<Otimes>\<^sub>M N)" + using X + by (safe intro!: sets.countable_UN' \<open>countable X\<close> subsetI pair_measureI) (auto simp: M N) + finally show "X \<in> sets (M \<Otimes>\<^sub>M N)" . +qed + +lemma pair_measure_countable: + assumes "countable S1" "countable S2" + shows "count_space S1 \<Otimes>\<^sub>M count_space S2 = count_space (S1 \<times> S2)" +proof (rule pair_measure_eqI) + show "sigma_finite_measure (count_space S1)" "sigma_finite_measure (count_space S2)" + using assms by (auto intro!: sigma_finite_measure_count_space_countable) + show "sets (count_space S1 \<Otimes>\<^sub>M count_space S2) = sets (count_space (S1 \<times> S2))" + by (subst sets_pair_countable[OF assms]) auto +next + fix A B assume "A \<in> sets (count_space S1)" "B \<in> sets (count_space S2)" + then show "emeasure (count_space S1) A * emeasure (count_space S2) B = + emeasure (count_space (S1 \<times> S2)) (A \<times> B)" + by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff ennreal_mult_top ennreal_top_mult) +qed + +lemma nn_integral_fst_count_space: + "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f" + (is "?lhs = ?rhs") +proof(cases) + assume *: "countable {xy. f xy \<noteq> 0}" + let ?A = "fst ` {xy. f xy \<noteq> 0}" + let ?B = "snd ` {xy. f xy \<noteq> 0}" + from * have [simp]: "countable ?A" "countable ?B" by(rule countable_image)+ + have "?lhs = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space ?A)" + by(rule nn_integral_count_space_eq) + (auto simp add: nn_integral_0_iff_AE AE_count_space not_le intro: rev_image_eqI) + also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space ?B \<partial>count_space ?A)" + by(intro nn_integral_count_space_eq nn_integral_cong)(auto intro: rev_image_eqI) + also have "\<dots> = (\<integral>\<^sup>+ xy. f xy \<partial>count_space (?A \<times> ?B))" + by(subst sigma_finite_measure.nn_integral_fst) + (simp_all add: sigma_finite_measure_count_space_countable pair_measure_countable) + also have "\<dots> = ?rhs" + by(rule nn_integral_count_space_eq)(auto intro: rev_image_eqI) + finally show ?thesis . +next + { fix xy assume "f xy \<noteq> 0" + then have "(\<exists>r\<ge>0. 0 < r \<and> f xy = ennreal r) \<or> f xy = \<infinity>" + by (cases "f xy" rule: ennreal_cases) (auto simp: less_le) + then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f xy" + by (auto elim!: nat_approx_posE intro!: less_imp_le) } + note * = this + + assume cntbl: "uncountable {xy. f xy \<noteq> 0}" + also have "{xy. f xy \<noteq> 0} = (\<Union>n. {xy. 1/Suc n \<le> f xy})" + using * by auto + finally obtain n where "infinite {xy. 1/Suc n \<le> f xy}" + by (meson countableI_type countable_UN uncountable_infinite) + then obtain C where C: "C \<subseteq> {xy. 1/Suc n \<le> f xy}" and "countable C" "infinite C" + by (metis infinite_countable_subset') + + have "\<infinity> = (\<integral>\<^sup>+ xy. ennreal (1 / Suc n) * indicator C xy \<partial>count_space UNIV)" + using \<open>infinite C\<close> by(simp add: nn_integral_cmult ennreal_mult_top) + also have "\<dots> \<le> ?rhs" using C + by(intro nn_integral_mono)(auto split: split_indicator) + finally have "?rhs = \<infinity>" by (simp add: top_unique) + moreover have "?lhs = \<infinity>" + proof(cases "finite (fst ` C)") + case True + then obtain x C' where x: "x \<in> fst ` C" + and C': "C' = fst -` {x} \<inter> C" + and "infinite C'" + using \<open>infinite C\<close> by(auto elim!: inf_img_fin_domE') + from x C C' have **: "C' \<subseteq> {xy. 1 / Suc n \<le> f xy}" by auto + + from C' \<open>infinite C'\<close> have "infinite (snd ` C')" + by(auto dest!: finite_imageD simp add: inj_on_def) + then have "\<infinity> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator (snd ` C') y \<partial>count_space UNIV)" + by(simp add: nn_integral_cmult ennreal_mult_top) + also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV)" + by(rule nn_integral_cong)(force split: split_indicator intro: rev_image_eqI simp add: C') + also have "\<dots> = (\<integral>\<^sup>+ x'. (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV) * indicator {x} x' \<partial>count_space UNIV)" + by(simp add: one_ereal_def[symmetric]) + also have "\<dots> \<le> (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV \<partial>count_space UNIV)" + by(rule nn_integral_mono)(simp split: split_indicator) + also have "\<dots> \<le> ?lhs" using ** + by(intro nn_integral_mono)(auto split: split_indicator) + finally show ?thesis by (simp add: top_unique) + next + case False + define C' where "C' = fst ` C" + have "\<infinity> = \<integral>\<^sup>+ x. ennreal (1 / Suc n) * indicator C' x \<partial>count_space UNIV" + using C'_def False by(simp add: nn_integral_cmult ennreal_mult_top) + also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \<in> C} y \<partial>count_space UNIV \<partial>count_space UNIV" + by(auto simp add: one_ereal_def[symmetric] max_def intro: nn_integral_cong) + also have "\<dots> \<le> \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C (x, y) \<partial>count_space UNIV \<partial>count_space UNIV" + by(intro nn_integral_mono)(auto simp add: C'_def split: split_indicator intro: someI) + also have "\<dots> \<le> ?lhs" using C + by(intro nn_integral_mono)(auto split: split_indicator) + finally show ?thesis by (simp add: top_unique) + qed + ultimately show ?thesis by simp +qed + +lemma nn_integral_snd_count_space: + "(\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f" + (is "?lhs = ?rhs") +proof - + have "?lhs = (\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. (\<lambda>(y, x). f (x, y)) (y, x) \<partial>count_space UNIV \<partial>count_space UNIV)" + by(simp) + also have "\<dots> = \<integral>\<^sup>+ yx. (\<lambda>(y, x). f (x, y)) yx \<partial>count_space UNIV" + by(rule nn_integral_fst_count_space) + also have "\<dots> = \<integral>\<^sup>+ xy. f xy \<partial>count_space ((\<lambda>(x, y). (y, x)) ` UNIV)" + by(subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric]) + (simp_all add: inj_on_def split_def) + also have "\<dots> = ?rhs" by(rule nn_integral_count_space_eq) auto + finally show ?thesis . +qed + +lemma measurable_pair_measure_countable1: + assumes "countable A" + and [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N K" + shows "f \<in> measurable (count_space A \<Otimes>\<^sub>M N) K" +using _ _ assms(1) +by(rule measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all + +subsection \<open>Product of Borel spaces\<close> + +lemma borel_Times: + fixes A :: "'a::topological_space set" and B :: "'b::topological_space set" + assumes A: "A \<in> sets borel" and B: "B \<in> sets borel" + shows "A \<times> B \<in> sets borel" +proof - + have "A \<times> B = (A\<times>UNIV) \<inter> (UNIV \<times> B)" + by auto + moreover + { have "A \<in> sigma_sets UNIV {S. open S}" using A by (simp add: sets_borel) + then have "A\<times>UNIV \<in> sets borel" + proof (induct A) + case (Basic S) then show ?case + by (auto intro!: borel_open open_Times) + next + case (Compl A) + moreover have *: "(UNIV - A) \<times> UNIV = UNIV - (A \<times> UNIV)" + by auto + ultimately show ?case + unfolding * by auto + next + case (Union A) + moreover have *: "(UNION UNIV A) \<times> UNIV = UNION UNIV (\<lambda>i. A i \<times> UNIV)" + by auto + ultimately show ?case + unfolding * by auto + qed simp } + moreover + { have "B \<in> sigma_sets UNIV {S. open S}" using B by (simp add: sets_borel) + then have "UNIV\<times>B \<in> sets borel" + proof (induct B) + case (Basic S) then show ?case + by (auto intro!: borel_open open_Times) + next + case (Compl B) + moreover have *: "UNIV \<times> (UNIV - B) = UNIV - (UNIV \<times> B)" + by auto + ultimately show ?case + unfolding * by auto + next + case (Union B) + moreover have *: "UNIV \<times> (UNION UNIV B) = UNION UNIV (\<lambda>i. UNIV \<times> B i)" + by auto + ultimately show ?case + unfolding * by auto + qed simp } + ultimately show ?thesis + by auto +qed + +lemma finite_measure_pair_measure: + assumes "finite_measure M" "finite_measure N" + shows "finite_measure (N \<Otimes>\<^sub>M M)" +proof (rule finite_measureI) + interpret M: finite_measure M by fact + interpret N: finite_measure N by fact + show "emeasure (N \<Otimes>\<^sub>M M) (space (N \<Otimes>\<^sub>M M)) \<noteq> \<infinity>" + by (auto simp: space_pair_measure M.emeasure_pair_measure_Times ennreal_mult_eq_top_iff) +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Bochner_Integration.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,3066 @@ +(* Title: HOL/Analysis/Bochner_Integration.thy + Author: Johannes Hölzl, TU München +*) + +section \<open>Bochner Integration for Vector-Valued Functions\<close> + +theory Bochner_Integration + imports Finite_Product_Measure +begin + +text \<open> + +In the following development of the Bochner integral we use second countable topologies instead +of separable spaces. A second countable topology is also separable. + +\<close> + +lemma borel_measurable_implies_sequence_metric: + fixes f :: "'a \<Rightarrow> 'b :: {metric_space, second_countable_topology}" + assumes [measurable]: "f \<in> borel_measurable M" + shows "\<exists>F. (\<forall>i. simple_function M (F i)) \<and> (\<forall>x\<in>space M. (\<lambda>i. F i x) \<longlonglongrightarrow> f x) \<and> + (\<forall>i. \<forall>x\<in>space M. dist (F i x) z \<le> 2 * dist (f x) z)" +proof - + obtain D :: "'b set" where "countable D" and D: "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d\<in>D. d \<in> X" + by (erule countable_dense_setE) + + define e where "e = from_nat_into D" + { fix n x + obtain d where "d \<in> D" and d: "d \<in> ball x (1 / Suc n)" + using D[of "ball x (1 / Suc n)"] by auto + from \<open>d \<in> D\<close> D[of UNIV] \<open>countable D\<close> obtain i where "d = e i" + unfolding e_def by (auto dest: from_nat_into_surj) + with d have "\<exists>i. dist x (e i) < 1 / Suc n" + by auto } + note e = this + + define A where [abs_def]: "A m n = + {x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> dist (f x) z}" for m n + define B where [abs_def]: "B m = disjointed (A m)" for m + + define m where [abs_def]: "m N x = Max {m. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}" for N x + define F where [abs_def]: "F N x = + (if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n) + then e (LEAST n. x \<in> B (m N x) n) else z)" for N x + + have B_imp_A[intro, simp]: "\<And>x m n. x \<in> B m n \<Longrightarrow> x \<in> A m n" + using disjointed_subset[of "A m" for m] unfolding B_def by auto + + { fix m + have "\<And>n. A m n \<in> sets M" + by (auto simp: A_def) + then have "\<And>n. B m n \<in> sets M" + using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) } + note this[measurable] + + { fix N i x assume "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" + then have "m N x \<in> {m::nat. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}" + unfolding m_def by (intro Max_in) auto + then have "m N x \<le> N" "\<exists>n\<le>N. x \<in> B (m N x) n" + by auto } + note m = this + + { fix j N i x assume "j \<le> N" "i \<le> N" "x \<in> B j i" + then have "j \<le> m N x" + unfolding m_def by (intro Max_ge) auto } + note m_upper = this + + show ?thesis + unfolding simple_function_def + proof (safe intro!: exI[of _ F]) + have [measurable]: "\<And>i. F i \<in> borel_measurable M" + unfolding F_def m_def by measurable + show "\<And>x i. F i -` {x} \<inter> space M \<in> sets M" + by measurable + + { fix i + { fix n x assume "x \<in> B (m i x) n" + then have "(LEAST n. x \<in> B (m i x) n) \<le> n" + by (intro Least_le) + also assume "n \<le> i" + finally have "(LEAST n. x \<in> B (m i x) n) \<le> i" . } + then have "F i ` space M \<subseteq> {z} \<union> e ` {.. i}" + by (auto simp: F_def) + then show "finite (F i ` space M)" + by (rule finite_subset) auto } + + { fix N i n x assume "i \<le> N" "n \<le> N" "x \<in> B i n" + then have 1: "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" by auto + from m[OF this] obtain n where n: "m N x \<le> N" "n \<le> N" "x \<in> B (m N x) n" by auto + moreover + define L where "L = (LEAST n. x \<in> B (m N x) n)" + have "dist (f x) (e L) < 1 / Suc (m N x)" + proof - + have "x \<in> B (m N x) L" + using n(3) unfolding L_def by (rule LeastI) + then have "x \<in> A (m N x) L" + by auto + then show ?thesis + unfolding A_def by simp + qed + ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)" + by (auto simp add: F_def L_def) } + note * = this + + fix x assume "x \<in> space M" + show "(\<lambda>i. F i x) \<longlonglongrightarrow> f x" + proof cases + assume "f x = z" + then have "\<And>i n. x \<notin> A i n" + unfolding A_def by auto + then have "\<And>i. F i x = z" + by (auto simp: F_def) + then show ?thesis + using \<open>f x = z\<close> by auto + next + assume "f x \<noteq> z" + + show ?thesis + proof (rule tendstoI) + fix e :: real assume "0 < e" + with \<open>f x \<noteq> z\<close> obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" + by (metis dist_nz order_less_trans neq_iff nat_approx_posE) + with \<open>x\<in>space M\<close> \<open>f x \<noteq> z\<close> have "x \<in> (\<Union>i. B n i)" + unfolding A_def B_def UN_disjointed_eq using e by auto + then obtain i where i: "x \<in> B n i" by auto + + show "eventually (\<lambda>i. dist (F i x) (f x) < e) sequentially" + using eventually_ge_at_top[of "max n i"] + proof eventually_elim + fix j assume j: "max n i \<le> j" + with i have "dist (f x) (F j x) < 1 / Suc (m j x)" + by (intro *[OF _ _ i]) auto + also have "\<dots> \<le> 1 / Suc n" + using j m_upper[OF _ _ i] + by (auto simp: field_simps) + also note \<open>1 / Suc n < e\<close> + finally show "dist (F j x) (f x) < e" + by (simp add: less_imp_le dist_commute) + qed + qed + qed + fix i + { fix n m assume "x \<in> A n m" + then have "dist (e m) (f x) + dist (f x) z \<le> 2 * dist (f x) z" + unfolding A_def by (auto simp: dist_commute) + also have "dist (e m) z \<le> dist (e m) (f x) + dist (f x) z" + by (rule dist_triangle) + finally (xtrans) have "dist (e m) z \<le> 2 * dist (f x) z" . } + then show "dist (F i x) z \<le> 2 * dist (f x) z" + unfolding F_def + apply auto + apply (rule LeastI2) + apply auto + done + qed +qed + +lemma + fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A" + shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)" + and setsum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator (B x) (g x) * f x) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)" + unfolding indicator_def + using assms by (auto intro!: setsum.mono_neutral_cong_right split: if_split_asm) + +lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]: + fixes P :: "('a \<Rightarrow> real) \<Rightarrow> bool" + assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x" + assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" + assumes mult: "\<And>u c. 0 \<le> c \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)" + assumes add: "\<And>u v. u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)" + assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. U i x) \<longlonglongrightarrow> u x) \<Longrightarrow> P u" + shows "P u" +proof - + have "(\<lambda>x. ennreal (u x)) \<in> borel_measurable M" using u by auto + from borel_measurable_implies_simple_function_sequence'[OF this] + obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and + sup: "\<And>x. (SUP i. U i x) = ennreal (u x)" + by blast + + define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x + then have U'_sf[measurable]: "\<And>i. simple_function M (U' i)" + using U by (auto intro!: simple_function_compose1[where g=enn2real]) + + show "P u" + proof (rule seq) + show U': "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> U' i x" for i + using U by (auto + intro: borel_measurable_simple_function + intro!: borel_measurable_enn2real borel_measurable_times + simp: U'_def zero_le_mult_iff enn2real_nonneg) + show "incseq U'" + using U(2,3) + by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono) + + fix x assume x: "x \<in> space M" + have "(\<lambda>i. U i x) \<longlonglongrightarrow> (SUP i. U i x)" + using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def) + moreover have "(\<lambda>i. U i x) = (\<lambda>i. ennreal (U' i x))" + using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute) + moreover have "(SUP i. U i x) = ennreal (u x)" + using sup u(2) by (simp add: max_def) + ultimately show "(\<lambda>i. U' i x) \<longlonglongrightarrow> u x" + using u U' by simp + next + fix i + have "U' i ` space M \<subseteq> enn2real ` (U i ` space M)" "finite (U i ` space M)" + unfolding U'_def using U(1) by (auto dest: simple_functionD) + then have fin: "finite (U' i ` space M)" + by (metis finite_subset finite_imageI) + moreover have "\<And>z. {y. U' i z = y \<and> y \<in> U' i ` space M \<and> z \<in> space M} = (if z \<in> space M then {U' i z} else {})" + by auto + ultimately have U': "(\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z) = U' i" + by (simp add: U'_def fun_eq_iff) + have "\<And>x. x \<in> U' i ` space M \<Longrightarrow> 0 \<le> x" + by (auto simp: U'_def enn2real_nonneg) + with fin have "P (\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z)" + proof induct + case empty from set[of "{}"] show ?case + by (simp add: indicator_def[abs_def]) + next + case (insert x F) + then show ?case + by (auto intro!: add mult set setsum_nonneg split: split_indicator split_indicator_asm + simp del: setsum_mult_indicator simp: setsum_nonneg_eq_0_iff) + qed + with U' show "P (U' i)" by simp + qed +qed + +lemma scaleR_cong_right: + fixes x :: "'a :: real_vector" + shows "(x \<noteq> 0 \<Longrightarrow> r = p) \<Longrightarrow> r *\<^sub>R x = p *\<^sub>R x" + by (cases "x = 0") auto + +inductive simple_bochner_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> bool" for M f where + "simple_function M f \<Longrightarrow> emeasure M {y\<in>space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow> + simple_bochner_integrable M f" + +lemma simple_bochner_integrable_compose2: + assumes p_0: "p 0 0 = 0" + shows "simple_bochner_integrable M f \<Longrightarrow> simple_bochner_integrable M g \<Longrightarrow> + simple_bochner_integrable M (\<lambda>x. p (f x) (g x))" +proof (safe intro!: simple_bochner_integrable.intros elim!: simple_bochner_integrable.cases del: notI) + assume sf: "simple_function M f" "simple_function M g" + then show "simple_function M (\<lambda>x. p (f x) (g x))" + by (rule simple_function_compose2) + + from sf have [measurable]: + "f \<in> measurable M (count_space UNIV)" + "g \<in> measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function) + + assume fin: "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" "emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity>" + + have "emeasure M {x\<in>space M. p (f x) (g x) \<noteq> 0} \<le> + emeasure M ({x\<in>space M. f x \<noteq> 0} \<union> {x\<in>space M. g x \<noteq> 0})" + by (intro emeasure_mono) (auto simp: p_0) + also have "\<dots> \<le> emeasure M {x\<in>space M. f x \<noteq> 0} + emeasure M {x\<in>space M. g x \<noteq> 0}" + by (intro emeasure_subadditive) auto + finally show "emeasure M {y \<in> space M. p (f y) (g y) \<noteq> 0} \<noteq> \<infinity>" + using fin by (auto simp: top_unique) +qed + +lemma simple_function_finite_support: + assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and nn: "\<And>x. 0 \<le> f x" + shows "emeasure M {x\<in>space M. f x \<noteq> 0} \<noteq> \<infinity>" +proof cases + from f have meas[measurable]: "f \<in> borel_measurable M" + by (rule borel_measurable_simple_function) + + assume non_empty: "\<exists>x\<in>space M. f x \<noteq> 0" + + define m where "m = Min (f`space M - {0})" + have "m \<in> f`space M - {0}" + unfolding m_def using f non_empty by (intro Min_in) (auto simp: simple_function_def) + then have m: "0 < m" + using nn by (auto simp: less_le) + + from m have "m * emeasure M {x\<in>space M. 0 \<noteq> f x} = + (\<integral>\<^sup>+x. m * indicator {x\<in>space M. 0 \<noteq> f x} x \<partial>M)" + using f by (intro nn_integral_cmult_indicator[symmetric]) auto + also have "\<dots> \<le> (\<integral>\<^sup>+x. f x \<partial>M)" + using AE_space + proof (intro nn_integral_mono_AE, eventually_elim) + fix x assume "x \<in> space M" + with nn show "m * indicator {x \<in> space M. 0 \<noteq> f x} x \<le> f x" + using f by (auto split: split_indicator simp: simple_function_def m_def) + qed + also note \<open>\<dots> < \<infinity>\<close> + finally show ?thesis + using m by (auto simp: ennreal_mult_less_top) +next + assume "\<not> (\<exists>x\<in>space M. f x \<noteq> 0)" + with nn have *: "{x\<in>space M. f x \<noteq> 0} = {}" + by auto + show ?thesis unfolding * by simp +qed + +lemma simple_bochner_integrableI_bounded: + assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" + shows "simple_bochner_integrable M f" +proof + have "emeasure M {y \<in> space M. ennreal (norm (f y)) \<noteq> 0} \<noteq> \<infinity>" + proof (rule simple_function_finite_support) + show "simple_function M (\<lambda>x. ennreal (norm (f x)))" + using f by (rule simple_function_compose1) + show "(\<integral>\<^sup>+ y. ennreal (norm (f y)) \<partial>M) < \<infinity>" by fact + qed simp + then show "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp +qed fact + +definition simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where + "simple_bochner_integral M f = (\<Sum>y\<in>f`space M. measure M {x\<in>space M. f x = y} *\<^sub>R y)" + +lemma simple_bochner_integral_partition: + assumes f: "simple_bochner_integrable M f" and g: "simple_function M g" + assumes sub: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> g x = g y \<Longrightarrow> f x = f y" + assumes v: "\<And>x. x \<in> space M \<Longrightarrow> f x = v (g x)" + shows "simple_bochner_integral M f = (\<Sum>y\<in>g ` space M. measure M {x\<in>space M. g x = y} *\<^sub>R v y)" + (is "_ = ?r") +proof - + from f g have [simp]: "finite (f`space M)" "finite (g`space M)" + by (auto simp: simple_function_def elim: simple_bochner_integrable.cases) + + from f have [measurable]: "f \<in> measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) + + from g have [measurable]: "g \<in> measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) + + { fix y assume "y \<in> space M" + then have "f ` space M \<inter> {i. \<exists>x\<in>space M. i = f x \<and> g y = g x} = {v (g y)}" + by (auto cong: sub simp: v[symmetric]) } + note eq = this + + have "simple_bochner_integral M f = + (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. + if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} else 0) *\<^sub>R y)" + unfolding simple_bochner_integral_def + proof (safe intro!: setsum.cong scaleR_cong_right) + fix y assume y: "y \<in> space M" "f y \<noteq> 0" + have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} = + {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}" + by auto + have eq:"{x \<in> space M. f x = f y} = + (\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> space M. g x = i})" + by (auto simp: eq_commute cong: sub rev_conj_cong) + have "finite (g`space M)" by simp + then have "finite {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}" + by (rule rev_finite_subset) auto + moreover + { fix x assume "x \<in> space M" "f x = f y" + then have "x \<in> space M" "f x \<noteq> 0" + using y by auto + then have "emeasure M {y \<in> space M. g y = g x} \<le> emeasure M {y \<in> space M. f y \<noteq> 0}" + by (auto intro!: emeasure_mono cong: sub) + then have "emeasure M {xa \<in> space M. g xa = g x} < \<infinity>" + using f by (auto simp: simple_bochner_integrable.simps less_top) } + ultimately + show "measure M {x \<in> space M. f x = f y} = + (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then measure M {x \<in> space M. g x = z} else 0)" + apply (simp add: setsum.If_cases eq) + apply (subst measure_finite_Union[symmetric]) + apply (auto simp: disjoint_family_on_def less_top) + done + qed + also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. + if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} *\<^sub>R y else 0))" + by (auto intro!: setsum.cong simp: scaleR_setsum_left) + also have "\<dots> = ?r" + by (subst setsum.commute) + (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq) + finally show "simple_bochner_integral M f = ?r" . +qed + +lemma simple_bochner_integral_add: + assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" + shows "simple_bochner_integral M (\<lambda>x. f x + g x) = + simple_bochner_integral M f + simple_bochner_integral M g" +proof - + from f g have "simple_bochner_integral M (\<lambda>x. f x + g x) = + (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> space M. (f x, g x) = y} *\<^sub>R (fst y + snd y))" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + moreover from f g have "simple_bochner_integral M f = + (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> space M. (f x, g x) = y} *\<^sub>R fst y)" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + moreover from f g have "simple_bochner_integral M g = + (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> space M. (f x, g x) = y} *\<^sub>R snd y)" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + ultimately show ?thesis + by (simp add: setsum.distrib[symmetric] scaleR_add_right) +qed + +lemma (in linear) simple_bochner_integral_linear: + assumes g: "simple_bochner_integrable M g" + shows "simple_bochner_integral M (\<lambda>x. f (g x)) = f (simple_bochner_integral M g)" +proof - + from g have "simple_bochner_integral M (\<lambda>x. f (g x)) = + (\<Sum>y\<in>g ` space M. measure M {x \<in> space M. g x = y} *\<^sub>R f y)" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"] zero + elim: simple_bochner_integrable.cases) + also have "\<dots> = f (simple_bochner_integral M g)" + by (simp add: simple_bochner_integral_def setsum scaleR) + finally show ?thesis . +qed + +lemma simple_bochner_integral_minus: + assumes f: "simple_bochner_integrable M f" + shows "simple_bochner_integral M (\<lambda>x. - f x) = - simple_bochner_integral M f" +proof - + interpret linear uminus by unfold_locales auto + from f show ?thesis + by (rule simple_bochner_integral_linear) +qed + +lemma simple_bochner_integral_diff: + assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" + shows "simple_bochner_integral M (\<lambda>x. f x - g x) = + simple_bochner_integral M f - simple_bochner_integral M g" + unfolding diff_conv_add_uminus using f g + by (subst simple_bochner_integral_add) + (auto simp: simple_bochner_integral_minus simple_bochner_integrable_compose2[where p="\<lambda>x y. - y"]) + +lemma simple_bochner_integral_norm_bound: + assumes f: "simple_bochner_integrable M f" + shows "norm (simple_bochner_integral M f) \<le> simple_bochner_integral M (\<lambda>x. norm (f x))" +proof - + have "norm (simple_bochner_integral M f) \<le> + (\<Sum>y\<in>f ` space M. norm (measure M {x \<in> space M. f x = y} *\<^sub>R y))" + unfolding simple_bochner_integral_def by (rule norm_setsum) + also have "\<dots> = (\<Sum>y\<in>f ` space M. measure M {x \<in> space M. f x = y} *\<^sub>R norm y)" + by simp + also have "\<dots> = simple_bochner_integral M (\<lambda>x. norm (f x))" + using f + by (intro simple_bochner_integral_partition[symmetric]) + (auto intro: f simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + finally show ?thesis . +qed + +lemma simple_bochner_integral_nonneg[simp]: + fixes f :: "'a \<Rightarrow> real" + shows "(\<And>x. 0 \<le> f x) \<Longrightarrow> 0 \<le> simple_bochner_integral M f" + by (simp add: setsum_nonneg simple_bochner_integral_def) + +lemma simple_bochner_integral_eq_nn_integral: + assumes f: "simple_bochner_integrable M f" "\<And>x. 0 \<le> f x" + shows "simple_bochner_integral M f = (\<integral>\<^sup>+x. f x \<partial>M)" +proof - + { fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ennreal x * y = ennreal x * z" + by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) } + note ennreal_cong_mult = this + + have [measurable]: "f \<in> borel_measurable M" + using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + { fix y assume y: "y \<in> space M" "f y \<noteq> 0" + have "ennreal (measure M {x \<in> space M. f x = f y}) = emeasure M {x \<in> space M. f x = f y}" + proof (rule emeasure_eq_ennreal_measure[symmetric]) + have "emeasure M {x \<in> space M. f x = f y} \<le> emeasure M {x \<in> space M. f x \<noteq> 0}" + using y by (intro emeasure_mono) auto + with f show "emeasure M {x \<in> space M. f x = f y} \<noteq> top" + by (auto simp: simple_bochner_integrable.simps top_unique) + qed + moreover have "{x \<in> space M. f x = f y} = (\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M" + using f by auto + ultimately have "ennreal (measure M {x \<in> space M. f x = f y}) = + emeasure M ((\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M)" by simp } + with f have "simple_bochner_integral M f = (\<integral>\<^sup>Sx. f x \<partial>M)" + unfolding simple_integral_def + by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real]) + (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases + intro!: setsum.cong ennreal_cong_mult + simp: setsum_ennreal[symmetric] ac_simps ennreal_mult + simp del: setsum_ennreal) + also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)" + using f + by (intro nn_integral_eq_simple_integral[symmetric]) + (auto simp: simple_function_compose1 simple_bochner_integrable.simps) + finally show ?thesis . +qed + +lemma simple_bochner_integral_bounded: + fixes f :: "'a \<Rightarrow> 'b::{real_normed_vector, second_countable_topology}" + assumes f[measurable]: "f \<in> borel_measurable M" + assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t" + shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le> + (\<integral>\<^sup>+ x. norm (f x - s x) \<partial>M) + (\<integral>\<^sup>+ x. norm (f x - t x) \<partial>M)" + (is "ennreal (norm (?s - ?t)) \<le> ?S + ?T") +proof - + have [measurable]: "s \<in> borel_measurable M" "t \<in> borel_measurable M" + using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\<lambda>x. s x - t x))" + using s t by (subst simple_bochner_integral_diff) auto + also have "\<dots> \<le> simple_bochner_integral M (\<lambda>x. norm (s x - t x))" + using simple_bochner_integrable_compose2[of "op -" M "s" "t"] s t + by (auto intro!: simple_bochner_integral_norm_bound) + also have "\<dots> = (\<integral>\<^sup>+x. norm (s x - t x) \<partial>M)" + using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t + by (auto intro!: simple_bochner_integral_eq_nn_integral) + also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)" + by (auto intro!: nn_integral_mono simp: ennreal_plus[symmetric] simp del: ennreal_plus) + (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans + norm_minus_commute norm_triangle_ineq4 order_refl) + also have "\<dots> = ?S + ?T" + by (rule nn_integral_add) auto + finally show ?thesis . +qed + +inductive has_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b::{real_normed_vector, second_countable_topology} \<Rightarrow> bool" + for M f x where + "f \<in> borel_measurable M \<Longrightarrow> + (\<And>i. simple_bochner_integrable M (s i)) \<Longrightarrow> + (\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0 \<Longrightarrow> + (\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x \<Longrightarrow> + has_bochner_integral M f x" + +lemma has_bochner_integral_cong: + assumes "M = N" "\<And>x. x \<in> space N \<Longrightarrow> f x = g x" "x = y" + shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral N g y" + unfolding has_bochner_integral.simps assms(1,3) + using assms(2) by (simp cong: measurable_cong_strong nn_integral_cong_strong) + +lemma has_bochner_integral_cong_AE: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> + has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x" + unfolding has_bochner_integral.simps + by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\<lambda>x. x \<longlonglongrightarrow> 0"] + nn_integral_cong_AE) + auto + +lemma borel_measurable_has_bochner_integral: + "has_bochner_integral M f x \<Longrightarrow> f \<in> borel_measurable M" + by (rule has_bochner_integral.cases) + +lemma borel_measurable_has_bochner_integral'[measurable_dest]: + "has_bochner_integral M f x \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N" + using borel_measurable_has_bochner_integral[measurable] by measurable + +lemma has_bochner_integral_simple_bochner_integrable: + "simple_bochner_integrable M f \<Longrightarrow> has_bochner_integral M f (simple_bochner_integral M f)" + by (rule has_bochner_integral.intros[where s="\<lambda>_. f"]) + (auto intro: borel_measurable_simple_function + elim: simple_bochner_integrable.cases + simp: zero_ennreal_def[symmetric]) + +lemma has_bochner_integral_real_indicator: + assumes [measurable]: "A \<in> sets M" and A: "emeasure M A < \<infinity>" + shows "has_bochner_integral M (indicator A) (measure M A)" +proof - + have sbi: "simple_bochner_integrable M (indicator A::'a \<Rightarrow> real)" + proof + have "{y \<in> space M. (indicator A y::real) \<noteq> 0} = A" + using sets.sets_into_space[OF \<open>A\<in>sets M\<close>] by (auto split: split_indicator) + then show "emeasure M {y \<in> space M. (indicator A y::real) \<noteq> 0} \<noteq> \<infinity>" + using A by auto + qed (rule simple_function_indicator assms)+ + moreover have "simple_bochner_integral M (indicator A) = measure M A" + using simple_bochner_integral_eq_nn_integral[OF sbi] A + by (simp add: ennreal_indicator emeasure_eq_ennreal_measure) + ultimately show ?thesis + by (metis has_bochner_integral_simple_bochner_integrable) +qed + +lemma has_bochner_integral_add[intro]: + "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow> + has_bochner_integral M (\<lambda>x. f x + g x) (x + y)" +proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) + fix sf sg + assume f_sf: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - sf i x) \<partial>M) \<longlonglongrightarrow> 0" + assume g_sg: "(\<lambda>i. \<integral>\<^sup>+ x. norm (g x - sg i x) \<partial>M) \<longlonglongrightarrow> 0" + + assume sf: "\<forall>i. simple_bochner_integrable M (sf i)" + and sg: "\<forall>i. simple_bochner_integrable M (sg i)" + then have [measurable]: "\<And>i. sf i \<in> borel_measurable M" "\<And>i. sg i \<in> borel_measurable M" + by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + assume [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" + + show "\<And>i. simple_bochner_integrable M (\<lambda>x. sf i x + sg i x)" + using sf sg by (simp add: simple_bochner_integrable_compose2) + + show "(\<lambda>i. \<integral>\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \<partial>M) \<longlonglongrightarrow> 0" + (is "?f \<longlonglongrightarrow> 0") + proof (rule tendsto_sandwich) + show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0" + by auto + show "eventually (\<lambda>i. ?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) \<partial>M) + \<integral>\<^sup>+ x. (norm (g x - sg i x)) \<partial>M) sequentially" + (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially") + proof (intro always_eventually allI) + fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)" + by (auto intro!: nn_integral_mono norm_diff_triangle_ineq + simp del: ennreal_plus simp add: ennreal_plus[symmetric]) + also have "\<dots> = ?g i" + by (intro nn_integral_add) auto + finally show "?f i \<le> ?g i" . + qed + show "?g \<longlonglongrightarrow> 0" + using tendsto_add[OF f_sf g_sg] by simp + qed +qed (auto simp: simple_bochner_integral_add tendsto_add) + +lemma has_bochner_integral_bounded_linear: + assumes "bounded_linear T" + shows "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M (\<lambda>x. T (f x)) (T x)" +proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) + interpret T: bounded_linear T by fact + have [measurable]: "T \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 T.continuous_on continuous_on_id) + assume [measurable]: "f \<in> borel_measurable M" + then show "(\<lambda>x. T (f x)) \<in> borel_measurable M" + by auto + + fix s assume f_s: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" + assume s: "\<forall>i. simple_bochner_integrable M (s i)" + then show "\<And>i. simple_bochner_integrable M (\<lambda>x. T (s i x))" + by (auto intro: simple_bochner_integrable_compose2 T.zero) + + have [measurable]: "\<And>i. s i \<in> borel_measurable M" + using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + obtain K where K: "K > 0" "\<And>x i. norm (T (f x) - T (s i x)) \<le> norm (f x - s i x) * K" + using T.pos_bounded by (auto simp: T.diff[symmetric]) + + show "(\<lambda>i. \<integral>\<^sup>+ x. norm (T (f x) - T (s i x)) \<partial>M) \<longlonglongrightarrow> 0" + (is "?f \<longlonglongrightarrow> 0") + proof (rule tendsto_sandwich) + show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0" + by auto + + show "eventually (\<lambda>i. ?f i \<le> K * (\<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M)) sequentially" + (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially") + proof (intro always_eventually allI) + fix i have "?f i \<le> (\<integral>\<^sup>+ x. ennreal K * norm (f x - s i x) \<partial>M)" + using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric]) + also have "\<dots> = ?g i" + using K by (intro nn_integral_cmult) auto + finally show "?f i \<le> ?g i" . + qed + show "?g \<longlonglongrightarrow> 0" + using ennreal_tendsto_cmult[OF _ f_s] by simp + qed + + assume "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" + with s show "(\<lambda>i. simple_bochner_integral M (\<lambda>x. T (s i x))) \<longlonglongrightarrow> T x" + by (auto intro!: T.tendsto simp: T.simple_bochner_integral_linear) +qed + +lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0" + by (auto intro!: has_bochner_integral.intros[where s="\<lambda>_ _. 0"] + simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps + simple_bochner_integral_def image_constant_conv) + +lemma has_bochner_integral_scaleR_left[intro]: + "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x *\<^sub>R c) (x *\<^sub>R c)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left]) + +lemma has_bochner_integral_scaleR_right[intro]: + "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c *\<^sub>R f x) (c *\<^sub>R x)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right]) + +lemma has_bochner_integral_mult_left[intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x * c) (x * c)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left]) + +lemma has_bochner_integral_mult_right[intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c * f x) (c * x)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right]) + +lemmas has_bochner_integral_divide = + has_bochner_integral_bounded_linear[OF bounded_linear_divide] + +lemma has_bochner_integral_divide_zero[intro]: + fixes c :: "_::{real_normed_field, field, second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x / c) (x / c)" + using has_bochner_integral_divide by (cases "c = 0") auto + +lemma has_bochner_integral_inner_left[intro]: + "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x \<bullet> c) (x \<bullet> c)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left]) + +lemma has_bochner_integral_inner_right[intro]: + "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c \<bullet> f x) (c \<bullet> x)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right]) + +lemmas has_bochner_integral_minus = + has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] +lemmas has_bochner_integral_Re = + has_bochner_integral_bounded_linear[OF bounded_linear_Re] +lemmas has_bochner_integral_Im = + has_bochner_integral_bounded_linear[OF bounded_linear_Im] +lemmas has_bochner_integral_cnj = + has_bochner_integral_bounded_linear[OF bounded_linear_cnj] +lemmas has_bochner_integral_of_real = + has_bochner_integral_bounded_linear[OF bounded_linear_of_real] +lemmas has_bochner_integral_fst = + has_bochner_integral_bounded_linear[OF bounded_linear_fst] +lemmas has_bochner_integral_snd = + has_bochner_integral_bounded_linear[OF bounded_linear_snd] + +lemma has_bochner_integral_indicator: + "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> + has_bochner_integral M (\<lambda>x. indicator A x *\<^sub>R c) (measure M A *\<^sub>R c)" + by (intro has_bochner_integral_scaleR_left has_bochner_integral_real_indicator) + +lemma has_bochner_integral_diff: + "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow> + has_bochner_integral M (\<lambda>x. f x - g x) (x - y)" + unfolding diff_conv_add_uminus + by (intro has_bochner_integral_add has_bochner_integral_minus) + +lemma has_bochner_integral_setsum: + "(\<And>i. i \<in> I \<Longrightarrow> has_bochner_integral M (f i) (x i)) \<Longrightarrow> + has_bochner_integral M (\<lambda>x. \<Sum>i\<in>I. f i x) (\<Sum>i\<in>I. x i)" + by (induct I rule: infinite_finite_induct) auto + +lemma has_bochner_integral_implies_finite_norm: + "has_bochner_integral M f x \<Longrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" +proof (elim has_bochner_integral.cases) + fix s v + assume [measurable]: "f \<in> borel_measurable M" and s: "\<And>i. simple_bochner_integrable M (s i)" and + lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" + from order_tendstoD[OF lim_0, of "\<infinity>"] + obtain i where f_s_fin: "(\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) < \<infinity>" + by (auto simp: eventually_sequentially) + + have [measurable]: "\<And>i. s i \<in> borel_measurable M" + using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + define m where "m = (if space M = {} then 0 else Max ((\<lambda>x. norm (s i x))`space M))" + have "finite (s i ` space M)" + using s by (auto simp: simple_function_def simple_bochner_integrable.simps) + then have "finite (norm ` s i ` space M)" + by (rule finite_imageI) + then have "\<And>x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> m" "0 \<le> m" + by (auto simp: m_def image_comp comp_def Max_ge_iff) + then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)" + by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:) + also have "\<dots> < \<infinity>" + using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps ennreal_mult_less_top less_top) + finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" . + + have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)" + by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric]) + (metis add.commute norm_triangle_sub) + also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)" + by (rule nn_integral_add) auto + also have "\<dots> < \<infinity>" + using s_fin f_s_fin by auto + finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" . +qed + +lemma has_bochner_integral_norm_bound: + assumes i: "has_bochner_integral M f x" + shows "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)" +using assms proof + fix s assume + x: "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" (is "?s \<longlonglongrightarrow> x") and + s[simp]: "\<And>i. simple_bochner_integrable M (s i)" and + lim: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" and + f[measurable]: "f \<in> borel_measurable M" + + have [measurable]: "\<And>i. s i \<in> borel_measurable M" + using s by (auto simp: simple_bochner_integrable.simps intro: borel_measurable_simple_function) + + show "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)" + proof (rule LIMSEQ_le) + show "(\<lambda>i. ennreal (norm (?s i))) \<longlonglongrightarrow> norm x" + using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros) + show "\<exists>N. \<forall>n\<ge>N. norm (?s n) \<le> (\<integral>\<^sup>+x. norm (f x - s n x) \<partial>M) + (\<integral>\<^sup>+x. norm (f x) \<partial>M)" + (is "\<exists>N. \<forall>n\<ge>N. _ \<le> ?t n") + proof (intro exI allI impI) + fix n + have "ennreal (norm (?s n)) \<le> simple_bochner_integral M (\<lambda>x. norm (s n x))" + by (auto intro!: simple_bochner_integral_norm_bound) + also have "\<dots> = (\<integral>\<^sup>+x. norm (s n x) \<partial>M)" + by (intro simple_bochner_integral_eq_nn_integral) + (auto intro: s simple_bochner_integrable_compose2) + also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)" + by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric]) + (metis add.commute norm_minus_commute norm_triangle_sub) + also have "\<dots> = ?t n" + by (rule nn_integral_add) auto + finally show "norm (?s n) \<le> ?t n" . + qed + have "?t \<longlonglongrightarrow> 0 + (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" + using has_bochner_integral_implies_finite_norm[OF i] + by (intro tendsto_add tendsto_const lim) + then show "?t \<longlonglongrightarrow> \<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M" + by simp + qed +qed + +lemma has_bochner_integral_eq: + "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M f y \<Longrightarrow> x = y" +proof (elim has_bochner_integral.cases) + assume f[measurable]: "f \<in> borel_measurable M" + + fix s t + assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0") + assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - t i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?T \<longlonglongrightarrow> 0") + assume s: "\<And>i. simple_bochner_integrable M (s i)" + assume t: "\<And>i. simple_bochner_integrable M (t i)" + + have [measurable]: "\<And>i. s i \<in> borel_measurable M" "\<And>i. t i \<in> borel_measurable M" + using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + let ?s = "\<lambda>i. simple_bochner_integral M (s i)" + let ?t = "\<lambda>i. simple_bochner_integral M (t i)" + assume "?s \<longlonglongrightarrow> x" "?t \<longlonglongrightarrow> y" + then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> norm (x - y)" + by (intro tendsto_intros) + moreover + have "(\<lambda>i. ennreal (norm (?s i - ?t i))) \<longlonglongrightarrow> ennreal 0" + proof (rule tendsto_sandwich) + show "eventually (\<lambda>i. 0 \<le> ennreal (norm (?s i - ?t i))) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> ennreal 0" + by auto + + show "eventually (\<lambda>i. norm (?s i - ?t i) \<le> ?S i + ?T i) sequentially" + by (intro always_eventually allI simple_bochner_integral_bounded s t f) + show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ennreal 0" + using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp + qed + then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0" + by (simp add: ennreal_0[symmetric] del: ennreal_0) + ultimately have "norm (x - y) = 0" + by (rule LIMSEQ_unique) + then show "x = y" by simp +qed + +lemma has_bochner_integralI_AE: + assumes f: "has_bochner_integral M f x" + and g: "g \<in> borel_measurable M" + and ae: "AE x in M. f x = g x" + shows "has_bochner_integral M g x" + using f +proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) + fix s assume "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" + also have "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) = (\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M)" + using ae + by (intro ext nn_integral_cong_AE, eventually_elim) simp + finally show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" . +qed (auto intro: g) + +lemma has_bochner_integral_eq_AE: + assumes f: "has_bochner_integral M f x" + and g: "has_bochner_integral M g y" + and ae: "AE x in M. f x = g x" + shows "x = y" +proof - + from assms have "has_bochner_integral M g x" + by (auto intro: has_bochner_integralI_AE) + from this g show "x = y" + by (rule has_bochner_integral_eq) +qed + +lemma simple_bochner_integrable_restrict_space: + fixes f :: "_ \<Rightarrow> 'b::real_normed_vector" + assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" + shows "simple_bochner_integrable (restrict_space M \<Omega>) f \<longleftrightarrow> + simple_bochner_integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" + by (simp add: simple_bochner_integrable.simps space_restrict_space + simple_function_restrict_space[OF \<Omega>] emeasure_restrict_space[OF \<Omega>] Collect_restrict + indicator_eq_0_iff conj_ac) + +lemma simple_bochner_integral_restrict_space: + fixes f :: "_ \<Rightarrow> 'b::real_normed_vector" + assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" + assumes f: "simple_bochner_integrable (restrict_space M \<Omega>) f" + shows "simple_bochner_integral (restrict_space M \<Omega>) f = + simple_bochner_integral M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" +proof - + have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x)`space M)" + using f simple_bochner_integrable_restrict_space[OF \<Omega>, of f] + by (simp add: simple_bochner_integrable.simps simple_function_def) + then show ?thesis + by (auto simp: space_restrict_space measure_restrict_space[OF \<Omega>(1)] le_infI2 + simple_bochner_integral_def Collect_restrict + split: split_indicator split_indicator_asm + intro!: setsum.mono_neutral_cong_left arg_cong2[where f=measure]) +qed + +context + notes [[inductive_internals]] +begin + +inductive integrable for M f where + "has_bochner_integral M f x \<Longrightarrow> integrable M f" + +end + +definition lebesgue_integral ("integral\<^sup>L") where + "integral\<^sup>L M f = (if \<exists>x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)" + +syntax + "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" ("\<integral>((2 _./ _)/ \<partial>_)" [60,61] 110) + +translations + "\<integral> x. f \<partial>M" == "CONST lebesgue_integral M (\<lambda>x. f)" + +syntax + "_ascii_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60) + +translations + "LINT x|M. f" == "CONST lebesgue_integral M (\<lambda>x. f)" + +lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \<Longrightarrow> integral\<^sup>L M f = x" + by (metis the_equality has_bochner_integral_eq lebesgue_integral_def) + +lemma has_bochner_integral_integrable: + "integrable M f \<Longrightarrow> has_bochner_integral M f (integral\<^sup>L M f)" + by (auto simp: has_bochner_integral_integral_eq integrable.simps) + +lemma has_bochner_integral_iff: + "has_bochner_integral M f x \<longleftrightarrow> integrable M f \<and> integral\<^sup>L M f = x" + by (metis has_bochner_integral_integrable has_bochner_integral_integral_eq integrable.intros) + +lemma simple_bochner_integrable_eq_integral: + "simple_bochner_integrable M f \<Longrightarrow> simple_bochner_integral M f = integral\<^sup>L M f" + using has_bochner_integral_simple_bochner_integrable[of M f] + by (simp add: has_bochner_integral_integral_eq) + +lemma not_integrable_integral_eq: "\<not> integrable M f \<Longrightarrow> integral\<^sup>L M f = 0" + unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The]) + +lemma integral_eq_cases: + "integrable M f \<longleftrightarrow> integrable N g \<Longrightarrow> + (integrable M f \<Longrightarrow> integrable N g \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g) \<Longrightarrow> + integral\<^sup>L M f = integral\<^sup>L N g" + by (metis not_integrable_integral_eq) + +lemma borel_measurable_integrable[measurable_dest]: "integrable M f \<Longrightarrow> f \<in> borel_measurable M" + by (auto elim: integrable.cases has_bochner_integral.cases) + +lemma borel_measurable_integrable'[measurable_dest]: + "integrable M f \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N" + using borel_measurable_integrable[measurable] by measurable + +lemma integrable_cong: + "M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable N g" + by (simp cong: has_bochner_integral_cong add: integrable.simps) + +lemma integrable_cong_AE: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow> + integrable M f \<longleftrightarrow> integrable M g" + unfolding integrable.simps + by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext) + +lemma integral_cong: + "M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g" + by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def) + +lemma integral_cong_AE: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow> + integral\<^sup>L M f = integral\<^sup>L M g" + unfolding lebesgue_integral_def + by (rule arg_cong[where x="has_bochner_integral M f"]) (intro has_bochner_integral_cong_AE ext) + +lemma integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)" + by (auto simp: integrable.simps) + +lemma integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)" + by (metis has_bochner_integral_zero integrable.simps) + +lemma integrable_setsum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)" + by (metis has_bochner_integral_setsum integrable.simps) + +lemma integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> + integrable M (\<lambda>x. indicator A x *\<^sub>R c)" + by (metis has_bochner_integral_indicator integrable.simps) + +lemma integrable_real_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> + integrable M (indicator A :: 'a \<Rightarrow> real)" + by (metis has_bochner_integral_real_indicator integrable.simps) + +lemma integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x - g x)" + by (auto simp: integrable.simps intro: has_bochner_integral_diff) + +lemma integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))" + by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear) + +lemma integrable_scaleR_left[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x *\<^sub>R c)" + unfolding integrable.simps by fastforce + +lemma integrable_scaleR_right[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c *\<^sub>R f x)" + unfolding integrable.simps by fastforce + +lemma integrable_mult_left[simp, intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x * c)" + unfolding integrable.simps by fastforce + +lemma integrable_mult_right[simp, intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c * f x)" + unfolding integrable.simps by fastforce + +lemma integrable_divide_zero[simp, intro]: + fixes c :: "_::{real_normed_field, field, second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x / c)" + unfolding integrable.simps by fastforce + +lemma integrable_inner_left[simp, intro]: + "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x \<bullet> c)" + unfolding integrable.simps by fastforce + +lemma integrable_inner_right[simp, intro]: + "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c \<bullet> f x)" + unfolding integrable.simps by fastforce + +lemmas integrable_minus[simp, intro] = + integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] +lemmas integrable_divide[simp, intro] = + integrable_bounded_linear[OF bounded_linear_divide] +lemmas integrable_Re[simp, intro] = + integrable_bounded_linear[OF bounded_linear_Re] +lemmas integrable_Im[simp, intro] = + integrable_bounded_linear[OF bounded_linear_Im] +lemmas integrable_cnj[simp, intro] = + integrable_bounded_linear[OF bounded_linear_cnj] +lemmas integrable_of_real[simp, intro] = + integrable_bounded_linear[OF bounded_linear_of_real] +lemmas integrable_fst[simp, intro] = + integrable_bounded_linear[OF bounded_linear_fst] +lemmas integrable_snd[simp, intro] = + integrable_bounded_linear[OF bounded_linear_snd] + +lemma integral_zero[simp]: "integral\<^sup>L M (\<lambda>x. 0) = 0" + by (intro has_bochner_integral_integral_eq has_bochner_integral_zero) + +lemma integral_add[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> + integral\<^sup>L M (\<lambda>x. f x + g x) = integral\<^sup>L M f + integral\<^sup>L M g" + by (intro has_bochner_integral_integral_eq has_bochner_integral_add has_bochner_integral_integrable) + +lemma integral_diff[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> + integral\<^sup>L M (\<lambda>x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g" + by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable) + +lemma integral_setsum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> + integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))" + by (intro has_bochner_integral_integral_eq has_bochner_integral_setsum has_bochner_integral_integrable) + +lemma integral_setsum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow> + integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))" + unfolding simp_implies_def by (rule integral_setsum) + +lemma integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> + integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)" + by (metis has_bochner_integral_bounded_linear has_bochner_integral_integrable has_bochner_integral_integral_eq) + +lemma integral_bounded_linear': + assumes T: "bounded_linear T" and T': "bounded_linear T'" + assumes *: "\<not> (\<forall>x. T x = 0) \<Longrightarrow> (\<forall>x. T' (T x) = x)" + shows "integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)" +proof cases + assume "(\<forall>x. T x = 0)" then show ?thesis + by simp +next + assume **: "\<not> (\<forall>x. T x = 0)" + show ?thesis + proof cases + assume "integrable M f" with T show ?thesis + by (rule integral_bounded_linear) + next + assume not: "\<not> integrable M f" + moreover have "\<not> integrable M (\<lambda>x. T (f x))" + proof + assume "integrable M (\<lambda>x. T (f x))" + from integrable_bounded_linear[OF T' this] not *[OF **] + show False + by auto + qed + ultimately show ?thesis + using T by (simp add: not_integrable_integral_eq linear_simps) + qed +qed + +lemma integral_scaleR_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x *\<^sub>R c \<partial>M) = integral\<^sup>L M f *\<^sub>R c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left) + +lemma integral_scaleR_right[simp]: "(\<integral> x. c *\<^sub>R f x \<partial>M) = c *\<^sub>R integral\<^sup>L M f" + by (rule integral_bounded_linear'[OF bounded_linear_scaleR_right bounded_linear_scaleR_right[of "1 / c"]]) simp + +lemma integral_mult_left[simp]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_left) + +lemma integral_mult_right[simp]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c * f x \<partial>M) = c * integral\<^sup>L M f" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_right) + +lemma integral_mult_left_zero[simp]: + fixes c :: "_::{real_normed_field,second_countable_topology}" + shows "(\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c" + by (rule integral_bounded_linear'[OF bounded_linear_mult_left bounded_linear_mult_left[of "1 / c"]]) simp + +lemma integral_mult_right_zero[simp]: + fixes c :: "_::{real_normed_field,second_countable_topology}" + shows "(\<integral> x. c * f x \<partial>M) = c * integral\<^sup>L M f" + by (rule integral_bounded_linear'[OF bounded_linear_mult_right bounded_linear_mult_right[of "1 / c"]]) simp + +lemma integral_inner_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x \<bullet> c \<partial>M) = integral\<^sup>L M f \<bullet> c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left) + +lemma integral_inner_right[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c \<bullet> f x \<partial>M) = c \<bullet> integral\<^sup>L M f" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right) + +lemma integral_divide_zero[simp]: + fixes c :: "_::{real_normed_field, field, second_countable_topology}" + shows "integral\<^sup>L M (\<lambda>x. f x / c) = integral\<^sup>L M f / c" + by (rule integral_bounded_linear'[OF bounded_linear_divide bounded_linear_mult_left[of c]]) simp + +lemma integral_minus[simp]: "integral\<^sup>L M (\<lambda>x. - f x) = - integral\<^sup>L M f" + by (rule integral_bounded_linear'[OF bounded_linear_minus[OF bounded_linear_ident] bounded_linear_minus[OF bounded_linear_ident]]) simp + +lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\<lambda>x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)" + by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp + +lemma integral_cnj[simp]: "integral\<^sup>L M (\<lambda>x. cnj (f x)) = cnj (integral\<^sup>L M f)" + by (rule integral_bounded_linear'[OF bounded_linear_cnj bounded_linear_cnj]) simp + +lemmas integral_divide[simp] = + integral_bounded_linear[OF bounded_linear_divide] +lemmas integral_Re[simp] = + integral_bounded_linear[OF bounded_linear_Re] +lemmas integral_Im[simp] = + integral_bounded_linear[OF bounded_linear_Im] +lemmas integral_of_real[simp] = + integral_bounded_linear[OF bounded_linear_of_real] +lemmas integral_fst[simp] = + integral_bounded_linear[OF bounded_linear_fst] +lemmas integral_snd[simp] = + integral_bounded_linear[OF bounded_linear_snd] + +lemma integral_norm_bound_ennreal: + "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)" + by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound) + +lemma integrableI_sequence: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes f[measurable]: "f \<in> borel_measurable M" + assumes s: "\<And>i. simple_bochner_integrable M (s i)" + assumes lim: "(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0") + shows "integrable M f" +proof - + let ?s = "\<lambda>n. simple_bochner_integral M (s n)" + + have "\<exists>x. ?s \<longlonglongrightarrow> x" + unfolding convergent_eq_cauchy + proof (rule metric_CauchyI) + fix e :: real assume "0 < e" + then have "0 < ennreal (e / 2)" by auto + from order_tendstoD(2)[OF lim this] + obtain M where M: "\<And>n. M \<le> n \<Longrightarrow> ?S n < e / 2" + by (auto simp: eventually_sequentially) + show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (?s m) (?s n) < e" + proof (intro exI allI impI) + fix m n assume m: "M \<le> m" and n: "M \<le> n" + have "?S n \<noteq> \<infinity>" + using M[OF n] by auto + have "norm (?s n - ?s m) \<le> ?S n + ?S m" + by (intro simple_bochner_integral_bounded s f) + also have "\<dots> < ennreal (e / 2) + e / 2" + by (intro add_strict_mono M n m) + also have "\<dots> = e" using \<open>0<e\<close> by (simp del: ennreal_plus add: ennreal_plus[symmetric]) + finally show "dist (?s n) (?s m) < e" + using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff) + qed + qed + then obtain x where "?s \<longlonglongrightarrow> x" .. + show ?thesis + by (rule, rule) fact+ +qed + +lemma nn_integral_dominated_convergence_norm: + fixes u' :: "_ \<Rightarrow> _::{real_normed_vector, second_countable_topology}" + assumes [measurable]: + "\<And>i. u i \<in> borel_measurable M" "u' \<in> borel_measurable M" "w \<in> borel_measurable M" + and bound: "\<And>j. AE x in M. norm (u j x) \<le> w x" + and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" + and u': "AE x in M. (\<lambda>i. u i x) \<longlonglongrightarrow> u' x" + shows "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> 0" +proof - + have "AE x in M. \<forall>j. norm (u j x) \<le> w x" + unfolding AE_all_countable by rule fact + with u' have bnd: "AE x in M. \<forall>j. norm (u' x - u j x) \<le> 2 * w x" + proof (eventually_elim, intro allI) + fix i x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x" "\<forall>j. norm (u j x) \<le> w x" "\<forall>j. norm (u j x) \<le> w x" + then have "norm (u' x) \<le> w x" "norm (u i x) \<le> w x" + by (auto intro: LIMSEQ_le_const2 tendsto_norm) + then have "norm (u' x) + norm (u i x) \<le> 2 * w x" + by simp + also have "norm (u' x - u i x) \<le> norm (u' x) + norm (u i x)" + by (rule norm_triangle_ineq4) + finally (xtrans) show "norm (u' x - u i x) \<le> 2 * w x" . + qed + have w_nonneg: "AE x in M. 0 \<le> w x" + using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero]) + + have "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. 0 \<partial>M)" + proof (rule nn_integral_dominated_convergence) + show "(\<integral>\<^sup>+x. 2 * w x \<partial>M) < \<infinity>" + by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult ) + show "AE x in M. (\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0" + using u' + proof eventually_elim + fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x" + from tendsto_diff[OF tendsto_const[of "u' x"] this] + show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0" + by (simp add: tendsto_norm_zero_iff ennreal_0[symmetric] del: ennreal_0) + qed + qed (insert bnd w_nonneg, auto) + then show ?thesis by simp +qed + +lemma integrableI_bounded: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes f[measurable]: "f \<in> borel_measurable M" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" + shows "integrable M f" +proof - + from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where + s: "\<And>i. simple_function M (s i)" and + pointwise: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" and + bound: "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)" + by simp metis + + show ?thesis + proof (rule integrableI_sequence) + { fix i + have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)" + by (intro nn_integral_mono) (simp add: bound) + also have "\<dots> = 2 * (\<integral>\<^sup>+x. ennreal (norm (f x)) \<partial>M)" + by (simp add: ennreal_mult nn_integral_cmult) + also have "\<dots> < top" + using fin by (simp add: ennreal_mult_less_top) + finally have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" + by simp } + note fin_s = this + + show "\<And>i. simple_bochner_integrable M (s i)" + by (rule simple_bochner_integrableI_bounded) fact+ + + show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" + proof (rule nn_integral_dominated_convergence_norm) + show "\<And>j. AE x in M. norm (s j x) \<le> 2 * norm (f x)" + using bound by auto + show "\<And>i. s i \<in> borel_measurable M" "(\<lambda>x. 2 * norm (f x)) \<in> borel_measurable M" + using s by (auto intro: borel_measurable_simple_function) + show "(\<integral>\<^sup>+ x. ennreal (2 * norm (f x)) \<partial>M) < \<infinity>" + using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top) + show "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x" + using pointwise by auto + qed fact + qed fact +qed + +lemma integrableI_bounded_set: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M" + assumes finite: "emeasure M A < \<infinity>" + and bnd: "AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B" + and null: "AE x in M. x \<notin> A \<longrightarrow> f x = 0" + shows "integrable M f" +proof (rule integrableI_bounded) + { fix x :: 'b have "norm x \<le> B \<Longrightarrow> 0 \<le> B" + using norm_ge_zero[of x] by arith } + with bnd null have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (max 0 B) * indicator A x \<partial>M)" + by (intro nn_integral_mono_AE) (auto split: split_indicator split_max) + also have "\<dots> < \<infinity>" + using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top) + finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" . +qed simp + +lemma integrableI_bounded_set_indicator: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "A \<in> sets M \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> + emeasure M A < \<infinity> \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B) \<Longrightarrow> + integrable M (\<lambda>x. indicator A x *\<^sub>R f x)" + by (rule integrableI_bounded_set[where A=A]) auto + +lemma integrableI_nonneg: + fixes f :: "'a \<Rightarrow> real" + assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" + shows "integrable M f" +proof - + have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)" + using assms by (intro nn_integral_cong_AE) auto + then show ?thesis + using assms by (intro integrableI_bounded) auto +qed + +lemma integrable_iff_bounded: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" + using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f] + unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto + +lemma integrable_bound: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + and g :: "'a \<Rightarrow> 'c::{banach, second_countable_topology}" + shows "integrable M f \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. norm (g x) \<le> norm (f x)) \<Longrightarrow> + integrable M g" + unfolding integrable_iff_bounded +proof safe + assume "f \<in> borel_measurable M" "g \<in> borel_measurable M" + assume "AE x in M. norm (g x) \<le> norm (f x)" + then have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" + by (intro nn_integral_mono_AE) auto + also assume "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" + finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" . +qed + +lemma integrable_mult_indicator: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)" + by (rule integrable_bound[of M f]) (auto split: split_indicator) + +lemma integrable_real_mult_indicator: + fixes f :: "'a \<Rightarrow> real" + shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. f x * indicator A x)" + using integrable_mult_indicator[of A M f] by (simp add: mult_ac) + +lemma integrable_abs[simp, intro]: + fixes f :: "'a \<Rightarrow> real" + assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. \<bar>f x\<bar>)" + using assms by (rule integrable_bound) auto + +lemma integrable_norm[simp, intro]: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. norm (f x))" + using assms by (rule integrable_bound) auto + +lemma integrable_norm_cancel: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes [measurable]: "integrable M (\<lambda>x. norm (f x))" "f \<in> borel_measurable M" shows "integrable M f" + using assms by (rule integrable_bound) auto + +lemma integrable_norm_iff: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. norm (f x)) \<longleftrightarrow> integrable M f" + by (auto intro: integrable_norm_cancel) + +lemma integrable_abs_cancel: + fixes f :: "'a \<Rightarrow> real" + assumes [measurable]: "integrable M (\<lambda>x. \<bar>f x\<bar>)" "f \<in> borel_measurable M" shows "integrable M f" + using assms by (rule integrable_bound) auto + +lemma integrable_abs_iff: + fixes f :: "'a \<Rightarrow> real" + shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. \<bar>f x\<bar>) \<longleftrightarrow> integrable M f" + by (auto intro: integrable_abs_cancel) + +lemma integrable_max[simp, intro]: + fixes f :: "'a \<Rightarrow> real" + assumes fg[measurable]: "integrable M f" "integrable M g" + shows "integrable M (\<lambda>x. max (f x) (g x))" + using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] + by (rule integrable_bound) auto + +lemma integrable_min[simp, intro]: + fixes f :: "'a \<Rightarrow> real" + assumes fg[measurable]: "integrable M f" "integrable M g" + shows "integrable M (\<lambda>x. min (f x) (g x))" + using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] + by (rule integrable_bound) auto + +lemma integral_minus_iff[simp]: + "integrable M (\<lambda>x. - f x ::'a::{banach, second_countable_topology}) \<longleftrightarrow> integrable M f" + unfolding integrable_iff_bounded + by (auto intro: borel_measurable_uminus[of "\<lambda>x. - f x" M, simplified]) + +lemma integrable_indicator_iff: + "integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>" + by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator' + cong: conj_cong) + +lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \<inter> space M)" +proof cases + assume *: "A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>" + have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M))" + by (intro integral_cong) (auto split: split_indicator) + also have "\<dots> = measure M (A \<inter> space M)" + using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto + finally show ?thesis . +next + assume *: "\<not> (A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>)" + have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M) :: _ \<Rightarrow> real)" + by (intro integral_cong) (auto split: split_indicator) + also have "\<dots> = 0" + using * by (subst not_integrable_integral_eq) (auto simp: integrable_indicator_iff) + also have "\<dots> = measure M (A \<inter> space M)" + using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique) + finally show ?thesis . +qed + +lemma integrable_discrete_difference: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes X: "countable X" + assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" + assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" + assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x" + shows "integrable M f \<longleftrightarrow> integrable M g" + unfolding integrable_iff_bounded +proof (rule conj_cong) + { assume "f \<in> borel_measurable M" then have "g \<in> borel_measurable M" + by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } + moreover + { assume "g \<in> borel_measurable M" then have "f \<in> borel_measurable M" + by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } + ultimately show "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M" .. +next + have "AE x in M. x \<notin> X" + by (rule AE_discrete_difference) fact+ + then have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. norm (g x) \<partial>M)" + by (intro nn_integral_cong_AE) (auto simp: eq) + then show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity> \<longleftrightarrow> (\<integral>\<^sup>+ x. norm (g x) \<partial>M) < \<infinity>" + by simp +qed + +lemma integral_discrete_difference: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes X: "countable X" + assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" + assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" + assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x" + shows "integral\<^sup>L M f = integral\<^sup>L M g" +proof (rule integral_eq_cases) + show eq: "integrable M f \<longleftrightarrow> integrable M g" + by (rule integrable_discrete_difference[where X=X]) fact+ + + assume f: "integrable M f" + show "integral\<^sup>L M f = integral\<^sup>L M g" + proof (rule integral_cong_AE) + show "f \<in> borel_measurable M" "g \<in> borel_measurable M" + using f eq by (auto intro: borel_measurable_integrable) + + have "AE x in M. x \<notin> X" + by (rule AE_discrete_difference) fact+ + with AE_space show "AE x in M. f x = g x" + by eventually_elim fact + qed +qed + +lemma has_bochner_integral_discrete_difference: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes X: "countable X" + assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" + assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" + assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x" + shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x" + using integrable_discrete_difference[of X M f g, OF assms] + using integral_discrete_difference[of X M f g, OF assms] + by (metis has_bochner_integral_iff) + +lemma + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real" + assumes "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M" "integrable M w" + assumes lim: "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x" + assumes bound: "\<And>i. AE x in M. norm (s i x) \<le> w x" + shows integrable_dominated_convergence: "integrable M f" + and integrable_dominated_convergence2: "\<And>i. integrable M (s i)" + and integral_dominated_convergence: "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f" +proof - + have w_nonneg: "AE x in M. 0 \<le> w x" + using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans) + then have "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)" + by (intro nn_integral_cong_AE) auto + with \<open>integrable M w\<close> have w: "w \<in> borel_measurable M" "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" + unfolding integrable_iff_bounded by auto + + show int_s: "\<And>i. integrable M (s i)" + unfolding integrable_iff_bounded + proof + fix i + have "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)" + using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto + with w show "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) < \<infinity>" by auto + qed fact + + have all_bound: "AE x in M. \<forall>i. norm (s i x) \<le> w x" + using bound unfolding AE_all_countable by auto + + show int_f: "integrable M f" + unfolding integrable_iff_bounded + proof + have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)" + using all_bound lim w_nonneg + proof (intro nn_integral_mono_AE, eventually_elim) + fix x assume "\<forall>i. norm (s i x) \<le> w x" "(\<lambda>i. s i x) \<longlonglongrightarrow> f x" "0 \<le> w x" + then show "ennreal (norm (f x)) \<le> ennreal (w x)" + by (intro LIMSEQ_le_const2[where X="\<lambda>i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros) + qed + with w show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" by auto + qed fact + + have "(\<lambda>n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \<longlonglongrightarrow> ennreal 0" (is "?d \<longlonglongrightarrow> ennreal 0") + proof (rule tendsto_sandwich) + show "eventually (\<lambda>n. ennreal 0 \<le> ?d n) sequentially" "(\<lambda>_. ennreal 0) \<longlonglongrightarrow> ennreal 0" by auto + show "eventually (\<lambda>n. ?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)) sequentially" + proof (intro always_eventually allI) + fix n + have "?d n = norm (integral\<^sup>L M (\<lambda>x. s n x - f x))" + using int_f int_s by simp + also have "\<dots> \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)" + by (intro int_f int_s integrable_diff integral_norm_bound_ennreal) + finally show "?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)" . + qed + show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) \<longlonglongrightarrow> ennreal 0" + unfolding ennreal_0 + apply (subst norm_minus_commute) + proof (rule nn_integral_dominated_convergence_norm[where w=w]) + show "\<And>n. s n \<in> borel_measurable M" + using int_s unfolding integrable_iff_bounded by auto + qed fact+ + qed + then have "(\<lambda>n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \<longlonglongrightarrow> 0" + by (simp add: tendsto_norm_zero_iff del: ennreal_0) + from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]] + show "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f" by simp +qed + +context + fixes s :: "real \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real" + and f :: "'a \<Rightarrow> 'b" and M + assumes "f \<in> borel_measurable M" "\<And>t. s t \<in> borel_measurable M" "integrable M w" + assumes lim: "AE x in M. ((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top" + assumes bound: "\<forall>\<^sub>F i in at_top. AE x in M. norm (s i x) \<le> w x" +begin + +lemma integral_dominated_convergence_at_top: "((\<lambda>t. integral\<^sup>L M (s t)) \<longlongrightarrow> integral\<^sup>L M f) at_top" +proof (rule tendsto_at_topI_sequentially) + fix X :: "nat \<Rightarrow> real" assume X: "filterlim X at_top sequentially" + from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] + obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s (X n) x) \<le> w x" + by (auto simp: eventually_sequentially) + + show "(\<lambda>n. integral\<^sup>L M (s (X n))) \<longlonglongrightarrow> integral\<^sup>L M f" + proof (rule LIMSEQ_offset, rule integral_dominated_convergence) + show "AE x in M. norm (s (X (n + N)) x) \<le> w x" for n + by (rule w) auto + show "AE x in M. (\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x" + using lim + proof eventually_elim + fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top" + then show "(\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x" + by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) + qed + qed fact+ +qed + +lemma integrable_dominated_convergence_at_top: "integrable M f" +proof - + from bound obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s n x) \<le> w x" + by (auto simp: eventually_at_top_linorder) + show ?thesis + proof (rule integrable_dominated_convergence) + show "AE x in M. norm (s (N + i) x) \<le> w x" for i :: nat + by (intro w) auto + show "AE x in M. (\<lambda>i. s (N + real i) x) \<longlonglongrightarrow> f x" + using lim + proof eventually_elim + fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top" + then show "(\<lambda>n. s (N + n) x) \<longlonglongrightarrow> f x" + by (rule filterlim_compose) + (auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially) + qed + qed fact+ +qed + +end + +lemma integrable_mult_left_iff: + fixes f :: "'a \<Rightarrow> real" + shows "integrable M (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> integrable M f" + using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\<lambda>x. c * f x"] + by (cases "c = 0") auto + +lemma integrableI_nn_integral_finite: + assumes [measurable]: "f \<in> borel_measurable M" + and nonneg: "AE x in M. 0 \<le> f x" + and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x" + shows "integrable M f" +proof (rule integrableI_bounded) + have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)" + using nonneg by (intro nn_integral_cong_AE) auto + with finite show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" + by auto +qed simp + +lemma integral_nonneg_AE: + fixes f :: "'a \<Rightarrow> real" + assumes nonneg: "AE x in M. 0 \<le> f x" + shows "0 \<le> integral\<^sup>L M f" +proof cases + assume f: "integrable M f" + then have [measurable]: "f \<in> M \<rightarrow>\<^sub>M borel" + by auto + have "(\<lambda>x. max 0 (f x)) \<in> M \<rightarrow>\<^sub>M borel" "\<And>x. 0 \<le> max 0 (f x)" "integrable M (\<lambda>x. max 0 (f x))" + using f by auto + from this have "0 \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))" + proof (induction rule: borel_measurable_induct_real) + case (add f g) + then have "integrable M f" "integrable M g" + by (auto intro!: integrable_bound[OF add.prems]) + with add show ?case + by (simp add: nn_integral_add) + next + case (seq U) + show ?case + proof (rule LIMSEQ_le_const) + have U_le: "x \<in> space M \<Longrightarrow> U i x \<le> max 0 (f x)" for x i + using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) + with seq nonneg show "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> LINT x|M. max 0 (f x)" + by (intro integral_dominated_convergence) auto + have "integrable M (U i)" for i + using seq.prems by (rule integrable_bound) (insert U_le seq, auto) + with seq show "\<exists>N. \<forall>n\<ge>N. 0 \<le> integral\<^sup>L M (U n)" + by auto + qed + qed (auto simp: measure_nonneg integrable_mult_left_iff) + also have "\<dots> = integral\<^sup>L M f" + using nonneg by (auto intro!: integral_cong_AE) + finally show ?thesis . +qed (simp add: not_integrable_integral_eq) + +lemma integral_nonneg[simp]: + fixes f :: "'a \<Rightarrow> real" + shows "(\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> 0 \<le> integral\<^sup>L M f" + by (intro integral_nonneg_AE) auto + +lemma nn_integral_eq_integral: + assumes f: "integrable M f" + assumes nonneg: "AE x in M. 0 \<le> f x" + shows "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f" +proof - + { fix f :: "'a \<Rightarrow> real" assume f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "integrable M f" + then have "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f" + proof (induct rule: borel_measurable_induct_real) + case (set A) then show ?case + by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure) + next + case (mult f c) then show ?case + by (auto simp add: integrable_mult_left_iff nn_integral_cmult ennreal_mult integral_nonneg_AE) + next + case (add g f) + then have "integrable M f" "integrable M g" + by (auto intro!: integrable_bound[OF add.prems]) + with add show ?case + by (simp add: nn_integral_add integral_nonneg_AE) + next + case (seq U) + show ?case + proof (rule LIMSEQ_unique) + have U_le_f: "x \<in> space M \<Longrightarrow> U i x \<le> f x" for x i + using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) + have int_U: "\<And>i. integrable M (U i)" + using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto + from U_le_f seq have "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> integral\<^sup>L M f" + by (intro integral_dominated_convergence) auto + then show "(\<lambda>i. ennreal (integral\<^sup>L M (U i))) \<longlonglongrightarrow> ennreal (integral\<^sup>L M f)" + using seq f int_U by (simp add: f integral_nonneg_AE) + have "(\<lambda>i. \<integral>\<^sup>+ x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+ x. f x \<partial>M" + using seq U_le_f f + by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded) + then show "(\<lambda>i. \<integral>x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+x. f x \<partial>M" + using seq int_U by simp + qed + qed } + from this[of "\<lambda>x. max 0 (f x)"] assms have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = integral\<^sup>L M (\<lambda>x. max 0 (f x))" + by simp + also have "\<dots> = integral\<^sup>L M f" + using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE) + also have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)" + using assms by (auto intro!: nn_integral_cong_AE simp: max_def) + finally show ?thesis . +qed + +lemma + fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {banach, second_countable_topology}" + assumes integrable[measurable]: "\<And>i. integrable M (f i)" + and summable: "AE x in M. summable (\<lambda>i. norm (f i x))" + and sums: "summable (\<lambda>i. (\<integral>x. norm (f i x) \<partial>M))" + shows integrable_suminf: "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S") + and sums_integral: "(\<lambda>i. integral\<^sup>L M (f i)) sums (\<integral>x. (\<Sum>i. f i x) \<partial>M)" (is "?f sums ?x") + and integral_suminf: "(\<integral>x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>L M (f i))" + and summable_integral: "summable (\<lambda>i. integral\<^sup>L M (f i))" +proof - + have 1: "integrable M (\<lambda>x. \<Sum>i. norm (f i x))" + proof (rule integrableI_bounded) + have "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ennreal (norm (f i x))) \<partial>M)" + apply (intro nn_integral_cong_AE) + using summable + apply eventually_elim + apply (simp add: suminf_nonneg ennreal_suminf_neq_top) + done + also have "\<dots> = (\<Sum>i. \<integral>\<^sup>+ x. norm (f i x) \<partial>M)" + by (intro nn_integral_suminf) auto + also have "\<dots> = (\<Sum>i. ennreal (\<integral>x. norm (f i x) \<partial>M))" + by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto + finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) < \<infinity>" + by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE) + qed simp + + have 2: "AE x in M. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> (\<Sum>i. f i x)" + using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel) + + have 3: "\<And>j. AE x in M. norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" + using summable + proof eventually_elim + fix j x assume [simp]: "summable (\<lambda>i. norm (f i x))" + have "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i<j. norm (f i x))" by (rule norm_setsum) + also have "\<dots> \<le> (\<Sum>i. norm (f i x))" + using setsum_le_suminf[of "\<lambda>i. norm (f i x)"] unfolding sums_iff by auto + finally show "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" by simp + qed + + note ibl = integrable_dominated_convergence[OF _ _ 1 2 3] + note int = integral_dominated_convergence[OF _ _ 1 2 3] + + show "integrable M ?S" + by (rule ibl) measurable + + show "?f sums ?x" unfolding sums_def + using int by (simp add: integrable) + then show "?x = suminf ?f" "summable ?f" + unfolding sums_iff by auto +qed + +lemma integral_norm_bound: + fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}" + shows "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)" + using nn_integral_eq_integral[of M "\<lambda>x. norm (f x)"] + using integral_norm_bound_ennreal[of M f] by (simp add: integral_nonneg_AE) + +lemma integral_eq_nn_integral: + assumes [measurable]: "f \<in> borel_measurable M" + assumes nonneg: "AE x in M. 0 \<le> f x" + shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)" +proof cases + assume *: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = \<infinity>" + also have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" + using nonneg by (intro nn_integral_cong_AE) auto + finally have "\<not> integrable M f" + by (auto simp: integrable_iff_bounded) + then show ?thesis + by (simp add: * not_integrable_integral_eq) +next + assume "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" + then have "integrable M f" + by (cases "\<integral>\<^sup>+ x. ennreal (f x) \<partial>M" rule: ennreal_cases) + (auto intro!: integrableI_nn_integral_finite assms) + from nn_integral_eq_integral[OF this] nonneg show ?thesis + by (simp add: integral_nonneg_AE) +qed + +lemma enn2real_nn_integral_eq_integral: + assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x" + and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < top" + and [measurable]: "g \<in> M \<rightarrow>\<^sub>M borel" + shows "enn2real (\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>x. g x \<partial>M)" +proof - + have "ennreal (enn2real (\<integral>\<^sup>+x. f x \<partial>M)) = (\<integral>\<^sup>+x. f x \<partial>M)" + using fin by (intro ennreal_enn2real) auto + also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M)" + using eq by (rule nn_integral_cong_AE) + also have "\<dots> = (\<integral>x. g x \<partial>M)" + proof (rule nn_integral_eq_integral) + show "integrable M g" + proof (rule integrableI_bounded) + have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)" + using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2) + also note fin + finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" + by simp + qed simp + qed fact + finally show ?thesis + using nn by (simp add: integral_nonneg_AE) +qed + +lemma has_bochner_integral_nn_integral: + assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> x" + assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x" + shows "has_bochner_integral M f x" + unfolding has_bochner_integral_iff + using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite) + +lemma integrableI_simple_bochner_integrable: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "simple_bochner_integrable M f \<Longrightarrow> integrable M f" + by (intro integrableI_sequence[where s="\<lambda>_. f"] borel_measurable_simple_function) + (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps) + +lemma integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes "integrable M f" + assumes base: "\<And>A c. A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> P (\<lambda>x. indicator A x *\<^sub>R c)" + assumes add: "\<And>f g. integrable M f \<Longrightarrow> P f \<Longrightarrow> integrable M g \<Longrightarrow> P g \<Longrightarrow> P (\<lambda>x. f x + g x)" + assumes lim: "\<And>f s. (\<And>i. integrable M (s i)) \<Longrightarrow> (\<And>i. P (s i)) \<Longrightarrow> + (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x) \<Longrightarrow> + (\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)) \<Longrightarrow> integrable M f \<Longrightarrow> P f" + shows "P f" +proof - + from \<open>integrable M f\<close> have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" + unfolding integrable_iff_bounded by auto + from borel_measurable_implies_sequence_metric[OF f(1)] + obtain s where s: "\<And>i. simple_function M (s i)" "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" + "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)" + unfolding norm_conv_dist by metis + + { fix f A + have [simp]: "P (\<lambda>x. 0)" + using base[of "{}" undefined] by simp + have "(\<And>i::'b. i \<in> A \<Longrightarrow> integrable M (f i::'a \<Rightarrow> 'b)) \<Longrightarrow> + (\<And>i. i \<in> A \<Longrightarrow> P (f i)) \<Longrightarrow> P (\<lambda>x. \<Sum>i\<in>A. f i x)" + by (induct A rule: infinite_finite_induct) (auto intro!: add) } + note setsum = this + + define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z + then have s'_eq_s: "\<And>i x. x \<in> space M \<Longrightarrow> s' i x = s i x" + by simp + + have sf[measurable]: "\<And>i. simple_function M (s' i)" + unfolding s'_def using s(1) + by (intro simple_function_compose2[where h="op *\<^sub>R"] simple_function_indicator) auto + + { fix i + have "\<And>z. {y. s' i z = y \<and> y \<in> s' i ` space M \<and> y \<noteq> 0 \<and> z \<in> space M} = + (if z \<in> space M \<and> s' i z \<noteq> 0 then {s' i z} else {})" + by (auto simp add: s'_def split: split_indicator) + then have "\<And>z. s' i = (\<lambda>z. \<Sum>y\<in>s' i`space M - {0}. indicator {x\<in>space M. s' i x = y} z *\<^sub>R y)" + using sf by (auto simp: fun_eq_iff simple_function_def s'_def) } + note s'_eq = this + + show "P f" + proof (rule lim) + fix i + + have "(\<integral>\<^sup>+x. norm (s' i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)" + using s by (intro nn_integral_mono) (auto simp: s'_eq_s) + also have "\<dots> < \<infinity>" + using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult) + finally have sbi: "simple_bochner_integrable M (s' i)" + using sf by (intro simple_bochner_integrableI_bounded) auto + then show "integrable M (s' i)" + by (rule integrableI_simple_bochner_integrable) + + { fix x assume"x \<in> space M" "s' i x \<noteq> 0" + then have "emeasure M {y \<in> space M. s' i y = s' i x} \<le> emeasure M {y \<in> space M. s' i y \<noteq> 0}" + by (intro emeasure_mono) auto + also have "\<dots> < \<infinity>" + using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top) + finally have "emeasure M {y \<in> space M. s' i y = s' i x} \<noteq> \<infinity>" by simp } + then show "P (s' i)" + by (subst s'_eq) (auto intro!: setsum base simp: less_top) + + fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i x) \<longlonglongrightarrow> f x" + by (simp add: s'_eq_s) + show "norm (s' i x) \<le> 2 * norm (f x)" + using \<open>x \<in> space M\<close> s by (simp add: s'_eq_s) + qed fact +qed + +lemma integral_eq_zero_AE: + "(AE x in M. f x = 0) \<Longrightarrow> integral\<^sup>L M f = 0" + using integral_cong_AE[of f M "\<lambda>_. 0"] + by (cases "integrable M f") (simp_all add: not_integrable_integral_eq) + +lemma integral_nonneg_eq_0_iff_AE: + fixes f :: "_ \<Rightarrow> real" + assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \<le> f x" + shows "integral\<^sup>L M f = 0 \<longleftrightarrow> (AE x in M. f x = 0)" +proof + assume "integral\<^sup>L M f = 0" + then have "integral\<^sup>N M f = 0" + using nn_integral_eq_integral[OF f nonneg] by simp + then have "AE x in M. ennreal (f x) \<le> 0" + by (simp add: nn_integral_0_iff_AE) + with nonneg show "AE x in M. f x = 0" + by auto +qed (auto simp add: integral_eq_zero_AE) + +lemma integral_mono_AE: + fixes f :: "'a \<Rightarrow> real" + assumes "integrable M f" "integrable M g" "AE x in M. f x \<le> g x" + shows "integral\<^sup>L M f \<le> integral\<^sup>L M g" +proof - + have "0 \<le> integral\<^sup>L M (\<lambda>x. g x - f x)" + using assms by (intro integral_nonneg_AE integrable_diff assms) auto + also have "\<dots> = integral\<^sup>L M g - integral\<^sup>L M f" + by (intro integral_diff assms) + finally show ?thesis by simp +qed + +lemma integral_mono: + fixes f :: "'a \<Rightarrow> real" + shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x) \<Longrightarrow> + integral\<^sup>L M f \<le> integral\<^sup>L M g" + by (intro integral_mono_AE) auto + +lemma (in finite_measure) integrable_measure: + assumes I: "disjoint_family_on X I" "countable I" + shows "integrable (count_space I) (\<lambda>i. measure M (X i))" +proof - + have "(\<integral>\<^sup>+i. measure M (X i) \<partial>count_space I) = (\<integral>\<^sup>+i. measure M (if X i \<in> sets M then X i else {}) \<partial>count_space I)" + by (auto intro!: nn_integral_cong measure_notin_sets) + also have "\<dots> = measure M (\<Union>i\<in>I. if X i \<in> sets M then X i else {})" + using I unfolding emeasure_eq_measure[symmetric] + by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def) + finally show ?thesis + by (auto intro!: integrableI_bounded) +qed + +lemma integrableI_real_bounded: + assumes f: "f \<in> borel_measurable M" and ae: "AE x in M. 0 \<le> f x" and fin: "integral\<^sup>N M f < \<infinity>" + shows "integrable M f" +proof (rule integrableI_bounded) + have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>M" + using ae by (auto intro: nn_integral_cong_AE) + also note fin + finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" . +qed fact + +lemma integral_real_bounded: + assumes "0 \<le> r" "integral\<^sup>N M f \<le> ennreal r" + shows "integral\<^sup>L M f \<le> r" +proof cases + assume [simp]: "integrable M f" + + have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (f x))" + by (intro nn_integral_eq_integral[symmetric]) auto + also have "\<dots> = integral\<^sup>N M f" + by (intro nn_integral_cong) (simp add: max_def ennreal_neg) + also have "\<dots> \<le> r" + by fact + finally have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) \<le> r" + using \<open>0 \<le> r\<close> by simp + + moreover have "integral\<^sup>L M f \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))" + by (rule integral_mono_AE) auto + ultimately show ?thesis + by simp +next + assume "\<not> integrable M f" then show ?thesis + using \<open>0 \<le> r\<close> by (simp add: not_integrable_integral_eq) +qed + +subsection \<open>Restricted measure spaces\<close> + +lemma integrable_restrict_space: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M" + shows "integrable (restrict_space M \<Omega>) f \<longleftrightarrow> integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" + unfolding integrable_iff_bounded + borel_measurable_restrict_space_iff[OF \<Omega>] + nn_integral_restrict_space[OF \<Omega>] + by (simp add: ac_simps ennreal_indicator ennreal_mult) + +lemma integral_restrict_space: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M" + shows "integral\<^sup>L (restrict_space M \<Omega>) f = integral\<^sup>L M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" +proof (rule integral_eq_cases) + assume "integrable (restrict_space M \<Omega>) f" + then show ?thesis + proof induct + case (base A c) then show ?case + by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff + emeasure_restrict_space Int_absorb1 measure_restrict_space) + next + case (add g f) then show ?case + by (simp add: scaleR_add_right integrable_restrict_space) + next + case (lim f s) + show ?case + proof (rule LIMSEQ_unique) + show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> integral\<^sup>L (restrict_space M \<Omega>) f" + using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) simp_all + + show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> (\<integral> x. indicator \<Omega> x *\<^sub>R f x \<partial>M)" + unfolding lim + using lim + by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (indicator \<Omega> x *\<^sub>R f x)"]) + (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR + split: split_indicator) + qed + qed +qed (simp add: integrable_restrict_space) + +lemma integral_empty: + assumes "space M = {}" + shows "integral\<^sup>L M f = 0" +proof - + have "(\<integral> x. f x \<partial>M) = (\<integral> x. 0 \<partial>M)" + by(rule integral_cong)(simp_all add: assms) + thus ?thesis by simp +qed + +subsection \<open>Measure spaces with an associated density\<close> + +lemma integrable_density: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real" + assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" + and nn: "AE x in M. 0 \<le> g x" + shows "integrable (density M g) f \<longleftrightarrow> integrable M (\<lambda>x. g x *\<^sub>R f x)" + unfolding integrable_iff_bounded using nn + apply (simp add: nn_integral_density less_top[symmetric]) + apply (intro arg_cong2[where f="op ="] refl nn_integral_cong_AE) + apply (auto simp: ennreal_mult) + done + +lemma integral_density: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real" + assumes f: "f \<in> borel_measurable M" + and g[measurable]: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" + shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)" +proof (rule integral_eq_cases) + assume "integrable (density M g) f" + then show ?thesis + proof induct + case (base A c) + then have [measurable]: "A \<in> sets M" by auto + + have int: "integrable M (\<lambda>x. g x * indicator A x)" + using g base integrable_density[of "indicator A :: 'a \<Rightarrow> real" M g] by simp + then have "integral\<^sup>L M (\<lambda>x. g x * indicator A x) = (\<integral>\<^sup>+ x. ennreal (g x * indicator A x) \<partial>M)" + using g by (subst nn_integral_eq_integral) auto + also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (g x) * indicator A x \<partial>M)" + by (intro nn_integral_cong) (auto split: split_indicator) + also have "\<dots> = emeasure (density M g) A" + by (rule emeasure_density[symmetric]) auto + also have "\<dots> = ennreal (measure (density M g) A)" + using base by (auto intro: emeasure_eq_ennreal_measure) + also have "\<dots> = integral\<^sup>L (density M g) (indicator A)" + using base by simp + finally show ?case + using base g + apply (simp add: int integral_nonneg_AE) + apply (subst (asm) ennreal_inj) + apply (auto intro!: integral_nonneg_AE) + done + next + case (add f h) + then have [measurable]: "f \<in> borel_measurable M" "h \<in> borel_measurable M" + by (auto dest!: borel_measurable_integrable) + from add g show ?case + by (simp add: scaleR_add_right integrable_density) + next + case (lim f s) + have [measurable]: "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M" + using lim(1,5)[THEN borel_measurable_integrable] by auto + + show ?case + proof (rule LIMSEQ_unique) + show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)" + proof (rule integral_dominated_convergence) + show "integrable M (\<lambda>x. 2 * norm (g x *\<^sub>R f x))" + by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto + show "AE x in M. (\<lambda>i. g x *\<^sub>R s i x) \<longlonglongrightarrow> g x *\<^sub>R f x" + using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M]) + show "\<And>i. AE x in M. norm (g x *\<^sub>R s i x) \<le> 2 * norm (g x *\<^sub>R f x)" + using lim(4) g by (auto intro!: AE_I2[of M] mult_left_mono simp: field_simps) + qed auto + show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L (density M g) f" + unfolding lim(2)[symmetric] + by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) + (insert lim(3-5), auto) + qed + qed +qed (simp add: f g integrable_density) + +lemma + fixes g :: "'a \<Rightarrow> real" + assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "g \<in> borel_measurable M" + shows integral_real_density: "integral\<^sup>L (density M f) g = (\<integral> x. f x * g x \<partial>M)" + and integrable_real_density: "integrable (density M f) g \<longleftrightarrow> integrable M (\<lambda>x. f x * g x)" + using assms integral_density[of g M f] integrable_density[of g M f] by auto + +lemma has_bochner_integral_density: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real" + shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. 0 \<le> g x) \<Longrightarrow> + has_bochner_integral M (\<lambda>x. g x *\<^sub>R f x) x \<Longrightarrow> has_bochner_integral (density M g) f x" + by (simp add: has_bochner_integral_iff integrable_density integral_density) + +subsection \<open>Distributions\<close> + +lemma integrable_distr_eq: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes [measurable]: "g \<in> measurable M N" "f \<in> borel_measurable N" + shows "integrable (distr M N g) f \<longleftrightarrow> integrable M (\<lambda>x. f (g x))" + unfolding integrable_iff_bounded by (simp_all add: nn_integral_distr) + +lemma integrable_distr: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "T \<in> measurable M M' \<Longrightarrow> integrable (distr M M' T) f \<Longrightarrow> integrable M (\<lambda>x. f (T x))" + by (subst integrable_distr_eq[symmetric, where g=T]) + (auto dest: borel_measurable_integrable) + +lemma integral_distr: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes g[measurable]: "g \<in> measurable M N" and f: "f \<in> borel_measurable N" + shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\<lambda>x. f (g x))" +proof (rule integral_eq_cases) + assume "integrable (distr M N g) f" + then show ?thesis + proof induct + case (base A c) + then have [measurable]: "A \<in> sets N" by auto + from base have int: "integrable (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c)" + by (intro integrable_indicator) + + have "integral\<^sup>L (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c" + using base by auto + also have "\<dots> = measure M (g -` A \<inter> space M) *\<^sub>R c" + by (subst measure_distr) auto + also have "\<dots> = integral\<^sup>L M (\<lambda>a. indicator (g -` A \<inter> space M) a *\<^sub>R c)" + using base by (auto simp: emeasure_distr) + also have "\<dots> = integral\<^sup>L M (\<lambda>a. indicator A (g a) *\<^sub>R c)" + using int base by (intro integral_cong_AE) (auto simp: emeasure_distr split: split_indicator) + finally show ?case . + next + case (add f h) + then have [measurable]: "f \<in> borel_measurable N" "h \<in> borel_measurable N" + by (auto dest!: borel_measurable_integrable) + from add g show ?case + by (simp add: scaleR_add_right integrable_distr_eq) + next + case (lim f s) + have [measurable]: "f \<in> borel_measurable N" "\<And>i. s i \<in> borel_measurable N" + using lim(1,5)[THEN borel_measurable_integrable] by auto + + show ?case + proof (rule LIMSEQ_unique) + show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. f (g x))" + proof (rule integral_dominated_convergence) + show "integrable M (\<lambda>x. 2 * norm (f (g x)))" + using lim by (auto simp: integrable_distr_eq) + show "AE x in M. (\<lambda>i. s i (g x)) \<longlonglongrightarrow> f (g x)" + using lim(3) g[THEN measurable_space] by auto + show "\<And>i. AE x in M. norm (s i (g x)) \<le> 2 * norm (f (g x))" + using lim(4) g[THEN measurable_space] by auto + qed auto + show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L (distr M N g) f" + unfolding lim(2)[symmetric] + by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) + (insert lim(3-5), auto) + qed + qed +qed (simp add: f g integrable_distr_eq) + +lemma has_bochner_integral_distr: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "f \<in> borel_measurable N \<Longrightarrow> g \<in> measurable M N \<Longrightarrow> + has_bochner_integral M (\<lambda>x. f (g x)) x \<Longrightarrow> has_bochner_integral (distr M N g) f x" + by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr) + +subsection \<open>Lebesgue integration on @{const count_space}\<close> + +lemma integrable_count_space: + fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}" + shows "finite X \<Longrightarrow> integrable (count_space X) f" + by (auto simp: nn_integral_count_space integrable_iff_bounded) + +lemma measure_count_space[simp]: + "B \<subseteq> A \<Longrightarrow> finite B \<Longrightarrow> measure (count_space A) B = card B" + unfolding measure_def by (subst emeasure_count_space ) auto + +lemma lebesgue_integral_count_space_finite_support: + assumes f: "finite {a\<in>A. f a \<noteq> 0}" + shows "(\<integral>x. f x \<partial>count_space A) = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. f a)" +proof - + have eq: "\<And>x. x \<in> A \<Longrightarrow> (\<Sum>a | x = a \<and> a \<in> A \<and> f a \<noteq> 0. f a) = (\<Sum>x\<in>{x}. f x)" + by (intro setsum.mono_neutral_cong_left) auto + + have "(\<integral>x. f x \<partial>count_space A) = (\<integral>x. (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. indicator {a} x *\<^sub>R f a) \<partial>count_space A)" + by (intro integral_cong refl) (simp add: f eq) + also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. measure (count_space A) {a} *\<^sub>R f a)" + by (subst integral_setsum) (auto intro!: setsum.cong) + finally show ?thesis + by auto +qed + +lemma lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)" + by (subst lebesgue_integral_count_space_finite_support) + (auto intro!: setsum.mono_neutral_cong_left) + +lemma integrable_count_space_nat_iff: + fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}" + shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))" + by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top + intro: summable_suminf_not_top) + +lemma sums_integral_count_space_nat: + fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}" + assumes *: "integrable (count_space UNIV) f" + shows "f sums (integral\<^sup>L (count_space UNIV) f)" +proof - + let ?f = "\<lambda>n i. indicator {n} i *\<^sub>R f i" + have f': "\<And>n i. ?f n i = indicator {n} i *\<^sub>R f n" + by (auto simp: fun_eq_iff split: split_indicator) + + have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) sums \<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV" + proof (rule sums_integral) + show "\<And>i. integrable (count_space UNIV) (?f i)" + using * by (intro integrable_mult_indicator) auto + show "AE n in count_space UNIV. summable (\<lambda>i. norm (?f i n))" + using summable_finite[of "{n}" "\<lambda>i. norm (?f i n)" for n] by simp + show "summable (\<lambda>i. \<integral> n. norm (?f i n) \<partial>count_space UNIV)" + using * by (subst f') (simp add: integrable_count_space_nat_iff) + qed + also have "(\<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV) = (\<integral>n. f n \<partial>count_space UNIV)" + using suminf_finite[of "{n}" "\<lambda>i. ?f i n" for n] by (auto intro!: integral_cong) + also have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) = f" + by (subst f') simp + finally show ?thesis . +qed + +lemma integral_count_space_nat: + fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}" + shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)" + using sums_integral_count_space_nat by (rule sums_unique) + +subsection \<open>Point measure\<close> + +lemma lebesgue_integral_point_measure_finite: + fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + shows "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> + integral\<^sup>L (point_measure A f) g = (\<Sum>a\<in>A. f a *\<^sub>R g a)" + by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def) + +lemma integrable_point_measure_finite: + fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and f :: "'a \<Rightarrow> real" + shows "finite A \<Longrightarrow> integrable (point_measure A f) g" + unfolding point_measure_def + apply (subst density_cong[where f'="\<lambda>x. ennreal (max 0 (f x))"]) + apply (auto split: split_max simp: ennreal_neg) + apply (subst integrable_density) + apply (auto simp: AE_count_space integrable_count_space) + done + +subsection \<open>Lebesgue integration on @{const null_measure}\<close> + +lemma has_bochner_integral_null_measure_iff[iff]: + "has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> borel_measurable M" + by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def] + intro!: exI[of _ "\<lambda>n x. 0"] simple_bochner_integrable.intros) + +lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M" + by (auto simp add: integrable.simps) + +lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0" + by (cases "integrable (null_measure M) f") + (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq) + +subsection \<open>Legacy lemmas for the real-valued Lebesgue integral\<close> + +lemma real_lebesgue_integral_def: + assumes f[measurable]: "integrable M f" + shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+x. f x \<partial>M) - enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)" +proof - + have "integral\<^sup>L M f = integral\<^sup>L M (\<lambda>x. max 0 (f x) - max 0 (- f x))" + by (auto intro!: arg_cong[where f="integral\<^sup>L M"]) + also have "\<dots> = integral\<^sup>L M (\<lambda>x. max 0 (f x)) - integral\<^sup>L M (\<lambda>x. max 0 (- f x))" + by (intro integral_diff integrable_max integrable_minus integrable_zero f) + also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = enn2real (\<integral>\<^sup>+x. ennreal (f x) \<partial>M)" + by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) + also have "integral\<^sup>L M (\<lambda>x. max 0 (- f x)) = enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)" + by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) + finally show ?thesis . +qed + +lemma real_integrable_def: + "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and> + (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" + unfolding integrable_iff_bounded +proof (safe del: notI) + assume *: "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" + have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" + by (intro nn_integral_mono) auto + also note * + finally show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" + by simp + have "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)" + by (intro nn_integral_mono) auto + also note * + finally show "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" + by simp +next + assume [measurable]: "f \<in> borel_measurable M" + assume fin: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" + have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) + ennreal (- f x) \<partial>M)" + by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg) + also have"\<dots> = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) + (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M)" + by (intro nn_integral_add) auto + also have "\<dots> < \<infinity>" + using fin by (auto simp: less_top) + finally show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity>" . +qed + +lemma integrableD[dest]: + assumes "integrable M f" + shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>" + using assms unfolding real_integrable_def by auto + +lemma integrableE: + assumes "integrable M f" + obtains r q where + "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M) = ennreal r" + "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M) = ennreal q" + "f \<in> borel_measurable M" "integral\<^sup>L M f = r - q" + using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms] + by (cases rule: ennreal2_cases[of "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M)"]) auto + +lemma integral_monotone_convergence_nonneg: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real" + assumes i: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)" + and pos: "\<And>i. AE x in M. 0 \<le> f i x" + and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x" + and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x" + and u: "u \<in> borel_measurable M" + shows "integrable M u" + and "integral\<^sup>L M u = x" +proof - + have nn: "AE x in M. \<forall>i. 0 \<le> f i x" + using pos unfolding AE_all_countable by auto + with lim have u_nn: "AE x in M. 0 \<le> u x" + by eventually_elim (auto intro: LIMSEQ_le_const) + have [simp]: "0 \<le> x" + by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos) + have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ennreal (f n x) \<partial>M))" + proof (subst nn_integral_monotone_convergence_SUP_AE[symmetric]) + fix i + from mono nn show "AE x in M. ennreal (f i x) \<le> ennreal (f (Suc i) x)" + by eventually_elim (auto simp: mono_def) + show "(\<lambda>x. ennreal (f i x)) \<in> borel_measurable M" + using i by auto + next + show "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ennreal (f i x)) \<partial>M" + apply (rule nn_integral_cong_AE) + using lim mono nn u_nn + apply eventually_elim + apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def) + done + qed + also have "\<dots> = ennreal x" + using mono i nn unfolding nn_integral_eq_integral[OF i pos] + by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim) + finally have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = ennreal x" . + moreover have "(\<integral>\<^sup>+ x. ennreal (- u x) \<partial>M) = 0" + using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg) + ultimately show "integrable M u" "integral\<^sup>L M u = x" + by (auto simp: real_integrable_def real_lebesgue_integral_def u) +qed + +lemma + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real" + assumes f: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)" + and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x" + and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x" + and u: "u \<in> borel_measurable M" + shows integrable_monotone_convergence: "integrable M u" + and integral_monotone_convergence: "integral\<^sup>L M u = x" + and has_bochner_integral_monotone_convergence: "has_bochner_integral M u x" +proof - + have 1: "\<And>i. integrable M (\<lambda>x. f i x - f 0 x)" + using f by auto + have 2: "AE x in M. mono (\<lambda>n. f n x - f 0 x)" + using mono by (auto simp: mono_def le_fun_def) + have 3: "\<And>n. AE x in M. 0 \<le> f n x - f 0 x" + using mono by (auto simp: field_simps mono_def le_fun_def) + have 4: "AE x in M. (\<lambda>i. f i x - f 0 x) \<longlonglongrightarrow> u x - f 0 x" + using lim by (auto intro!: tendsto_diff) + have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) \<longlonglongrightarrow> x - integral\<^sup>L M (f 0)" + using f ilim by (auto intro!: tendsto_diff) + have 6: "(\<lambda>x. u x - f 0 x) \<in> borel_measurable M" + using f[of 0] u by auto + note diff = integral_monotone_convergence_nonneg[OF 1 2 3 4 5 6] + have "integrable M (\<lambda>x. (u x - f 0 x) + f 0 x)" + using diff(1) f by (rule integrable_add) + with diff(2) f show "integrable M u" "integral\<^sup>L M u = x" + by auto + then show "has_bochner_integral M u x" + by (metis has_bochner_integral_integrable) +qed + +lemma integral_norm_eq_0_iff: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes f[measurable]: "integrable M f" + shows "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0" +proof - + have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>x. norm (f x) \<partial>M)" + using f by (intro nn_integral_eq_integral integrable_norm) auto + then have "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) = 0" + by simp + also have "\<dots> \<longleftrightarrow> emeasure M {x\<in>space M. ennreal (norm (f x)) \<noteq> 0} = 0" + by (intro nn_integral_0_iff) auto + finally show ?thesis + by simp +qed + +lemma integral_0_iff: + fixes f :: "'a \<Rightarrow> real" + shows "integrable M f \<Longrightarrow> (\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0" + using integral_norm_eq_0_iff[of M f] by simp + +lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)" + using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric]) + +lemma lebesgue_integral_const[simp]: + fixes a :: "'a :: {banach, second_countable_topology}" + shows "(\<integral>x. a \<partial>M) = measure M (space M) *\<^sub>R a" +proof - + { assume "emeasure M (space M) = \<infinity>" "a \<noteq> 0" + then have ?thesis + by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) } + moreover + { assume "a = 0" then have ?thesis by simp } + moreover + { assume "emeasure M (space M) \<noteq> \<infinity>" + interpret finite_measure M + proof qed fact + have "(\<integral>x. a \<partial>M) = (\<integral>x. indicator (space M) x *\<^sub>R a \<partial>M)" + by (intro integral_cong) auto + also have "\<dots> = measure M (space M) *\<^sub>R a" + by (simp add: less_top[symmetric]) + finally have ?thesis . } + ultimately show ?thesis by blast +qed + +lemma (in finite_measure) integrable_const_bound: + fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}" + shows "AE x in M. norm (f x) \<le> B \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> integrable M f" + apply (rule integrable_bound[OF integrable_const[of B], of f]) + apply assumption + apply (cases "0 \<le> B") + apply auto + done + +lemma integral_indicator_finite_real: + fixes f :: "'a \<Rightarrow> real" + assumes [simp]: "finite A" + assumes [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M" + assumes finite: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} < \<infinity>" + shows "(\<integral>x. f x * indicator A x \<partial>M) = (\<Sum>a\<in>A. f a * measure M {a})" +proof - + have "(\<integral>x. f x * indicator A x \<partial>M) = (\<integral>x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)" + proof (intro integral_cong refl) + fix x show "f x * indicator A x = (\<Sum>a\<in>A. f a * indicator {a} x)" + by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong) + qed + also have "\<dots> = (\<Sum>a\<in>A. f a * measure M {a})" + using finite by (subst integral_setsum) (auto simp add: integrable_mult_left_iff) + finally show ?thesis . +qed + +lemma (in finite_measure) ennreal_integral_real: + assumes [measurable]: "f \<in> borel_measurable M" + assumes ae: "AE x in M. f x \<le> ennreal B" "0 \<le> B" + shows "ennreal (\<integral>x. enn2real (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)" +proof (subst nn_integral_eq_integral[symmetric]) + show "integrable M (\<lambda>x. enn2real (f x))" + using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI enn2real_nonneg) + show "(\<integral>\<^sup>+ x. ennreal (enn2real (f x)) \<partial>M) = integral\<^sup>N M f" + using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top]) +qed (auto simp: enn2real_nonneg) + +lemma (in finite_measure) integral_less_AE: + fixes X Y :: "'a \<Rightarrow> real" + assumes int: "integrable M X" "integrable M Y" + assumes A: "(emeasure M) A \<noteq> 0" "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> X x \<noteq> Y x" + assumes gt: "AE x in M. X x \<le> Y x" + shows "integral\<^sup>L M X < integral\<^sup>L M Y" +proof - + have "integral\<^sup>L M X \<le> integral\<^sup>L M Y" + using gt int by (intro integral_mono_AE) auto + moreover + have "integral\<^sup>L M X \<noteq> integral\<^sup>L M Y" + proof + assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y" + have "integral\<^sup>L M (\<lambda>x. \<bar>Y x - X x\<bar>) = integral\<^sup>L M (\<lambda>x. Y x - X x)" + using gt int by (intro integral_cong_AE) auto + also have "\<dots> = 0" + using eq int by simp + finally have "(emeasure M) {x \<in> space M. Y x - X x \<noteq> 0} = 0" + using int by (simp add: integral_0_iff) + moreover + have "(\<integral>\<^sup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^sup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)" + using A by (intro nn_integral_mono_AE) auto + then have "(emeasure M) A \<le> (emeasure M) {x \<in> space M. Y x - X x \<noteq> 0}" + using int A by (simp add: integrable_def) + ultimately have "emeasure M A = 0" + by simp + with \<open>(emeasure M) A \<noteq> 0\<close> show False by auto + qed + ultimately show ?thesis by auto +qed + +lemma (in finite_measure) integral_less_AE_space: + fixes X Y :: "'a \<Rightarrow> real" + assumes int: "integrable M X" "integrable M Y" + assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \<noteq> 0" + shows "integral\<^sup>L M X < integral\<^sup>L M Y" + using gt by (intro integral_less_AE[OF int, where A="space M"]) auto + +lemma tendsto_integral_at_top: + fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}" + assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f" + shows "((\<lambda>y. \<integral> x. indicator {.. y} x *\<^sub>R f x \<partial>M) \<longlongrightarrow> \<integral> x. f x \<partial>M) at_top" +proof (rule tendsto_at_topI_sequentially) + fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially" + show "(\<lambda>n. \<integral>x. indicator {..X n} x *\<^sub>R f x \<partial>M) \<longlonglongrightarrow> integral\<^sup>L M f" + proof (rule integral_dominated_convergence) + show "integrable M (\<lambda>x. norm (f x))" + by (rule integrable_norm) fact + show "AE x in M. (\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x" + proof + fix x + from \<open>filterlim X at_top sequentially\<close> + have "eventually (\<lambda>n. x \<le> X n) sequentially" + unfolding filterlim_at_top_ge[where c=x] by auto + then show "(\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x" + by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono) + qed + fix n show "AE x in M. norm (indicator {..X n} x *\<^sub>R f x) \<le> norm (f x)" + by (auto split: split_indicator) + qed auto +qed + +lemma + fixes f :: "real \<Rightarrow> real" + assumes M: "sets M = sets borel" + assumes nonneg: "AE x in M. 0 \<le> f x" + assumes borel: "f \<in> borel_measurable borel" + assumes int: "\<And>y. integrable M (\<lambda>x. f x * indicator {.. y} x)" + assumes conv: "((\<lambda>y. \<integral> x. f x * indicator {.. y} x \<partial>M) \<longlongrightarrow> x) at_top" + shows has_bochner_integral_monotone_convergence_at_top: "has_bochner_integral M f x" + and integrable_monotone_convergence_at_top: "integrable M f" + and integral_monotone_convergence_at_top:"integral\<^sup>L M f = x" +proof - + from nonneg have "AE x in M. mono (\<lambda>n::nat. f x * indicator {..real n} x)" + by (auto split: split_indicator intro!: monoI) + { fix x have "eventually (\<lambda>n. f x * indicator {..real n} x = f x) sequentially" + by (rule eventually_sequentiallyI[of "nat \<lceil>x\<rceil>"]) + (auto split: split_indicator simp: nat_le_iff ceiling_le_iff) } + from filterlim_cong[OF refl refl this] + have "AE x in M. (\<lambda>i. f x * indicator {..real i} x) \<longlonglongrightarrow> f x" + by simp + have "(\<lambda>i. \<integral> x. f x * indicator {..real i} x \<partial>M) \<longlonglongrightarrow> x" + using conv filterlim_real_sequentially by (rule filterlim_compose) + have M_measure[simp]: "borel_measurable M = borel_measurable borel" + using M by (simp add: sets_eq_imp_space_eq measurable_def) + have "f \<in> borel_measurable M" + using borel by simp + show "has_bochner_integral M f x" + by (rule has_bochner_integral_monotone_convergence) fact+ + then show "integrable M f" "integral\<^sup>L M f = x" + by (auto simp: _has_bochner_integral_iff) +qed + +subsection \<open>Product measure\<close> + +lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]: + fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" + assumes [measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)" + shows "Measurable.pred N (\<lambda>x. integrable M (f x))" +proof - + have [simp]: "\<And>x. x \<in> space N \<Longrightarrow> integrable M (f x) \<longleftrightarrow> (\<integral>\<^sup>+y. norm (f x y) \<partial>M) < \<infinity>" + unfolding integrable_iff_bounded by simp + show ?thesis + by (simp cong: measurable_cong) +qed + +lemma Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> A" by auto + +lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]: + "(\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M) \<Longrightarrow> + {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M) \<Longrightarrow> + (\<lambda>x. measure M (A x)) \<in> borel_measurable N" + unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto + +lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]: + fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f[measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)" + shows "(\<lambda>x. \<integral>y. f x y \<partial>M) \<in> borel_measurable N" +proof - + from borel_measurable_implies_sequence_metric[OF f, of 0] guess s .. + then have s: "\<And>i. simple_function (N \<Otimes>\<^sub>M M) (s i)" + "\<And>x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> (\<lambda>i. s i (x, y)) \<longlonglongrightarrow> f x y" + "\<And>i x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> norm (s i (x, y)) \<le> 2 * norm (f x y)" + by (auto simp: space_pair_measure) + + have [measurable]: "\<And>i. s i \<in> borel_measurable (N \<Otimes>\<^sub>M M)" + by (rule borel_measurable_simple_function) fact + + have "\<And>i. s i \<in> measurable (N \<Otimes>\<^sub>M M) (count_space UNIV)" + by (rule measurable_simple_function) fact + + define f' where [abs_def]: "f' i x = + (if integrable M (f x) then simple_bochner_integral M (\<lambda>y. s i (x, y)) else 0)" for i x + + { fix i x assume "x \<in> space N" + then have "simple_bochner_integral M (\<lambda>y. s i (x, y)) = + (\<Sum>z\<in>s i ` (space N \<times> space M). measure M {y \<in> space M. s i (x, y) = z} *\<^sub>R z)" + using s(1)[THEN simple_functionD(1)] + unfolding simple_bochner_integral_def + by (intro setsum.mono_neutral_cong_left) + (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) } + note eq = this + + show ?thesis + proof (rule borel_measurable_LIMSEQ_metric) + fix i show "f' i \<in> borel_measurable N" + unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong) + next + fix x assume x: "x \<in> space N" + { assume int_f: "integrable M (f x)" + have int_2f: "integrable M (\<lambda>y. 2 * norm (f x y))" + by (intro integrable_norm integrable_mult_right int_f) + have "(\<lambda>i. integral\<^sup>L M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)" + proof (rule integral_dominated_convergence) + from int_f show "f x \<in> borel_measurable M" by auto + show "\<And>i. (\<lambda>y. s i (x, y)) \<in> borel_measurable M" + using x by simp + show "AE xa in M. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f x xa" + using x s(2) by auto + show "\<And>i. AE xa in M. norm (s i (x, xa)) \<le> 2 * norm (f x xa)" + using x s(3) by auto + qed fact + moreover + { fix i + have "simple_bochner_integrable M (\<lambda>y. s i (x, y))" + proof (rule simple_bochner_integrableI_bounded) + have "(\<lambda>y. s i (x, y)) ` space M \<subseteq> s i ` (space N \<times> space M)" + using x by auto + then show "simple_function M (\<lambda>y. s i (x, y))" + using simple_functionD(1)[OF s(1), of i] x + by (intro simple_function_borel_measurable) + (auto simp: space_pair_measure dest: finite_subset) + have "(\<integral>\<^sup>+ y. ennreal (norm (s i (x, y))) \<partial>M) \<le> (\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M)" + using x s by (intro nn_integral_mono) auto + also have "(\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M) < \<infinity>" + using int_2f by (simp add: integrable_iff_bounded) + finally show "(\<integral>\<^sup>+ xa. ennreal (norm (s i (x, xa))) \<partial>M) < \<infinity>" . + qed + then have "integral\<^sup>L M (\<lambda>y. s i (x, y)) = simple_bochner_integral M (\<lambda>y. s i (x, y))" + by (rule simple_bochner_integrable_eq_integral[symmetric]) } + ultimately have "(\<lambda>i. simple_bochner_integral M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)" + by simp } + then + show "(\<lambda>i. f' i x) \<longlonglongrightarrow> integral\<^sup>L M (f x)" + unfolding f'_def + by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq) + qed +qed + +lemma (in pair_sigma_finite) integrable_product_swap: + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes "integrable (M1 \<Otimes>\<^sub>M M2) f" + shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x))" +proof - + interpret Q: pair_sigma_finite M2 M1 .. + have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff) + show ?thesis unfolding * + by (rule integrable_distr[OF measurable_pair_swap']) + (simp add: distr_pair_swap[symmetric] assms) +qed + +lemma (in pair_sigma_finite) integrable_product_swap_iff: + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^sub>M M2) f" +proof - + interpret Q: pair_sigma_finite M2 M1 .. + from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f] + show ?thesis by auto +qed + +lemma (in pair_sigma_finite) integral_product_swap: + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" + shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" +proof - + have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff) + show ?thesis unfolding * + by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) +qed + +lemma (in pair_sigma_finite) Fubini_integrable: + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" + and integ1: "integrable M1 (\<lambda>x. \<integral> y. norm (f (x, y)) \<partial>M2)" + and integ2: "AE x in M1. integrable M2 (\<lambda>y. f (x, y))" + shows "integrable (M1 \<Otimes>\<^sub>M M2) f" +proof (rule integrableI_bounded) + have "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. norm (f (x, y)) \<partial>M2) \<partial>M1)" + by (simp add: M2.nn_integral_fst [symmetric]) + also have "\<dots> = (\<integral>\<^sup>+ x. \<bar>\<integral>y. norm (f (x, y)) \<partial>M2\<bar> \<partial>M1)" + apply (intro nn_integral_cong_AE) + using integ2 + proof eventually_elim + fix x assume "integrable M2 (\<lambda>y. f (x, y))" + then have f: "integrable M2 (\<lambda>y. norm (f (x, y)))" + by simp + then have "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal (LINT y|M2. norm (f (x, y)))" + by (rule nn_integral_eq_integral) simp + also have "\<dots> = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>" + using f by simp + finally show "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>" . + qed + also have "\<dots> < \<infinity>" + using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE) + finally show "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>" . +qed fact + +lemma (in pair_sigma_finite) emeasure_pair_measure_finite: + assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" and finite: "emeasure (M1 \<Otimes>\<^sub>M M2) A < \<infinity>" + shows "AE x in M1. emeasure M2 {y\<in>space M2. (x, y) \<in> A} < \<infinity>" +proof - + from M2.emeasure_pair_measure_alt[OF A] finite + have "(\<integral>\<^sup>+ x. emeasure M2 (Pair x -` A) \<partial>M1) \<noteq> \<infinity>" + by simp + then have "AE x in M1. emeasure M2 (Pair x -` A) \<noteq> \<infinity>" + by (rule nn_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A) + moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> Pair x -` A = {y\<in>space M2. (x, y) \<in> A}" + using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) + ultimately show ?thesis by (auto simp: less_top) +qed + +lemma (in pair_sigma_finite) AE_integrable_fst': + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f" + shows "AE x in M1. integrable M2 (\<lambda>y. f (x, y))" +proof - + have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1) = (\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2))" + by (rule M2.nn_integral_fst) simp + also have "(\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2)) \<noteq> \<infinity>" + using f unfolding integrable_iff_bounded by simp + finally have "AE x in M1. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<noteq> \<infinity>" + by (intro nn_integral_PInf_AE M2.borel_measurable_nn_integral ) + (auto simp: measurable_split_conv) + with AE_space show ?thesis + by eventually_elim + (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top) +qed + +lemma (in pair_sigma_finite) integrable_fst': + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f" + shows "integrable M1 (\<lambda>x. \<integral>y. f (x, y) \<partial>M2)" + unfolding integrable_iff_bounded +proof + show "(\<lambda>x. \<integral> y. f (x, y) \<partial>M2) \<in> borel_measurable M1" + by (rule M2.borel_measurable_lebesgue_integral) simp + have "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) \<le> (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1)" + using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal) + also have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1) = (\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2))" + by (rule M2.nn_integral_fst) simp + also have "(\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>" + using f unfolding integrable_iff_bounded by simp + finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" . +qed + +lemma (in pair_sigma_finite) integral_fst': + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f" + shows "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" +using f proof induct + case (base A c) + have A[measurable]: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" by fact + + have eq: "\<And>x y. x \<in> space M1 \<Longrightarrow> indicator A (x, y) = indicator {y\<in>space M2. (x, y) \<in> A} y" + using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure) + + have int_A: "integrable (M1 \<Otimes>\<^sub>M M2) (indicator A :: _ \<Rightarrow> real)" + using base by (rule integrable_real_indicator) + + have "(\<integral> x. \<integral> y. indicator A (x, y) *\<^sub>R c \<partial>M2 \<partial>M1) = (\<integral>x. measure M2 {y\<in>space M2. (x, y) \<in> A} *\<^sub>R c \<partial>M1)" + proof (intro integral_cong_AE, simp, simp) + from AE_integrable_fst'[OF int_A] AE_space + show "AE x in M1. (\<integral>y. indicator A (x, y) *\<^sub>R c \<partial>M2) = measure M2 {y\<in>space M2. (x, y) \<in> A} *\<^sub>R c" + by eventually_elim (simp add: eq integrable_indicator_iff) + qed + also have "\<dots> = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c" + proof (subst integral_scaleR_left) + have "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = + (\<integral>\<^sup>+x. emeasure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1)" + using emeasure_pair_measure_finite[OF base] + by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure) + also have "\<dots> = emeasure (M1 \<Otimes>\<^sub>M M2) A" + using sets.sets_into_space[OF A] + by (subst M2.emeasure_pair_measure_alt) + (auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure) + finally have *: "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = emeasure (M1 \<Otimes>\<^sub>M M2) A" . + + from base * show "integrable M1 (\<lambda>x. measure M2 {y \<in> space M2. (x, y) \<in> A})" + by (simp add: integrable_iff_bounded) + then have "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) = + (\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1)" + by (rule nn_integral_eq_integral[symmetric]) simp + also note * + finally show "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) *\<^sub>R c = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c" + using base by (simp add: emeasure_eq_ennreal_measure) + qed + also have "\<dots> = (\<integral> a. indicator A a *\<^sub>R c \<partial>(M1 \<Otimes>\<^sub>M M2))" + using base by simp + finally show ?case . +next + case (add f g) + then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "g \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" + by auto + have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) = + (\<integral> x. (\<integral> y. f (x, y) \<partial>M2) + (\<integral> y. g (x, y) \<partial>M2) \<partial>M1)" + apply (rule integral_cong_AE) + apply simp_all + using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)] + apply eventually_elim + apply simp + done + also have "\<dots> = (\<integral> x. f x \<partial>(M1 \<Otimes>\<^sub>M M2)) + (\<integral> x. g x \<partial>(M1 \<Otimes>\<^sub>M M2))" + using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp + finally show ?case + using add by simp +next + case (lim f s) + then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "\<And>i. s i \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" + by auto + + show ?case + proof (rule LIMSEQ_unique) + show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" + proof (rule integral_dominated_convergence) + show "integrable (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. 2 * norm (f x))" + using lim(5) by auto + qed (insert lim, auto) + have "(\<lambda>i. \<integral> x. \<integral> y. s i (x, y) \<partial>M2 \<partial>M1) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1" + proof (rule integral_dominated_convergence) + have "AE x in M1. \<forall>i. integrable M2 (\<lambda>y. s i (x, y))" + unfolding AE_all_countable using AE_integrable_fst'[OF lim(1)] .. + with AE_space AE_integrable_fst'[OF lim(5)] + show "AE x in M1. (\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2" + proof eventually_elim + fix x assume x: "x \<in> space M1" and + s: "\<forall>i. integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))" + show "(\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2" + proof (rule integral_dominated_convergence) + show "integrable M2 (\<lambda>y. 2 * norm (f (x, y)))" + using f by auto + show "AE xa in M2. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f (x, xa)" + using x lim(3) by (auto simp: space_pair_measure) + show "\<And>i. AE xa in M2. norm (s i (x, xa)) \<le> 2 * norm (f (x, xa))" + using x lim(4) by (auto simp: space_pair_measure) + qed (insert x, measurable) + qed + show "integrable M1 (\<lambda>x. (\<integral> y. 2 * norm (f (x, y)) \<partial>M2))" + by (intro integrable_mult_right integrable_norm integrable_fst' lim) + fix i show "AE x in M1. norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral> y. 2 * norm (f (x, y)) \<partial>M2)" + using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)] + proof eventually_elim + fix x assume x: "x \<in> space M1" + and s: "integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))" + from s have "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral>\<^sup>+y. norm (s i (x, y)) \<partial>M2)" + by (rule integral_norm_bound_ennreal) + also have "\<dots> \<le> (\<integral>\<^sup>+y. 2 * norm (f (x, y)) \<partial>M2)" + using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure) + also have "\<dots> = (\<integral>y. 2 * norm (f (x, y)) \<partial>M2)" + using f by (intro nn_integral_eq_integral) auto + finally show "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral> y. 2 * norm (f (x, y)) \<partial>M2)" + by simp + qed + qed simp_all + then show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1" + using lim by simp + qed +qed + +lemma (in pair_sigma_finite) + fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)" + shows AE_integrable_fst: "AE x in M1. integrable M2 (\<lambda>y. f x y)" (is "?AE") + and integrable_fst: "integrable M1 (\<lambda>x. \<integral>y. f x y \<partial>M2)" (is "?INT") + and integral_fst: "(\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). f x y)" (is "?EQ") + using AE_integrable_fst'[OF f] integrable_fst'[OF f] integral_fst'[OF f] by auto + +lemma (in pair_sigma_finite) + fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}" + assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)" + shows AE_integrable_snd: "AE y in M2. integrable M1 (\<lambda>x. f x y)" (is "?AE") + and integrable_snd: "integrable M2 (\<lambda>y. \<integral>x. f x y \<partial>M1)" (is "?INT") + and integral_snd: "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (case_prod f)" (is "?EQ") +proof - + interpret Q: pair_sigma_finite M2 M1 .. + have Q_int: "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x, y). f y x)" + using f unfolding integrable_product_swap_iff[symmetric] by simp + show ?AE using Q.AE_integrable_fst'[OF Q_int] by simp + show ?INT using Q.integrable_fst'[OF Q_int] by simp + show ?EQ using Q.integral_fst'[OF Q_int] + using integral_product_swap[of "case_prod f"] by simp +qed + +lemma (in pair_sigma_finite) Fubini_integral: + fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: {banach, second_countable_topology}" + assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)" + shows "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1)" + unfolding integral_snd[OF assms] integral_fst[OF assms] .. + +lemma (in product_sigma_finite) product_integral_singleton: + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + shows "f \<in> borel_measurable (M i) \<Longrightarrow> (\<integral>x. f (x i) \<partial>Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" + apply (subst distr_singleton[symmetric]) + apply (subst integral_distr) + apply simp_all + done + +lemma (in product_sigma_finite) product_integral_fold: + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J" + and f: "integrable (Pi\<^sub>M (I \<union> J) M) f" + shows "integral\<^sup>L (Pi\<^sub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^sub>M J M) \<partial>Pi\<^sub>M I M)" +proof - + interpret I: finite_product_sigma_finite M I by standard fact + interpret J: finite_product_sigma_finite M J by standard fact + have "finite (I \<union> J)" using fin by auto + interpret IJ: finite_product_sigma_finite M "I \<union> J" by standard fact + interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" .. + let ?M = "merge I J" + let ?f = "\<lambda>x. f (?M x)" + from f have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)" + by auto + have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)" + using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) + have f_int: "integrable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) ?f" + by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) + show ?thesis + apply (subst distr_merge[symmetric, OF IJ fin]) + apply (subst integral_distr[OF measurable_merge f_borel]) + apply (subst P.integral_fst'[symmetric, OF f_int]) + apply simp + done +qed + +lemma (in product_sigma_finite) product_integral_insert: + fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" + assumes I: "finite I" "i \<notin> I" + and f: "integrable (Pi\<^sub>M (insert i I) M) f" + shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^sub>M I M)" +proof - + have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \<union> {i}) M) f" + by simp + also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) \<partial>Pi\<^sub>M I M)" + using f I by (intro product_integral_fold) auto + also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^sub>M I M)" + proof (rule integral_cong[OF refl], subst product_integral_singleton[symmetric]) + fix x assume x: "x \<in> space (Pi\<^sub>M I M)" + have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)" + using f by auto + show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)" + using measurable_comp[OF measurable_component_update f_borel, OF x \<open>i \<notin> I\<close>] + unfolding comp_def . + from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^sub>M {i} M)" + by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) + qed + finally show ?thesis . +qed + +lemma (in product_sigma_finite) product_integrable_setprod: + fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}" + assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)" + shows "integrable (Pi\<^sub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f") +proof (unfold integrable_iff_bounded, intro conjI) + interpret finite_product_sigma_finite M I by standard fact + + show "?f \<in> borel_measurable (Pi\<^sub>M I M)" + using assms by simp + have "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) = + (\<integral>\<^sup>+ x. (\<Prod>i\<in>I. ennreal (norm (f i (x i)))) \<partial>Pi\<^sub>M I M)" + by (simp add: setprod_norm setprod_ennreal) + also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+ x. ennreal (norm (f i x)) \<partial>M i)" + using assms by (intro product_nn_integral_setprod) auto + also have "\<dots> < \<infinity>" + using integrable by (simp add: less_top[symmetric] ennreal_setprod_eq_top integrable_iff_bounded) + finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" . +qed + +lemma (in product_sigma_finite) product_integral_setprod: + fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}" + assumes "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)" + shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>L (M i) (f i))" +using assms proof induct + case empty + interpret finite_measure "Pi\<^sub>M {} M" + by rule (simp add: space_PiM) + show ?case by (simp add: space_PiM measure_def) +next + case (insert i I) + then have iI: "finite (insert i I)" by auto + then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> + integrable (Pi\<^sub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))" + by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset) + interpret I: finite_product_sigma_finite M I by standard fact + have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" + using \<open>i \<notin> I\<close> by (auto intro!: setprod.cong) + show ?case + unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] + by (simp add: * insert prod subset_insertI) +qed + +lemma integrable_subalgebra: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes borel: "f \<in> borel_measurable N" + and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A" + shows "integrable N f \<longleftrightarrow> integrable M f" (is ?P) +proof - + have "f \<in> borel_measurable M" + using assms by (auto simp: measurable_def) + with assms show ?thesis + using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra) +qed + +lemma integral_subalgebra: + fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes borel: "f \<in> borel_measurable N" + and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A" + shows "integral\<^sup>L N f = integral\<^sup>L M f" +proof cases + assume "integrable N f" + then show ?thesis + proof induct + case base with assms show ?case by (auto simp: subset_eq measure_def) + next + case (add f g) + then have "(\<integral> a. f a + g a \<partial>N) = integral\<^sup>L M f + integral\<^sup>L M g" + by simp + also have "\<dots> = (\<integral> a. f a + g a \<partial>M)" + using add integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of g] by simp + finally show ?case . + next + case (lim f s) + then have M: "\<And>i. integrable M (s i)" "integrable M f" + using integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of "s i" for i] by simp_all + show ?case + proof (intro LIMSEQ_unique) + show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L N f" + apply (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) + using lim + apply auto + done + show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L M f" + unfolding lim + apply (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) + using lim M N(2) + apply auto + done + qed + qed +qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms]) + +hide_const (open) simple_bochner_integral +hide_const (open) simple_bochner_integrable + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Borel_Space.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,1915 @@ +(* Title: HOL/Analysis/Borel_Space.thy + Author: Johannes Hölzl, TU München + Author: Armin Heller, TU München +*) + +section \<open>Borel spaces\<close> + +theory Borel_Space +imports + Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits +begin + +lemma sets_Collect_eventually_sequentially[measurable]: + "(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M" + unfolding eventually_sequentially by simp + +lemma topological_basis_trivial: "topological_basis {A. open A}" + by (auto simp: topological_basis_def) + +lemma open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}" +proof - + have "{A \<times> B :: ('a \<times> 'b) set | A B. open A \<and> open B} = ((\<lambda>(a, b). a \<times> b) ` ({A. open A} \<times> {A. open A}))" + by auto + then show ?thesis + by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis) +qed + +definition "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s" + +lemma mono_onI: + "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A" + unfolding mono_on_def by simp + +lemma mono_onD: + "\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s" + unfolding mono_on_def by simp + +lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A" + unfolding mono_def mono_on_def by auto + +lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B" + unfolding mono_on_def by auto + +definition "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s" + +lemma strict_mono_onI: + "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A" + unfolding strict_mono_on_def by simp + +lemma strict_mono_onD: + "\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s" + unfolding strict_mono_on_def by simp + +lemma mono_on_greaterD: + assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)" + shows "x > y" +proof (rule ccontr) + assume "\<not>x > y" + hence "x \<le> y" by (simp add: not_less) + from assms(1-3) and this have "g x \<le> g y" by (rule mono_onD) + with assms(4) show False by simp +qed + +lemma strict_mono_inv: + fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)" + assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x" + shows "strict_mono g" +proof + fix x y :: 'b assume "x < y" + from \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast + with \<open>x < y\<close> and \<open>strict_mono f\<close> have "x' < y'" by (simp add: strict_mono_less) + with inv show "g x < g y" by simp +qed + +lemma strict_mono_on_imp_inj_on: + assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A" + shows "inj_on f A" +proof (rule inj_onI) + fix x y assume "x \<in> A" "y \<in> A" "f x = f y" + thus "x = y" + by (cases x y rule: linorder_cases) + (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) +qed + +lemma strict_mono_on_leD: + assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y" + shows "f x \<le> f y" +proof (insert le_less_linear[of y x], elim disjE) + assume "x < y" + with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all + thus ?thesis by (rule less_imp_le) +qed (insert assms, simp) + +lemma strict_mono_on_eqD: + fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)" + assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A" + shows "y = x" + using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD) + +lemma mono_on_imp_deriv_nonneg: + assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)" + assumes "x \<in> interior A" + shows "D \<ge> 0" +proof (rule tendsto_le_const) + let ?A' = "(\<lambda>y. y - x) ` interior A" + from deriv show "((\<lambda>h. (f (x + h) - f x) / h) \<longlongrightarrow> D) (at 0)" + by (simp add: field_has_derivative_at has_field_derivative_def) + from mono have mono': "mono_on f (interior A)" by (rule mono_on_subset) (rule interior_subset) + + show "eventually (\<lambda>h. (f (x + h) - f x) / h \<ge> 0) (at 0)" + proof (subst eventually_at_topological, intro exI conjI ballI impI) + have "open (interior A)" by simp + hence "open (op + (-x) ` interior A)" by (rule open_translation) + also have "(op + (-x) ` interior A) = ?A'" by auto + finally show "open ?A'" . + next + from \<open>x \<in> interior A\<close> show "0 \<in> ?A'" by auto + next + fix h assume "h \<in> ?A'" + hence "x + h \<in> interior A" by auto + with mono' and \<open>x \<in> interior A\<close> show "(f (x + h) - f x) / h \<ge> 0" + by (cases h rule: linorder_cases[of _ 0]) + (simp_all add: divide_nonpos_neg divide_nonneg_pos mono_onD field_simps) + qed +qed simp + +lemma strict_mono_on_imp_mono_on: + "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A" + by (rule mono_onI, rule strict_mono_on_leD) + +lemma mono_on_ctble_discont: + fixes f :: "real \<Rightarrow> real" + fixes A :: "real set" + assumes "mono_on f A" + shows "countable {a\<in>A. \<not> continuous (at a within A) f}" +proof - + have mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y" + using \<open>mono_on f A\<close> by (simp add: mono_on_def) + have "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. \<exists>q :: nat \<times> rat. + (fst q = 0 \<and> of_rat (snd q) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd q))) \<or> + (fst q = 1 \<and> of_rat (snd q) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd q)))" + proof (clarsimp simp del: One_nat_def) + fix a assume "a \<in> A" assume "\<not> continuous (at a within A) f" + thus "\<exists>q1 q2. + q1 = 0 \<and> real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2) \<or> + q1 = 1 \<and> f a < real_of_rat q2 \<and> (\<forall>x\<in>A. a < x \<longrightarrow> real_of_rat q2 < f x)" + proof (auto simp add: continuous_within order_tendsto_iff eventually_at) + fix l assume "l < f a" + then obtain q2 where q2: "l < of_rat q2" "of_rat q2 < f a" + using of_rat_dense by blast + assume * [rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> l < f x" + from q2 have "real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2)" + proof auto + fix x assume "x \<in> A" "x < a" + with q2 *[of "a - x"] show "f x < real_of_rat q2" + apply (auto simp add: dist_real_def not_less) + apply (subgoal_tac "f x \<le> f xa") + by (auto intro: mono) + qed + thus ?thesis by auto + next + fix u assume "u > f a" + then obtain q2 where q2: "f a < of_rat q2" "of_rat q2 < u" + using of_rat_dense by blast + assume *[rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> u > f x" + from q2 have "real_of_rat q2 > f a \<and> (\<forall>x\<in>A. x > a \<longrightarrow> f x > real_of_rat q2)" + proof auto + fix x assume "x \<in> A" "x > a" + with q2 *[of "x - a"] show "f x > real_of_rat q2" + apply (auto simp add: dist_real_def) + apply (subgoal_tac "f x \<ge> f xa") + by (auto intro: mono) + qed + thus ?thesis by auto + qed + qed + hence "\<exists>g :: real \<Rightarrow> nat \<times> rat . \<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. + (fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd (g a)))) | + (fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd (g a))))" + by (rule bchoice) + then guess g .. + hence g: "\<And>a x. a \<in> A \<Longrightarrow> \<not> continuous (at a within A) f \<Longrightarrow> x \<in> A \<Longrightarrow> + (fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (x < a \<longrightarrow> f x < of_rat (snd (g a)))) | + (fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (x > a \<longrightarrow> f x > of_rat (snd (g a))))" + by auto + have "inj_on g {a\<in>A. \<not> continuous (at a within A) f}" + proof (auto simp add: inj_on_def) + fix w z + assume 1: "w \<in> A" and 2: "\<not> continuous (at w within A) f" and + 3: "z \<in> A" and 4: "\<not> continuous (at z within A) f" and + 5: "g w = g z" + from g [OF 1 2 3] g [OF 3 4 1] 5 + show "w = z" by auto + qed + thus ?thesis + by (rule countableI') +qed + +lemma mono_on_ctble_discont_open: + fixes f :: "real \<Rightarrow> real" + fixes A :: "real set" + assumes "open A" "mono_on f A" + shows "countable {a\<in>A. \<not>isCont f a}" +proof - + have "{a\<in>A. \<not>isCont f a} = {a\<in>A. \<not>(continuous (at a within A) f)}" + by (auto simp add: continuous_within_open [OF _ \<open>open A\<close>]) + thus ?thesis + apply (elim ssubst) + by (rule mono_on_ctble_discont, rule assms) +qed + +lemma mono_ctble_discont: + fixes f :: "real \<Rightarrow> real" + assumes "mono f" + shows "countable {a. \<not> isCont f a}" +using assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto + +lemma has_real_derivative_imp_continuous_on: + assumes "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)" + shows "continuous_on A f" + apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def) + apply (intro ballI Deriv.differentiableI) + apply (rule has_field_derivative_subset[OF assms]) + apply simp_all + done + +lemma closure_contains_Sup: + fixes S :: "real set" + assumes "S \<noteq> {}" "bdd_above S" + shows "Sup S \<in> closure S" +proof- + have "Inf (uminus ` S) \<in> closure (uminus ` S)" + using assms by (intro closure_contains_Inf) auto + also have "Inf (uminus ` S) = -Sup S" by (simp add: Inf_real_def) + also have "closure (uminus ` S) = uminus ` closure S" + by (rule sym, intro closure_injective_linear_image) (auto intro: linearI) + finally show ?thesis by auto +qed + +lemma closed_contains_Sup: + fixes S :: "real set" + shows "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> closed S \<Longrightarrow> Sup S \<in> S" + by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup) + +lemma deriv_nonneg_imp_mono: + assumes deriv: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)" + assumes nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0" + assumes ab: "a \<le> b" + shows "g a \<le> g b" +proof (cases "a < b") + assume "a < b" + from deriv have "\<forall>x. x \<ge> a \<and> x \<le> b \<longrightarrow> (g has_real_derivative g' x) (at x)" by simp + from MVT2[OF \<open>a < b\<close> this] and deriv + obtain \<xi> where \<xi>_ab: "\<xi> > a" "\<xi> < b" and g_ab: "g b - g a = (b - a) * g' \<xi>" by blast + from \<xi>_ab ab nonneg have "(b - a) * g' \<xi> \<ge> 0" by simp + with g_ab show ?thesis by simp +qed (insert ab, simp) + +lemma continuous_interval_vimage_Int: + assumes "continuous_on {a::real..b} g" and mono: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y" + assumes "a \<le> b" "(c::real) \<le> d" "{c..d} \<subseteq> {g a..g b}" + obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d" +proof- + let ?A = "{a..b} \<inter> g -` {c..d}" + from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5) + obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto + from IVT'[of g a d b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5) + obtain d'' where d'': "d'' \<in> ?A" "g d'' = d" by auto + hence [simp]: "?A \<noteq> {}" by blast + + define c' where "c' = Inf ?A" + define d' where "d' = Sup ?A" + have "?A \<subseteq> {c'..d'}" unfolding c'_def d'_def + by (intro subsetI) (auto intro: cInf_lower cSup_upper) + moreover from assms have "closed ?A" + using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp + hence c'd'_in_set: "c' \<in> ?A" "d' \<in> ?A" unfolding c'_def d'_def + by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+ + hence "{c'..d'} \<subseteq> ?A" using assms + by (intro subsetI) + (auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x] + intro!: mono) + moreover have "c' \<le> d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto + moreover have "g c' \<le> c" "g d' \<ge> d" + apply (insert c'' d'' c'd'_in_set) + apply (subst c''(2)[symmetric]) + apply (auto simp: c'_def intro!: mono cInf_lower c'') [] + apply (subst d''(2)[symmetric]) + apply (auto simp: d'_def intro!: mono cSup_upper d'') [] + done + with c'd'_in_set have "g c' = c" "g d' = d" by auto + ultimately show ?thesis using that by blast +qed + +subsection \<open>Generic Borel spaces\<close> + +definition (in topological_space) borel :: "'a measure" where + "borel = sigma UNIV {S. open S}" + +abbreviation "borel_measurable M \<equiv> measurable M borel" + +lemma in_borel_measurable: + "f \<in> borel_measurable M \<longleftrightarrow> + (\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)" + by (auto simp add: measurable_def borel_def) + +lemma in_borel_measurable_borel: + "f \<in> borel_measurable M \<longleftrightarrow> + (\<forall>S \<in> sets borel. + f -` S \<inter> space M \<in> sets M)" + by (auto simp add: measurable_def borel_def) + +lemma space_borel[simp]: "space borel = UNIV" + unfolding borel_def by auto + +lemma space_in_borel[measurable]: "UNIV \<in> sets borel" + unfolding borel_def by auto + +lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}" + unfolding borel_def by (rule sets_measure_of) simp + +lemma measurable_sets_borel: + "\<lbrakk>f \<in> measurable borel M; A \<in> sets M\<rbrakk> \<Longrightarrow> f -` A \<in> sets borel" + by (drule (1) measurable_sets) simp + +lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel" + unfolding borel_def pred_def by auto + +lemma borel_open[measurable (raw generic)]: + assumes "open A" shows "A \<in> sets borel" +proof - + have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms . + thus ?thesis unfolding borel_def by auto +qed + +lemma borel_closed[measurable (raw generic)]: + assumes "closed A" shows "A \<in> sets borel" +proof - + have "space borel - (- A) \<in> sets borel" + using assms unfolding closed_def by (blast intro: borel_open) + thus ?thesis by simp +qed + +lemma borel_singleton[measurable]: + "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets (borel :: 'a::t1_space measure)" + unfolding insert_def by (rule sets.Un) auto + +lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel" + unfolding Compl_eq_Diff_UNIV by simp + +lemma borel_measurable_vimage: + fixes f :: "'a \<Rightarrow> 'x::t2_space" + assumes borel[measurable]: "f \<in> borel_measurable M" + shows "f -` {x} \<inter> space M \<in> sets M" + by simp + +lemma borel_measurableI: + fixes f :: "'a \<Rightarrow> 'x::topological_space" + assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M" + shows "f \<in> borel_measurable M" + unfolding borel_def +proof (rule measurable_measure_of, simp_all) + fix S :: "'x set" assume "open S" thus "f -` S \<inter> space M \<in> sets M" + using assms[of S] by simp +qed + +lemma borel_measurable_const: + "(\<lambda>x. c) \<in> borel_measurable M" + by auto + +lemma borel_measurable_indicator: + assumes A: "A \<in> sets M" + shows "indicator A \<in> borel_measurable M" + unfolding indicator_def [abs_def] using A + by (auto intro!: measurable_If_set) + +lemma borel_measurable_count_space[measurable (raw)]: + "f \<in> borel_measurable (count_space S)" + unfolding measurable_def by auto + +lemma borel_measurable_indicator'[measurable (raw)]: + assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M" + shows "(\<lambda>x. indicator (A x) (f x)) \<in> borel_measurable M" + unfolding indicator_def[abs_def] + by (auto intro!: measurable_If) + +lemma borel_measurable_indicator_iff: + "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M" + (is "?I \<in> borel_measurable M \<longleftrightarrow> _") +proof + assume "?I \<in> borel_measurable M" + then have "?I -` {1} \<inter> space M \<in> sets M" + unfolding measurable_def by auto + also have "?I -` {1} \<inter> space M = A \<inter> space M" + unfolding indicator_def [abs_def] by auto + finally show "A \<inter> space M \<in> sets M" . +next + assume "A \<inter> space M \<in> sets M" + moreover have "?I \<in> borel_measurable M \<longleftrightarrow> + (indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M" + by (intro measurable_cong) (auto simp: indicator_def) + ultimately show "?I \<in> borel_measurable M" by auto +qed + +lemma borel_measurable_subalgebra: + assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N" + shows "f \<in> borel_measurable M" + using assms unfolding measurable_def by auto + +lemma borel_measurable_restrict_space_iff_ereal: + fixes f :: "'a \<Rightarrow> ereal" + assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" + shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> + (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M" + by (subst measurable_restrict_space_iff) + (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong) + +lemma borel_measurable_restrict_space_iff_ennreal: + fixes f :: "'a \<Rightarrow> ennreal" + assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" + shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> + (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M" + by (subst measurable_restrict_space_iff) + (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong) + +lemma borel_measurable_restrict_space_iff: + fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" + assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M" + shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow> + (\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> borel_measurable M" + by (subst measurable_restrict_space_iff) + (auto simp: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps + cong del: if_weak_cong) + +lemma cbox_borel[measurable]: "cbox a b \<in> sets borel" + by (auto intro: borel_closed) + +lemma box_borel[measurable]: "box a b \<in> sets borel" + by (auto intro: borel_open) + +lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel" + by (auto intro: borel_closed dest!: compact_imp_closed) + +lemma borel_sigma_sets_subset: + "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel" + using sets.sigma_sets_subset[of A borel] by simp + +lemma borel_eq_sigmaI1: + fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set" + assumes borel_eq: "borel = sigma UNIV X" + assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))" + assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel" + shows "borel = sigma UNIV (F ` A)" + unfolding borel_def +proof (intro sigma_eqI antisym) + have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel" + unfolding borel_def by simp + also have "\<dots> = sigma_sets UNIV X" + unfolding borel_eq by simp + also have "\<dots> \<subseteq> sigma_sets UNIV (F`A)" + using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto + finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (F`A)" . + show "sigma_sets UNIV (F`A) \<subseteq> sigma_sets UNIV {S. open S}" + unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto +qed auto + +lemma borel_eq_sigmaI2: + fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" + and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)" + assumes X: "\<And>i j. (i, j) \<in> B \<Longrightarrow> G i j \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))" + assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel" + shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)" + using assms + by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto + +lemma borel_eq_sigmaI3: + fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set" + assumes borel_eq: "borel = sigma UNIV X" + assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))" + assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel" + shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)" + using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto + +lemma borel_eq_sigmaI4: + fixes F :: "'i \<Rightarrow> 'a::topological_space set" + and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)" + assumes X: "\<And>i j. (i, j) \<in> A \<Longrightarrow> G i j \<in> sets (sigma UNIV (range F))" + assumes F: "\<And>i. F i \<in> sets borel" + shows "borel = sigma UNIV (range F)" + using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto + +lemma borel_eq_sigmaI5: + fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV (range G)" + assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))" + assumes F: "\<And>i j. F i j \<in> sets borel" + shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))" + using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto + +lemma second_countable_borel_measurable: + fixes X :: "'a::second_countable_topology set set" + assumes eq: "open = generate_topology X" + shows "borel = sigma UNIV X" + unfolding borel_def +proof (intro sigma_eqI sigma_sets_eqI) + interpret X: sigma_algebra UNIV "sigma_sets UNIV X" + by (rule sigma_algebra_sigma_sets) simp + + fix S :: "'a set" assume "S \<in> Collect open" + then have "generate_topology X S" + by (auto simp: eq) + then show "S \<in> sigma_sets UNIV X" + proof induction + case (UN K) + then have K: "\<And>k. k \<in> K \<Longrightarrow> open k" + unfolding eq by auto + from ex_countable_basis obtain B :: "'a set set" where + B: "\<And>b. b \<in> B \<Longrightarrow> open b" "\<And>X. open X \<Longrightarrow> \<exists>b\<subseteq>B. (\<Union>b) = X" and "countable B" + by (auto simp: topological_basis_def) + from B(2)[OF K] obtain m where m: "\<And>k. k \<in> K \<Longrightarrow> m k \<subseteq> B" "\<And>k. k \<in> K \<Longrightarrow> (\<Union>m k) = k" + by metis + define U where "U = (\<Union>k\<in>K. m k)" + with m have "countable U" + by (intro countable_subset[OF _ \<open>countable B\<close>]) auto + have "\<Union>U = (\<Union>A\<in>U. A)" by simp + also have "\<dots> = \<Union>K" + unfolding U_def UN_simps by (simp add: m) + finally have "\<Union>U = \<Union>K" . + + have "\<forall>b\<in>U. \<exists>k\<in>K. b \<subseteq> k" + using m by (auto simp: U_def) + then obtain u where u: "\<And>b. b \<in> U \<Longrightarrow> u b \<in> K" and "\<And>b. b \<in> U \<Longrightarrow> b \<subseteq> u b" + by metis + then have "(\<Union>b\<in>U. u b) \<subseteq> \<Union>K" "\<Union>U \<subseteq> (\<Union>b\<in>U. u b)" + by auto + then have "\<Union>K = (\<Union>b\<in>U. u b)" + unfolding \<open>\<Union>U = \<Union>K\<close> by auto + also have "\<dots> \<in> sigma_sets UNIV X" + using u UN by (intro X.countable_UN' \<open>countable U\<close>) auto + finally show "\<Union>K \<in> sigma_sets UNIV X" . + qed auto +qed (auto simp: eq intro: generate_topology.Basis) + +lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" + unfolding borel_def +proof (intro sigma_eqI sigma_sets_eqI, safe) + fix x :: "'a set" assume "open x" + hence "x = UNIV - (UNIV - x)" by auto + also have "\<dots> \<in> sigma_sets UNIV (Collect closed)" + by (force intro: sigma_sets.Compl simp: \<open>open x\<close>) + finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp +next + fix x :: "'a set" assume "closed x" + hence "x = UNIV - (UNIV - x)" by auto + also have "\<dots> \<in> sigma_sets UNIV (Collect open)" + by (force intro: sigma_sets.Compl simp: \<open>closed x\<close>) + finally show "x \<in> sigma_sets UNIV (Collect open)" by simp +qed simp_all + +lemma borel_eq_countable_basis: + fixes B::"'a::topological_space set set" + assumes "countable B" + assumes "topological_basis B" + shows "borel = sigma UNIV B" + unfolding borel_def +proof (intro sigma_eqI sigma_sets_eqI, safe) + interpret countable_basis using assms by unfold_locales + fix X::"'a set" assume "open X" + from open_countable_basisE[OF this] guess B' . note B' = this + then show "X \<in> sigma_sets UNIV B" + by (blast intro: sigma_sets_UNION \<open>countable B\<close> countable_subset) +next + fix b assume "b \<in> B" + hence "open b" by (rule topological_basis_open[OF assms(2)]) + thus "b \<in> sigma_sets UNIV (Collect open)" by auto +qed simp_all + +lemma borel_measurable_continuous_on_restrict: + fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" + assumes f: "continuous_on A f" + shows "f \<in> borel_measurable (restrict_space borel A)" +proof (rule borel_measurableI) + fix S :: "'b set" assume "open S" + with f obtain T where "f -` S \<inter> A = T \<inter> A" "open T" + by (metis continuous_on_open_invariant) + then show "f -` S \<inter> space (restrict_space borel A) \<in> sets (restrict_space borel A)" + by (force simp add: sets_restrict_space space_restrict_space) +qed + +lemma borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel" + by (drule borel_measurable_continuous_on_restrict) simp + +lemma borel_measurable_continuous_on_if: + "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on (- A) g \<Longrightarrow> + (\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel" + by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq + intro!: borel_measurable_continuous_on_restrict) + +lemma borel_measurable_continuous_countable_exceptions: + fixes f :: "'a::t1_space \<Rightarrow> 'b::topological_space" + assumes X: "countable X" + assumes "continuous_on (- X) f" + shows "f \<in> borel_measurable borel" +proof (rule measurable_discrete_difference[OF _ X]) + have "X \<in> sets borel" + by (rule sets.countable[OF _ X]) auto + then show "(\<lambda>x. if x \<in> X then undefined else f x) \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on_if assms continuous_intros) +qed auto + +lemma borel_measurable_continuous_on: + assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M" + shows "(\<lambda>x. f (g x)) \<in> borel_measurable M" + using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def) + +lemma borel_measurable_continuous_on_indicator: + fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" + shows "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel" + by (subst borel_measurable_restrict_space_iff[symmetric]) + (auto intro: borel_measurable_continuous_on_restrict) + +lemma borel_measurable_Pair[measurable (raw)]: + fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology" + assumes f[measurable]: "f \<in> borel_measurable M" + assumes g[measurable]: "g \<in> borel_measurable M" + shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M" +proof (subst borel_eq_countable_basis) + let ?B = "SOME B::'b set set. countable B \<and> topological_basis B" + let ?C = "SOME B::'c set set. countable B \<and> topological_basis B" + let ?P = "(\<lambda>(b, c). b \<times> c) ` (?B \<times> ?C)" + show "countable ?P" "topological_basis ?P" + by (auto intro!: countable_basis topological_basis_prod is_basis) + + show "(\<lambda>x. (f x, g x)) \<in> measurable M (sigma UNIV ?P)" + proof (rule measurable_measure_of) + fix S assume "S \<in> ?P" + then obtain b c where "b \<in> ?B" "c \<in> ?C" and S: "S = b \<times> c" by auto + then have borel: "open b" "open c" + by (auto intro: is_basis topological_basis_open) + have "(\<lambda>x. (f x, g x)) -` S \<inter> space M = (f -` b \<inter> space M) \<inter> (g -` c \<inter> space M)" + unfolding S by auto + also have "\<dots> \<in> sets M" + using borel by simp + finally show "(\<lambda>x. (f x, g x)) -` S \<inter> space M \<in> sets M" . + qed auto +qed + +lemma borel_measurable_continuous_Pair: + fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology" + assumes [measurable]: "f \<in> borel_measurable M" + assumes [measurable]: "g \<in> borel_measurable M" + assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))" + shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M" +proof - + have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto + show ?thesis + unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto +qed + +subsection \<open>Borel spaces on order topologies\<close> + +lemma [measurable]: + fixes a b :: "'a::linorder_topology" + shows lessThan_borel: "{..< a} \<in> sets borel" + and greaterThan_borel: "{a <..} \<in> sets borel" + and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel" + and atMost_borel: "{..a} \<in> sets borel" + and atLeast_borel: "{a..} \<in> sets borel" + and atLeastAtMost_borel: "{a..b} \<in> sets borel" + and greaterThanAtMost_borel: "{a<..b} \<in> sets borel" + and atLeastLessThan_borel: "{a..<b} \<in> sets borel" + unfolding greaterThanAtMost_def atLeastLessThan_def + by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan + closed_atMost closed_atLeast closed_atLeastAtMost)+ + +lemma borel_Iio: + "borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)" + unfolding second_countable_borel_measurable[OF open_generated_order] +proof (intro sigma_eqI sigma_sets_eqI) + from countable_dense_setE guess D :: "'a set" . note D = this + + interpret L: sigma_algebra UNIV "sigma_sets UNIV (range lessThan)" + by (rule sigma_algebra_sigma_sets) simp + + fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan" + then obtain y where "A = {y <..} \<or> A = {..< y}" + by blast + then show "A \<in> sigma_sets UNIV (range lessThan)" + proof + assume A: "A = {y <..}" + show ?thesis + proof cases + assume "\<forall>x>y. \<exists>d. y < d \<and> d < x" + with D(2)[of "{y <..< x}" for x] have "\<forall>x>y. \<exists>d\<in>D. y < d \<and> d < x" + by (auto simp: set_eq_iff) + then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. y < d}. {..< d})" + by (auto simp: A) (metis less_asym) + also have "\<dots> \<in> sigma_sets UNIV (range lessThan)" + using D(1) by (intro L.Diff L.top L.countable_INT'') auto + finally show ?thesis . + next + assume "\<not> (\<forall>x>y. \<exists>d. y < d \<and> d < x)" + then obtain x where "y < x" "\<And>d. y < d \<Longrightarrow> \<not> d < x" + by auto + then have "A = UNIV - {..< x}" + unfolding A by (auto simp: not_less[symmetric]) + also have "\<dots> \<in> sigma_sets UNIV (range lessThan)" + by auto + finally show ?thesis . + qed + qed auto +qed auto + +lemma borel_Ioi: + "borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)" + unfolding second_countable_borel_measurable[OF open_generated_order] +proof (intro sigma_eqI sigma_sets_eqI) + from countable_dense_setE guess D :: "'a set" . note D = this + + interpret L: sigma_algebra UNIV "sigma_sets UNIV (range greaterThan)" + by (rule sigma_algebra_sigma_sets) simp + + fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan" + then obtain y where "A = {y <..} \<or> A = {..< y}" + by blast + then show "A \<in> sigma_sets UNIV (range greaterThan)" + proof + assume A: "A = {..< y}" + show ?thesis + proof cases + assume "\<forall>x<y. \<exists>d. x < d \<and> d < y" + with D(2)[of "{x <..< y}" for x] have "\<forall>x<y. \<exists>d\<in>D. x < d \<and> d < y" + by (auto simp: set_eq_iff) + then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. d < y}. {d <..})" + by (auto simp: A) (metis less_asym) + also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)" + using D(1) by (intro L.Diff L.top L.countable_INT'') auto + finally show ?thesis . + next + assume "\<not> (\<forall>x<y. \<exists>d. x < d \<and> d < y)" + then obtain x where "x < y" "\<And>d. y > d \<Longrightarrow> x \<ge> d" + by (auto simp: not_less[symmetric]) + then have "A = UNIV - {x <..}" + unfolding A Compl_eq_Diff_UNIV[symmetric] by auto + also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)" + by auto + finally show ?thesis . + qed + qed auto +qed auto + +lemma borel_measurableI_less: + fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}" + shows "(\<And>y. {x\<in>space M. f x < y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M" + unfolding borel_Iio + by (rule measurable_measure_of) (auto simp: Int_def conj_commute) + +lemma borel_measurableI_greater: + fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}" + shows "(\<And>y. {x\<in>space M. y < f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M" + unfolding borel_Ioi + by (rule measurable_measure_of) (auto simp: Int_def conj_commute) + +lemma borel_measurableI_le: + fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}" + shows "(\<And>y. {x\<in>space M. f x \<le> y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M" + by (rule borel_measurableI_greater) (auto simp: not_le[symmetric]) + +lemma borel_measurableI_ge: + fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}" + shows "(\<And>y. {x\<in>space M. y \<le> f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M" + by (rule borel_measurableI_less) (auto simp: not_le[symmetric]) + +lemma borel_measurable_less[measurable]: + fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}" + assumes "f \<in> borel_measurable M" + assumes "g \<in> borel_measurable M" + shows "{w \<in> space M. f w < g w} \<in> sets M" +proof - + have "{w \<in> space M. f w < g w} = (\<lambda>x. (f x, g x)) -` {x. fst x < snd x} \<inter> space M" + by auto + also have "\<dots> \<in> sets M" + by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less] + continuous_intros) + finally show ?thesis . +qed + +lemma + fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}" + assumes f[measurable]: "f \<in> borel_measurable M" + assumes g[measurable]: "g \<in> borel_measurable M" + shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M" + and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M" + and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M" + unfolding eq_iff not_less[symmetric] + by measurable + +lemma borel_measurable_SUP[measurable (raw)]: + fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}" + assumes [simp]: "countable I" + assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" + shows "(\<lambda>x. SUP i:I. F i x) \<in> borel_measurable M" + by (rule borel_measurableI_greater) (simp add: less_SUP_iff) + +lemma borel_measurable_INF[measurable (raw)]: + fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}" + assumes [simp]: "countable I" + assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" + shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M" + by (rule borel_measurableI_less) (simp add: INF_less_iff) + +lemma borel_measurable_cSUP[measurable (raw)]: + fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" + assumes [simp]: "countable I" + assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" + assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_above ((\<lambda>i. F i x) ` I)" + shows "(\<lambda>x. SUP i:I. F i x) \<in> borel_measurable M" +proof cases + assume "I = {}" then show ?thesis + unfolding \<open>I = {}\<close> image_empty by simp +next + assume "I \<noteq> {}" + show ?thesis + proof (rule borel_measurableI_le) + fix y + have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} \<in> sets M" + by measurable + also have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} = {x \<in> space M. (SUP i:I. F i x) \<le> y}" + by (simp add: cSUP_le_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong) + finally show "{x \<in> space M. (SUP i:I. F i x) \<le> y} \<in> sets M" . + qed +qed + +lemma borel_measurable_cINF[measurable (raw)]: + fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" + assumes [simp]: "countable I" + assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M" + assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_below ((\<lambda>i. F i x) ` I)" + shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M" +proof cases + assume "I = {}" then show ?thesis + unfolding \<open>I = {}\<close> image_empty by simp +next + assume "I \<noteq> {}" + show ?thesis + proof (rule borel_measurableI_ge) + fix y + have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} \<in> sets M" + by measurable + also have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} = {x \<in> space M. y \<le> (INF i:I. F i x)}" + by (simp add: le_cINF_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong) + finally show "{x \<in> space M. y \<le> (INF i:I. F i x)} \<in> sets M" . + qed +qed + +lemma borel_measurable_lfp[consumes 1, case_names continuity step]: + fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})" + assumes "sup_continuous F" + assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M" + shows "lfp F \<in> borel_measurable M" +proof - + { fix i have "((F ^^ i) bot) \<in> borel_measurable M" + by (induct i) (auto intro!: *) } + then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> borel_measurable M" + by measurable + also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = (SUP i. (F ^^ i) bot)" + by auto + also have "(SUP i. (F ^^ i) bot) = lfp F" + by (rule sup_continuous_lfp[symmetric]) fact + finally show ?thesis . +qed + +lemma borel_measurable_gfp[consumes 1, case_names continuity step]: + fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})" + assumes "inf_continuous F" + assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M" + shows "gfp F \<in> borel_measurable M" +proof - + { fix i have "((F ^^ i) top) \<in> borel_measurable M" + by (induct i) (auto intro!: * simp: bot_fun_def) } + then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> borel_measurable M" + by measurable + also have "(\<lambda>x. INF i. (F ^^ i) top x) = (INF i. (F ^^ i) top)" + by auto + also have "\<dots> = gfp F" + by (rule inf_continuous_gfp[symmetric]) fact + finally show ?thesis . +qed + +lemma borel_measurable_max[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M" + by (rule borel_measurableI_less) simp + +lemma borel_measurable_min[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M" + by (rule borel_measurableI_greater) simp + +lemma borel_measurable_Min[measurable (raw)]: + "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Min ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M" +proof (induct I rule: finite_induct) + case (insert i I) then show ?case + by (cases "I = {}") auto +qed auto + +lemma borel_measurable_Max[measurable (raw)]: + "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Max ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M" +proof (induct I rule: finite_induct) + case (insert i I) then show ?case + by (cases "I = {}") auto +qed auto + +lemma borel_measurable_sup[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M" + unfolding sup_max by measurable + +lemma borel_measurable_inf[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M" + unfolding inf_min by measurable + +lemma [measurable (raw)]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}" + assumes "\<And>i. f i \<in> borel_measurable M" + shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M" + and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M" + unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto + +lemma measurable_convergent[measurable (raw)]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}" + assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "Measurable.pred M (\<lambda>x. convergent (\<lambda>i. f i x))" + unfolding convergent_ereal by measurable + +lemma sets_Collect_convergent[measurable]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}" + assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M" + by measurable + +lemma borel_measurable_lim[measurable (raw)]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}" + assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" +proof - + have "\<And>x. lim (\<lambda>i. f i x) = (if convergent (\<lambda>i. f i x) then limsup (\<lambda>i. f i x) else (THE i. False))" + by (simp add: lim_def convergent_def convergent_limsup_cl) + then show ?thesis + by simp +qed + +lemma borel_measurable_LIMSEQ_order: + fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}" + assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x" + and u: "\<And>i. u i \<in> borel_measurable M" + shows "u' \<in> borel_measurable M" +proof - + have "\<And>x. x \<in> space M \<Longrightarrow> u' x = liminf (\<lambda>n. u n x)" + using u' by (simp add: lim_imp_Liminf[symmetric]) + with u show ?thesis by (simp cong: measurable_cong) +qed + +subsection \<open>Borel spaces on topological monoids\<close> + +lemma borel_measurable_add[measurable (raw)]: + fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, topological_monoid_add}" + assumes f: "f \<in> borel_measurable M" + assumes g: "g \<in> borel_measurable M" + shows "(\<lambda>x. f x + g x) \<in> borel_measurable M" + using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) + +lemma borel_measurable_setsum[measurable (raw)]: + fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, topological_comm_monoid_add}" + assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" + shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" +proof cases + assume "finite S" + thus ?thesis using assms by induct auto +qed simp + +lemma borel_measurable_suminf_order[measurable (raw)]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}" + assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M" + unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp + +subsection \<open>Borel spaces on Euclidean spaces\<close> + +lemma borel_measurable_inner[measurable (raw)]: + fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_inner}" + assumes "f \<in> borel_measurable M" + assumes "g \<in> borel_measurable M" + shows "(\<lambda>x. f x \<bullet> g x) \<in> borel_measurable M" + using assms + by (rule borel_measurable_continuous_Pair) (intro continuous_intros) + +notation + eucl_less (infix "<e" 50) + +lemma box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}" + and box_co: "{x. a \<le> x \<and> x <e b} = {a..} \<inter> {x. x <e b}" + by auto + +lemma eucl_ivals[measurable]: + fixes a b :: "'a::ordered_euclidean_space" + shows "{x. x <e a} \<in> sets borel" + and "{x. a <e x} \<in> sets borel" + and "{..a} \<in> sets borel" + and "{a..} \<in> sets borel" + and "{a..b} \<in> sets borel" + and "{x. a <e x \<and> x \<le> b} \<in> sets borel" + and "{x. a \<le> x \<and> x <e b} \<in> sets borel" + unfolding box_oc box_co + by (auto intro: borel_open borel_closed) + +lemma + fixes i :: "'a::{second_countable_topology, real_inner}" + shows hafspace_less_borel: "{x. a < x \<bullet> i} \<in> sets borel" + and hafspace_greater_borel: "{x. x \<bullet> i < a} \<in> sets borel" + and hafspace_less_eq_borel: "{x. a \<le> x \<bullet> i} \<in> sets borel" + and hafspace_greater_eq_borel: "{x. x \<bullet> i \<le> a} \<in> sets borel" + by simp_all + +lemma borel_eq_box: + "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a :: euclidean_space set))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI1[OF borel_def]) + fix M :: "'a set" assume "M \<in> {S. open S}" + then have "open M" by simp + show "M \<in> ?SIGMA" + apply (subst open_UNION_box[OF \<open>open M\<close>]) + apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect) + apply (auto intro: countable_rat) + done +qed (auto simp: box_def) + +lemma halfspace_gt_in_halfspace: + assumes i: "i \<in> A" + shows "{x::'a. a < x \<bullet> i} \<in> + sigma_sets UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))" + (is "?set \<in> ?SIGMA") +proof - + interpret sigma_algebra UNIV ?SIGMA + by (intro sigma_algebra_sigma_sets) simp_all + have *: "?set = (\<Union>n. UNIV - {x::'a. x \<bullet> i < a + 1 / real (Suc n)})" + proof (safe, simp_all add: not_less del: of_nat_Suc) + fix x :: 'a assume "a < x \<bullet> i" + with reals_Archimedean[of "x \<bullet> i - a"] + obtain n where "a + 1 / real (Suc n) < x \<bullet> i" + by (auto simp: field_simps) + then show "\<exists>n. a + 1 / real (Suc n) \<le> x \<bullet> i" + by (blast intro: less_imp_le) + next + fix x n + have "a < a + 1 / real (Suc n)" by auto + also assume "\<dots> \<le> x" + finally show "a < x" . + qed + show "?set \<in> ?SIGMA" unfolding * + by (auto intro!: Diff sigma_sets_Inter i) +qed + +lemma borel_eq_halfspace_less: + "borel = sigma UNIV ((\<lambda>(a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_box]) + fix a b :: 'a + have "box a b = {x\<in>space ?SIGMA. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}" + by (auto simp: box_def) + also have "\<dots> \<in> sets ?SIGMA" + by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const) + (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat) + finally show "box a b \<in> sets ?SIGMA" . +qed auto + +lemma borel_eq_halfspace_le: + "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i \<le> a}) ` (UNIV \<times> Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) + fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis" + then have i: "i \<in> Basis" by auto + have *: "{x::'a. x\<bullet>i < a} = (\<Union>n. {x. x\<bullet>i \<le> a - 1/real (Suc n)})" + proof (safe, simp_all del: of_nat_Suc) + fix x::'a assume *: "x\<bullet>i < a" + with reals_Archimedean[of "a - x\<bullet>i"] + obtain n where "x \<bullet> i < a - 1 / (real (Suc n))" + by (auto simp: field_simps) + then show "\<exists>n. x \<bullet> i \<le> a - 1 / (real (Suc n))" + by (blast intro: less_imp_le) + next + fix x::'a and n + assume "x\<bullet>i \<le> a - 1 / real (Suc n)" + also have "\<dots> < a" by auto + finally show "x\<bullet>i < a" . + qed + show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding * + by (intro sets.countable_UN) (auto intro: i) +qed auto + +lemma borel_eq_halfspace_ge: + "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) + fix a :: real and i :: 'a assume i: "(a, i) \<in> UNIV \<times> Basis" + have *: "{x::'a. x\<bullet>i < a} = space ?SIGMA - {x::'a. a \<le> x\<bullet>i}" by auto + show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding * + using i by (intro sets.compl_sets) auto +qed auto + +lemma borel_eq_halfspace_greater: + "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) + fix a :: real and i :: 'a assume "(a, i) \<in> (UNIV \<times> Basis)" + then have i: "i \<in> Basis" by auto + have *: "{x::'a. x\<bullet>i \<le> a} = space ?SIGMA - {x::'a. a < x\<bullet>i}" by auto + show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding * + by (intro sets.compl_sets) (auto intro: i) +qed auto + +lemma borel_eq_atMost: + "borel = sigma UNIV (range (\<lambda>a. {..a::'a::ordered_euclidean_space}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) + fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis" + then have "i \<in> Basis" by auto + then have *: "{x::'a. x\<bullet>i \<le> a} = (\<Union>k::nat. {.. (\<Sum>n\<in>Basis. (if n = i then a else real k)*\<^sub>R n)})" + proof (safe, simp_all add: eucl_le[where 'a='a] split: if_split_asm) + fix x :: 'a + from real_arch_simple[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"] guess k::nat .. + then have "\<And>i. i \<in> Basis \<Longrightarrow> x\<bullet>i \<le> real k" + by (subst (asm) Max_le_iff) auto + then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia \<le> real k" + by (auto intro!: exI[of _ k]) + qed + show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding * + by (intro sets.countable_UN) auto +qed auto + +lemma borel_eq_greaterThan: + "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. a <e x}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) + fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis" + then have i: "i \<in> Basis" by auto + have "{x::'a. x\<bullet>i \<le> a} = UNIV - {x::'a. a < x\<bullet>i}" by auto + also have *: "{x::'a. a < x\<bullet>i} = + (\<Union>k::nat. {x. (\<Sum>n\<in>Basis. (if n = i then a else -real k) *\<^sub>R n) <e x})" using i + proof (safe, simp_all add: eucl_less_def split: if_split_asm) + fix x :: 'a + from reals_Archimedean2[of "Max ((\<lambda>i. -x\<bullet>i)`Basis)"] + guess k::nat .. note k = this + { fix i :: 'a assume "i \<in> Basis" + then have "-x\<bullet>i < real k" + using k by (subst (asm) Max_less_iff) auto + then have "- real k < x\<bullet>i" by simp } + then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> -real k < x \<bullet> ia" + by (auto intro!: exI[of _ k]) + qed + finally show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" + apply (simp only:) + apply (intro sets.countable_UN sets.Diff) + apply (auto intro: sigma_sets_top) + done +qed auto + +lemma borel_eq_lessThan: + "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. x <e a}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge]) + fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis" + then have i: "i \<in> Basis" by auto + have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto + also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" using \<open>i\<in> Basis\<close> + proof (safe, simp_all add: eucl_less_def split: if_split_asm) + fix x :: 'a + from reals_Archimedean2[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"] + guess k::nat .. note k = this + { fix i :: 'a assume "i \<in> Basis" + then have "x\<bullet>i < real k" + using k by (subst (asm) Max_less_iff) auto + then have "x\<bullet>i < real k" by simp } + then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia < real k" + by (auto intro!: exI[of _ k]) + qed + finally show "{x. a \<le> x\<bullet>i} \<in> ?SIGMA" + apply (simp only:) + apply (intro sets.countable_UN sets.Diff) + apply (auto intro: sigma_sets_top ) + done +qed auto + +lemma borel_eq_atLeastAtMost: + "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) + fix a::'a + have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})" + proof (safe, simp_all add: eucl_le[where 'a='a]) + fix x :: 'a + from real_arch_simple[of "Max ((\<lambda>i. - x\<bullet>i)`Basis)"] + guess k::nat .. note k = this + { fix i :: 'a assume "i \<in> Basis" + with k have "- x\<bullet>i \<le> real k" + by (subst (asm) Max_le_iff) (auto simp: field_simps) + then have "- real k \<le> x\<bullet>i" by simp } + then show "\<exists>n::nat. \<forall>i\<in>Basis. - real n \<le> x \<bullet> i" + by (auto intro!: exI[of _ k]) + qed + show "{..a} \<in> ?SIGMA" unfolding * + by (intro sets.countable_UN) + (auto intro!: sigma_sets_top) +qed auto + +lemma borel_set_induct[consumes 1, case_names empty interval compl union]: + assumes "A \<in> sets borel" + assumes empty: "P {}" and int: "\<And>a b. a \<le> b \<Longrightarrow> P {a..b}" and compl: "\<And>A. A \<in> sets borel \<Longrightarrow> P A \<Longrightarrow> P (-A)" and + un: "\<And>f. disjoint_family f \<Longrightarrow> (\<And>i. f i \<in> sets borel) \<Longrightarrow> (\<And>i. P (f i)) \<Longrightarrow> P (\<Union>i::nat. f i)" + shows "P (A::real set)" +proof- + let ?G = "range (\<lambda>(a,b). {a..b::real})" + have "Int_stable ?G" "?G \<subseteq> Pow UNIV" "A \<in> sigma_sets UNIV ?G" + using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def) + thus ?thesis + proof (induction rule: sigma_sets_induct_disjoint) + case (union f) + from union.hyps(2) have "\<And>i. f i \<in> sets borel" by (auto simp: borel_eq_atLeastAtMost) + with union show ?case by (auto intro: un) + next + case (basic A) + then obtain a b where "A = {a .. b}" by auto + then show ?case + by (cases "a \<le> b") (auto intro: int empty) + qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost) +qed + +lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))" +proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) + fix i :: real + have "{..i} = (\<Union>j::nat. {-j <.. i})" + by (auto simp: minus_less_iff reals_Archimedean2) + also have "\<dots> \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" + by (intro sets.countable_nat_UN) auto + finally show "{..i} \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" . +qed simp + +lemma eucl_lessThan: "{x::real. x <e a} = lessThan a" + by (simp add: eucl_less_def lessThan_def) + +lemma borel_eq_atLeastLessThan: + "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) + have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto + fix x :: real + have "{..<x} = (\<Union>i::nat. {-real i ..< x})" + by (auto simp: move_uminus real_arch_simple) + then show "{y. y <e x} \<in> ?SIGMA" + by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan) +qed auto + +lemma borel_measurable_halfspacesI: + fixes f :: "'a \<Rightarrow> 'c::euclidean_space" + assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))" + and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M" + shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a::real. S a i \<in> sets M)" +proof safe + fix a :: real and i :: 'b assume i: "i \<in> Basis" and f: "f \<in> borel_measurable M" + then show "S a i \<in> sets M" unfolding assms + by (auto intro!: measurable_sets simp: assms(1)) +next + assume a: "\<forall>i\<in>Basis. \<forall>a. S a i \<in> sets M" + then show "f \<in> borel_measurable M" + by (auto intro!: measurable_measure_of simp: S_eq F) +qed + +lemma borel_measurable_iff_halfspace_le: + fixes f :: "'a \<Rightarrow> 'c::euclidean_space" + shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i \<le> a} \<in> sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto + +lemma borel_measurable_iff_halfspace_less: + fixes f :: "'a \<Rightarrow> 'c::euclidean_space" + shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i < a} \<in> sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto + +lemma borel_measurable_iff_halfspace_ge: + fixes f :: "'a \<Rightarrow> 'c::euclidean_space" + shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a \<le> f w \<bullet> i} \<in> sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto + +lemma borel_measurable_iff_halfspace_greater: + fixes f :: "'a \<Rightarrow> 'c::euclidean_space" + shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a < f w \<bullet> i} \<in> sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto + +lemma borel_measurable_iff_le: + "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)" + using borel_measurable_iff_halfspace_le[where 'c=real] by simp + +lemma borel_measurable_iff_less: + "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)" + using borel_measurable_iff_halfspace_less[where 'c=real] by simp + +lemma borel_measurable_iff_ge: + "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)" + using borel_measurable_iff_halfspace_ge[where 'c=real] + by simp + +lemma borel_measurable_iff_greater: + "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)" + using borel_measurable_iff_halfspace_greater[where 'c=real] by simp + +lemma borel_measurable_euclidean_space: + fixes f :: "'a \<Rightarrow> 'c::euclidean_space" + shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)" +proof safe + assume f: "\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M" + then show "f \<in> borel_measurable M" + by (subst borel_measurable_iff_halfspace_le) auto +qed auto + +subsection "Borel measurable operators" + +lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel" + by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"]) + (auto intro!: continuous_on_sgn continuous_on_id) + +lemma borel_measurable_uminus[measurable (raw)]: + fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}" + assumes g: "g \<in> borel_measurable M" + shows "(\<lambda>x. - g x) \<in> borel_measurable M" + by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros) + +lemma borel_measurable_diff[measurable (raw)]: + fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}" + assumes f: "f \<in> borel_measurable M" + assumes g: "g \<in> borel_measurable M" + shows "(\<lambda>x. f x - g x) \<in> borel_measurable M" + using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def) + +lemma borel_measurable_times[measurable (raw)]: + fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}" + assumes f: "f \<in> borel_measurable M" + assumes g: "g \<in> borel_measurable M" + shows "(\<lambda>x. f x * g x) \<in> borel_measurable M" + using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) + +lemma borel_measurable_setprod[measurable (raw)]: + fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, real_normed_field}" + assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" + shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" +proof cases + assume "finite S" + thus ?thesis using assms by induct auto +qed simp + +lemma borel_measurable_dist[measurable (raw)]: + fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}" + assumes f: "f \<in> borel_measurable M" + assumes g: "g \<in> borel_measurable M" + shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M" + using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) + +lemma borel_measurable_scaleR[measurable (raw)]: + fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}" + assumes f: "f \<in> borel_measurable M" + assumes g: "g \<in> borel_measurable M" + shows "(\<lambda>x. f x *\<^sub>R g x) \<in> borel_measurable M" + using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) + +lemma affine_borel_measurable_vector: + fixes f :: "'a \<Rightarrow> 'x::real_normed_vector" + assumes "f \<in> borel_measurable M" + shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M" +proof (rule borel_measurableI) + fix S :: "'x set" assume "open S" + show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M" + proof cases + assume "b \<noteq> 0" + with \<open>open S\<close> have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") + using open_affinity [of S "inverse b" "- a /\<^sub>R b"] + by (auto simp: algebra_simps) + hence "?S \<in> sets borel" by auto + moreover + from \<open>b \<noteq> 0\<close> have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S" + apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) + ultimately show ?thesis using assms unfolding in_borel_measurable_borel + by auto + qed simp +qed + +lemma borel_measurable_const_scaleR[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. b *\<^sub>R f x ::'a::real_normed_vector) \<in> borel_measurable M" + using affine_borel_measurable_vector[of f M 0 b] by simp + +lemma borel_measurable_const_add[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. a + f x ::'a::real_normed_vector) \<in> borel_measurable M" + using affine_borel_measurable_vector[of f M a 1] by simp + +lemma borel_measurable_inverse[measurable (raw)]: + fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra" + assumes f: "f \<in> borel_measurable M" + shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M" + apply (rule measurable_compose[OF f]) + apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) + apply (auto intro!: continuous_on_inverse continuous_on_id) + done + +lemma borel_measurable_divide[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> + (\<lambda>x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \<in> borel_measurable M" + by (simp add: divide_inverse) + +lemma borel_measurable_abs[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M" + unfolding abs_real_def by simp + +lemma borel_measurable_nth[measurable (raw)]: + "(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel" + by (simp add: cart_eq_inner_axis) + +lemma convex_measurable: + fixes A :: "'a :: euclidean_space set" + shows "X \<in> borel_measurable M \<Longrightarrow> X ` space M \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> convex_on A q \<Longrightarrow> + (\<lambda>x. q (X x)) \<in> borel_measurable M" + by (rule measurable_compose[where f=X and N="restrict_space borel A"]) + (auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2) + +lemma borel_measurable_ln[measurable (raw)]: + assumes f: "f \<in> borel_measurable M" + shows "(\<lambda>x. ln (f x :: real)) \<in> borel_measurable M" + apply (rule measurable_compose[OF f]) + apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) + apply (auto intro!: continuous_on_ln continuous_on_id) + done + +lemma borel_measurable_log[measurable (raw)]: + "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M" + unfolding log_def by auto + +lemma borel_measurable_exp[measurable]: + "(exp::'a::{real_normed_field,banach}\<Rightarrow>'a) \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp) + +lemma measurable_real_floor[measurable]: + "(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" +proof - + have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real_of_int a \<le> x \<and> x < real_of_int (a + 1))" + by (auto intro: floor_eq2) + then show ?thesis + by (auto simp: vimage_def measurable_count_space_eq2_countable) +qed + +lemma measurable_real_ceiling[measurable]: + "(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)" + unfolding ceiling_def[abs_def] by simp + +lemma borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel" + by simp + +lemma borel_measurable_root [measurable]: "root n \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_power [measurable (raw)]: + fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}" + assumes f: "f \<in> borel_measurable M" + shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M" + by (intro borel_measurable_continuous_on [OF _ f] continuous_intros) + +lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + +lemma borel_measurable_complex_iff: + "f \<in> borel_measurable M \<longleftrightarrow> + (\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M" + apply auto + apply (subst fun_complex_eq) + apply (intro borel_measurable_add) + apply auto + done + +subsection "Borel space on the extended reals" + +lemma borel_measurable_ereal[measurable (raw)]: + assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" + using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id) + +lemma borel_measurable_real_of_ereal[measurable (raw)]: + fixes f :: "'a \<Rightarrow> ereal" + assumes f: "f \<in> borel_measurable M" + shows "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" + apply (rule measurable_compose[OF f]) + apply (rule borel_measurable_continuous_countable_exceptions[of "{\<infinity>, -\<infinity> }"]) + apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV) + done + +lemma borel_measurable_ereal_cases: + fixes f :: "'a \<Rightarrow> ereal" + assumes f: "f \<in> borel_measurable M" + assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x)))) \<in> borel_measurable M" + shows "(\<lambda>x. H (f x)) \<in> borel_measurable M" +proof - + let ?F = "\<lambda>x. if f x = \<infinity> then H \<infinity> else if f x = - \<infinity> then H (-\<infinity>) else H (ereal (real_of_ereal (f x)))" + { fix x have "H (f x) = ?F x" by (cases "f x") auto } + with f H show ?thesis by simp +qed + +lemma + fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M" + shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" + and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M" + and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M" + by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) + +lemma borel_measurable_uminus_eq_ereal[simp]: + "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r") +proof + assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp +qed auto + +lemma set_Collect_ereal2: + fixes f g :: "'a \<Rightarrow> ereal" + assumes f: "f \<in> borel_measurable M" + assumes g: "g \<in> borel_measurable M" + assumes H: "{x \<in> space M. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))} \<in> sets M" + "{x \<in> space borel. H (-\<infinity>) (ereal x)} \<in> sets borel" + "{x \<in> space borel. H (\<infinity>) (ereal x)} \<in> sets borel" + "{x \<in> space borel. H (ereal x) (-\<infinity>)} \<in> sets borel" + "{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel" + shows "{x \<in> space M. H (f x) (g x)} \<in> sets M" +proof - + let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = -\<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))" + let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = -\<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x" + { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } + note * = this + from assms show ?thesis + by (subst *) (simp del: space_borel split del: if_split) +qed + +lemma borel_measurable_ereal_iff: + shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" +proof + assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M" + from borel_measurable_real_of_ereal[OF this] + show "f \<in> borel_measurable M" by auto +qed auto + +lemma borel_measurable_erealD[measurable_dest]: + "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N" + unfolding borel_measurable_ereal_iff by simp + +lemma borel_measurable_ereal_iff_real: + fixes f :: "'a \<Rightarrow> ereal" + shows "f \<in> borel_measurable M \<longleftrightarrow> + ((\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)" +proof safe + assume *: "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M" + have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto + with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all + let ?f = "\<lambda>x. if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real_of_ereal (f x))" + have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto + also have "?f = f" by (auto simp: fun_eq_iff ereal_real) + finally show "f \<in> borel_measurable M" . +qed simp_all + +lemma borel_measurable_ereal_iff_Iio: + "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)" + by (auto simp: borel_Iio measurable_iff_measure_of) + +lemma borel_measurable_ereal_iff_Ioi: + "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)" + by (auto simp: borel_Ioi measurable_iff_measure_of) + +lemma vimage_sets_compl_iff: + "f -` A \<inter> space M \<in> sets M \<longleftrightarrow> f -` (- A) \<inter> space M \<in> sets M" +proof - + { fix A assume "f -` A \<inter> space M \<in> sets M" + moreover have "f -` (- A) \<inter> space M = space M - f -` A \<inter> space M" by auto + ultimately have "f -` (- A) \<inter> space M \<in> sets M" by auto } + from this[of A] this[of "-A"] show ?thesis + by (metis double_complement) +qed + +lemma borel_measurable_iff_Iic_ereal: + "(f::'a\<Rightarrow>ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)" + unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp + +lemma borel_measurable_iff_Ici_ereal: + "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)" + unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp + +lemma borel_measurable_ereal2: + fixes f g :: "'a \<Rightarrow> ereal" + assumes f: "f \<in> borel_measurable M" + assumes g: "g \<in> borel_measurable M" + assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M" + "(\<lambda>x. H (-\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M" + "(\<lambda>x. H (\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M" + "(\<lambda>x. H (ereal (real_of_ereal (f x))) (-\<infinity>)) \<in> borel_measurable M" + "(\<lambda>x. H (ereal (real_of_ereal (f x))) (\<infinity>)) \<in> borel_measurable M" + shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M" +proof - + let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = - \<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))" + let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = - \<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x" + { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } + note * = this + from assms show ?thesis unfolding * by simp +qed + +lemma [measurable(raw)]: + fixes f :: "'a \<Rightarrow> ereal" + assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" + shows borel_measurable_ereal_add: "(\<lambda>x. f x + g x) \<in> borel_measurable M" + and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M" + by (simp_all add: borel_measurable_ereal2) + +lemma [measurable(raw)]: + fixes f g :: "'a \<Rightarrow> ereal" + assumes "f \<in> borel_measurable M" + assumes "g \<in> borel_measurable M" + shows borel_measurable_ereal_diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" + and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M" + using assms by (simp_all add: minus_ereal_def divide_ereal_def) + +lemma borel_measurable_ereal_setsum[measurable (raw)]: + fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" + assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" + shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M" + using assms by (induction S rule: infinite_finite_induct) auto + +lemma borel_measurable_ereal_setprod[measurable (raw)]: + fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal" + assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" + shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" + using assms by (induction S rule: infinite_finite_induct) auto + +lemma borel_measurable_extreal_suminf[measurable (raw)]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" + assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M" + unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp + +subsection "Borel space on the extended non-negative reals" + +text \<open> @{type ennreal} is a topological monoid, so no rules for plus are required, also all order + statements are usually done on type classes. \<close> + +lemma measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel" + by (intro borel_measurable_continuous_on1 continuous_on_enn2ereal) + +lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel" + by (intro borel_measurable_continuous_on1 continuous_on_e2ennreal) + +lemma borel_measurable_enn2real[measurable (raw)]: + "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel" + unfolding enn2real_def[abs_def] by measurable + +definition [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M" + +lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) op = is_borel is_borel" + unfolding is_borel_def[abs_def] +proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1]) + fix f and M :: "'a measure" + show "f \<in> borel_measurable M" if f: "enn2ereal \<circ> f \<in> borel_measurable M" + using measurable_compose[OF f measurable_e2ennreal] by simp +qed simp + +context + includes ennreal.lifting +begin + +lemma measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel" + unfolding is_borel_def[symmetric] + by transfer simp + +lemma borel_measurable_ennreal_iff[simp]: + assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" + shows "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel \<longleftrightarrow> f \<in> M \<rightarrow>\<^sub>M borel" +proof safe + assume "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel" + then have "(\<lambda>x. enn2real (ennreal (f x))) \<in> M \<rightarrow>\<^sub>M borel" + by measurable + then show "f \<in> M \<rightarrow>\<^sub>M borel" + by (rule measurable_cong[THEN iffD1, rotated]) auto +qed measurable + +lemma borel_measurable_times_ennreal[measurable (raw)]: + fixes f g :: "'a \<Rightarrow> ennreal" + shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x * g x) \<in> M \<rightarrow>\<^sub>M borel" + unfolding is_borel_def[symmetric] by transfer simp + +lemma borel_measurable_inverse_ennreal[measurable (raw)]: + fixes f :: "'a \<Rightarrow> ennreal" + shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. inverse (f x)) \<in> M \<rightarrow>\<^sub>M borel" + unfolding is_borel_def[symmetric] by transfer simp + +lemma borel_measurable_divide_ennreal[measurable (raw)]: + fixes f :: "'a \<Rightarrow> ennreal" + shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x / g x) \<in> M \<rightarrow>\<^sub>M borel" + unfolding divide_ennreal_def by simp + +lemma borel_measurable_minus_ennreal[measurable (raw)]: + fixes f :: "'a \<Rightarrow> ennreal" + shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x - g x) \<in> M \<rightarrow>\<^sub>M borel" + unfolding is_borel_def[symmetric] by transfer simp + +lemma borel_measurable_setprod_ennreal[measurable (raw)]: + fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ennreal" + assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M" + shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M" + using assms by (induction S rule: infinite_finite_induct) auto + +end + +hide_const (open) is_borel + +subsection \<open>LIMSEQ is borel measurable\<close> + +lemma borel_measurable_LIMSEQ_real: + fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" + assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x" + and u: "\<And>i. u i \<in> borel_measurable M" + shows "u' \<in> borel_measurable M" +proof - + have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)" + using u' by (simp add: lim_imp_Liminf) + moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M" + by auto + ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) +qed + +lemma borel_measurable_LIMSEQ_metric: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b :: metric_space" + assumes [measurable]: "\<And>i. f i \<in> borel_measurable M" + assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) \<longlonglongrightarrow> g x" + shows "g \<in> borel_measurable M" + unfolding borel_eq_closed +proof (safe intro!: measurable_measure_of) + fix A :: "'b set" assume "closed A" + + have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M" + proof (rule borel_measurable_LIMSEQ_real) + show "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. infdist (f i x) A) \<longlonglongrightarrow> infdist (g x) A" + by (intro tendsto_infdist lim) + show "\<And>i. (\<lambda>x. infdist (f i x) A) \<in> borel_measurable M" + by (intro borel_measurable_continuous_on[where f="\<lambda>x. infdist x A"] + continuous_at_imp_continuous_on ballI continuous_infdist continuous_ident) auto + qed + + show "g -` A \<inter> space M \<in> sets M" + proof cases + assume "A \<noteq> {}" + then have "\<And>x. infdist x A = 0 \<longleftrightarrow> x \<in> A" + using \<open>closed A\<close> by (simp add: in_closed_iff_infdist_zero) + then have "g -` A \<inter> space M = {x\<in>space M. infdist (g x) A = 0}" + by auto + also have "\<dots> \<in> sets M" + by measurable + finally show ?thesis . + qed simp +qed auto + +lemma sets_Collect_Cauchy[measurable]: + fixes f :: "nat \<Rightarrow> 'a => 'b::{metric_space, second_countable_topology}" + assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M" + unfolding metric_Cauchy_iff2 using f by auto + +lemma borel_measurable_lim_metric[measurable (raw)]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M" +proof - + define u' where "u' x = lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" for x + then have *: "\<And>x. lim (\<lambda>i. f i x) = (if Cauchy (\<lambda>i. f i x) then u' x else (THE x. False))" + by (auto simp: lim_def convergent_eq_cauchy[symmetric]) + have "u' \<in> borel_measurable M" + proof (rule borel_measurable_LIMSEQ_metric) + fix x + have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" + by (cases "Cauchy (\<lambda>i. f i x)") + (auto simp add: convergent_eq_cauchy[symmetric] convergent_def) + then show "(\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) \<longlonglongrightarrow> u' x" + unfolding u'_def + by (rule convergent_LIMSEQ_iff[THEN iffD1]) + qed measurable + then show ?thesis + unfolding * by measurable +qed + +lemma borel_measurable_suminf[measurable (raw)]: + fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" + assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M" + shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M" + unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp + +lemma Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P" + by (simp add: pred_def) + +(* Proof by Jeremy Avigad and Luke Serafin *) +lemma isCont_borel_pred[measurable]: + fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space" + shows "Measurable.pred borel (isCont f)" +proof (subst measurable_cong) + let ?I = "\<lambda>j. inverse(real (Suc j))" + show "isCont f x = (\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i)" for x + unfolding continuous_at_eps_delta + proof safe + fix i assume "\<forall>e>0. \<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e" + moreover have "0 < ?I i / 2" + by simp + ultimately obtain d where d: "0 < d" "\<And>y. dist x y < d \<Longrightarrow> dist (f y) (f x) < ?I i / 2" + by (metis dist_commute) + then obtain j where j: "?I j < d" + by (metis reals_Archimedean) + + show "\<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i" + proof (safe intro!: exI[where x=j]) + fix y z assume *: "dist x y < ?I j" "dist x z < ?I j" + have "dist (f y) (f z) \<le> dist (f y) (f x) + dist (f z) (f x)" + by (rule dist_triangle2) + also have "\<dots> < ?I i / 2 + ?I i / 2" + by (intro add_strict_mono d less_trans[OF _ j] *) + also have "\<dots> \<le> ?I i" + by (simp add: field_simps of_nat_Suc) + finally show "dist (f y) (f z) \<le> ?I i" + by simp + qed + next + fix e::real assume "0 < e" + then obtain n where n: "?I n < e" + by (metis reals_Archimedean) + assume "\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i" + from this[THEN spec, of "Suc n"] + obtain j where j: "\<And>y z. dist x y < ?I j \<Longrightarrow> dist x z < ?I j \<Longrightarrow> dist (f y) (f z) \<le> ?I (Suc n)" + by auto + + show "\<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e" + proof (safe intro!: exI[of _ "?I j"]) + fix y assume "dist y x < ?I j" + then have "dist (f y) (f x) \<le> ?I (Suc n)" + by (intro j) (auto simp: dist_commute) + also have "?I (Suc n) < ?I n" + by simp + also note n + finally show "dist (f y) (f x) < e" . + qed simp + qed +qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less + Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros) + +lemma isCont_borel: + fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space" + shows "{x. isCont f x} \<in> sets borel" + by simp + +lemma is_real_interval: + assumes S: "is_interval S" + shows "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or> S = {a<..} \<or> S = {a..} \<or> + S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}" + using S unfolding is_interval_1 by (blast intro: interval_cases) + +lemma real_interval_borel_measurable: + assumes "is_interval (S::real set)" + shows "S \<in> sets borel" +proof - + from assms is_real_interval have "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or> + S = {a<..} \<or> S = {a..} \<or> S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}" by auto + then guess a .. + then guess b .. + thus ?thesis + by auto +qed + +lemma borel_measurable_mono_on_fnc: + fixes f :: "real \<Rightarrow> real" and A :: "real set" + assumes "mono_on f A" + shows "f \<in> borel_measurable (restrict_space borel A)" + apply (rule measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]]) + apply (auto intro!: image_eqI[where x="{x}" for x] simp: sets_restrict_space) + apply (auto simp add: sets_restrict_restrict_space continuous_on_eq_continuous_within + cong: measurable_cong_sets + intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset) + done + +lemma borel_measurable_mono: + fixes f :: "real \<Rightarrow> real" + shows "mono f \<Longrightarrow> f \<in> borel_measurable borel" + using borel_measurable_mono_on_fnc[of f UNIV] by (simp add: mono_def mono_on_def) + +no_notation + eucl_less (infix "<e" 50) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Bounded_Continuous_Function.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,518 @@ +section \<open>Bounded Continuous Functions\<close> + +theory Bounded_Continuous_Function +imports Henstock_Kurzweil_Integration +begin + +subsection \<open>Definition\<close> + +definition bcontfun :: "('a::topological_space \<Rightarrow> 'b::metric_space) set" + where "bcontfun = {f. continuous_on UNIV f \<and> bounded (range f)}" + +typedef (overloaded) ('a, 'b) bcontfun = + "bcontfun :: ('a::topological_space \<Rightarrow> 'b::metric_space) set" + by (auto simp: bcontfun_def intro: continuous_intros simp: bounded_def) + +lemma bcontfunE: + assumes "f \<in> bcontfun" + obtains y where "continuous_on UNIV f" "\<And>x. dist (f x) u \<le> y" + using assms unfolding bcontfun_def + by (metis (lifting) bounded_any_center dist_commute mem_Collect_eq rangeI) + +lemma bcontfunE': + assumes "f \<in> bcontfun" + obtains y where "continuous_on UNIV f" "\<And>x. dist (f x) undefined \<le> y" + using assms bcontfunE + by metis + +lemma bcontfunI: "continuous_on UNIV f \<Longrightarrow> (\<And>x. dist (f x) u \<le> b) \<Longrightarrow> f \<in> bcontfun" + unfolding bcontfun_def + by (metis (lifting, no_types) bounded_def dist_commute mem_Collect_eq rangeE) + +lemma bcontfunI': "continuous_on UNIV f \<Longrightarrow> (\<And>x. dist (f x) undefined \<le> b) \<Longrightarrow> f \<in> bcontfun" + using bcontfunI by metis + +lemma continuous_on_Rep_bcontfun[intro, simp]: "continuous_on T (Rep_bcontfun x)" + using Rep_bcontfun[of x] + by (auto simp: bcontfun_def intro: continuous_on_subset) + +(* TODO: Generalize to uniform spaces? *) +instantiation bcontfun :: (topological_space, metric_space) metric_space +begin + +definition dist_bcontfun :: "('a, 'b) bcontfun \<Rightarrow> ('a, 'b) bcontfun \<Rightarrow> real" + where "dist_bcontfun f g = (SUP x. dist (Rep_bcontfun f x) (Rep_bcontfun g x))" + +definition uniformity_bcontfun :: "(('a, 'b) bcontfun \<times> ('a, 'b) bcontfun) filter" + where "uniformity_bcontfun = (INF e:{0 <..}. principal {(x, y). dist x y < e})" + +definition open_bcontfun :: "('a, 'b) bcontfun set \<Rightarrow> bool" + where "open_bcontfun S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)" + +lemma dist_bounded: + fixes f :: "('a, 'b) bcontfun" + shows "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f g" +proof - + have "Rep_bcontfun f \<in> bcontfun" by (rule Rep_bcontfun) + from bcontfunE'[OF this] obtain y where y: + "continuous_on UNIV (Rep_bcontfun f)" + "\<And>x. dist (Rep_bcontfun f x) undefined \<le> y" + by auto + have "Rep_bcontfun g \<in> bcontfun" by (rule Rep_bcontfun) + from bcontfunE'[OF this] obtain z where z: + "continuous_on UNIV (Rep_bcontfun g)" + "\<And>x. dist (Rep_bcontfun g x) undefined \<le> z" + by auto + show ?thesis + unfolding dist_bcontfun_def + proof (intro cSUP_upper bdd_aboveI2) + fix x + have "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> + dist (Rep_bcontfun f x) undefined + dist (Rep_bcontfun g x) undefined" + by (rule dist_triangle2) + also have "\<dots> \<le> y + z" + using y(2)[of x] z(2)[of x] by (rule add_mono) + finally show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> y + z" . + qed simp +qed + +lemma dist_bound: + fixes f :: "('a, 'b) bcontfun" + assumes "\<And>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> b" + shows "dist f g \<le> b" + using assms by (auto simp: dist_bcontfun_def intro: cSUP_least) + +lemma dist_bounded_Abs: + fixes f g :: "'a \<Rightarrow> 'b" + assumes "f \<in> bcontfun" "g \<in> bcontfun" + shows "dist (f x) (g x) \<le> dist (Abs_bcontfun f) (Abs_bcontfun g)" + by (metis Abs_bcontfun_inverse assms dist_bounded) + +lemma const_bcontfun: "(\<lambda>x::'a. b::'b) \<in> bcontfun" + by (auto intro: bcontfunI continuous_on_const) + +lemma dist_fun_lt_imp_dist_val_lt: + assumes "dist f g < e" + shows "dist (Rep_bcontfun f x) (Rep_bcontfun g x) < e" + using dist_bounded assms by (rule le_less_trans) + +lemma dist_val_lt_imp_dist_fun_le: + assumes "\<forall>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) < e" + shows "dist f g \<le> e" + unfolding dist_bcontfun_def +proof (intro cSUP_least) + fix x + show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> e" + using assms[THEN spec[where x=x]] by (simp add: dist_norm) +qed simp + +instance +proof + fix f g h :: "('a, 'b) bcontfun" + show "dist f g = 0 \<longleftrightarrow> f = g" + proof + have "\<And>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f g" + by (rule dist_bounded) + also assume "dist f g = 0" + finally show "f = g" + by (auto simp: Rep_bcontfun_inject[symmetric] Abs_bcontfun_inverse) + qed (auto simp: dist_bcontfun_def intro!: cSup_eq) + show "dist f g \<le> dist f h + dist g h" + proof (subst dist_bcontfun_def, safe intro!: cSUP_least) + fix x + have "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> + dist (Rep_bcontfun f x) (Rep_bcontfun h x) + dist (Rep_bcontfun g x) (Rep_bcontfun h x)" + by (rule dist_triangle2) + also have "dist (Rep_bcontfun f x) (Rep_bcontfun h x) \<le> dist f h" + by (rule dist_bounded) + also have "dist (Rep_bcontfun g x) (Rep_bcontfun h x) \<le> dist g h" + by (rule dist_bounded) + finally show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f h + dist g h" + by simp + qed +qed (rule open_bcontfun_def uniformity_bcontfun_def)+ + +end + +lemma closed_Pi_bcontfun: + fixes I :: "'a::metric_space set" + and X :: "'a \<Rightarrow> 'b::complete_space set" + assumes "\<And>i. i \<in> I \<Longrightarrow> closed (X i)" + shows "closed (Abs_bcontfun ` (Pi I X \<inter> bcontfun))" + unfolding closed_sequential_limits +proof safe + fix f l + assume seq: "\<forall>n. f n \<in> Abs_bcontfun ` (Pi I X \<inter> bcontfun)" and lim: "f \<longlonglongrightarrow> l" + have lim_fun: "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (Rep_bcontfun l x) < e" + using LIMSEQ_imp_Cauchy[OF lim, simplified Cauchy_def] metric_LIMSEQ_D[OF lim] + by (intro uniformly_cauchy_imp_uniformly_convergent[where P="\<lambda>_. True", simplified]) + (metis dist_fun_lt_imp_dist_val_lt)+ + show "l \<in> Abs_bcontfun ` (Pi I X \<inter> bcontfun)" + proof (rule, safe) + fix x assume "x \<in> I" + then have "closed (X x)" + using assms by simp + moreover have "eventually (\<lambda>xa. Rep_bcontfun (f xa) x \<in> X x) sequentially" + proof (rule always_eventually, safe) + fix i + from seq[THEN spec, of i] \<open>x \<in> I\<close> + show "Rep_bcontfun (f i) x \<in> X x" + by (auto simp: Abs_bcontfun_inverse) + qed + moreover note sequentially_bot + moreover have "(\<lambda>n. Rep_bcontfun (f n) x) \<longlonglongrightarrow> Rep_bcontfun l x" + using lim_fun by (blast intro!: metric_LIMSEQ_I) + ultimately show "Rep_bcontfun l x \<in> X x" + by (rule Lim_in_closed_set) + qed (auto simp: Rep_bcontfun Rep_bcontfun_inverse) +qed + + +subsection \<open>Complete Space\<close> + +instance bcontfun :: (metric_space, complete_space) complete_space +proof + fix f :: "nat \<Rightarrow> ('a, 'b) bcontfun" + assume "Cauchy f" \<comment> \<open>Cauchy equals uniform convergence\<close> + then obtain g where limit_function: + "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e" + using uniformly_convergent_eq_cauchy[of "\<lambda>_. True" + "\<lambda>n. Rep_bcontfun (f n)"] + unfolding Cauchy_def + by (metis dist_fun_lt_imp_dist_val_lt) + + then obtain N where fg_dist: \<comment> \<open>for an upper bound on @{term g}\<close> + "\<forall>n\<ge>N. \<forall>x. dist (g x) ( Rep_bcontfun (f n) x) < 1" + by (force simp add: dist_commute) + from bcontfunE'[OF Rep_bcontfun, of "f N"] obtain b where + f_bound: "\<forall>x. dist (Rep_bcontfun (f N) x) undefined \<le> b" + by force + have "g \<in> bcontfun" \<comment> \<open>The limit function is bounded and continuous\<close> + proof (intro bcontfunI) + show "continuous_on UNIV g" + using bcontfunE[OF Rep_bcontfun] limit_function + by (intro continuous_uniform_limit[where f="\<lambda>n. Rep_bcontfun (f n)" and F="sequentially"]) + (auto simp add: eventually_sequentially trivial_limit_def dist_norm) + next + fix x + from fg_dist have "dist (g x) (Rep_bcontfun (f N) x) < 1" + by (simp add: dist_norm norm_minus_commute) + with dist_triangle[of "g x" undefined "Rep_bcontfun (f N) x"] + show "dist (g x) undefined \<le> 1 + b" using f_bound[THEN spec, of x] + by simp + qed + show "convergent f" + proof (rule convergentI, subst lim_sequentially, safe) + \<comment> \<open>The limit function converges according to its norm\<close> + fix e :: real + assume "e > 0" + then have "e/2 > 0" by simp + with limit_function[THEN spec, of"e/2"] + have "\<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e/2" + by simp + then obtain N where N: "\<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e / 2" by auto + show "\<exists>N. \<forall>n\<ge>N. dist (f n) (Abs_bcontfun g) < e" + proof (rule, safe) + fix n + assume "N \<le> n" + with N show "dist (f n) (Abs_bcontfun g) < e" + using dist_val_lt_imp_dist_fun_le[of + "f n" "Abs_bcontfun g" "e/2"] + Abs_bcontfun_inverse[OF \<open>g \<in> bcontfun\<close>] \<open>e > 0\<close> by simp + qed + qed +qed + + +subsection \<open>Supremum norm for a normed vector space\<close> + +instantiation bcontfun :: (topological_space, real_normed_vector) real_vector +begin + +definition "-f = Abs_bcontfun (\<lambda>x. -(Rep_bcontfun f x))" + +definition "f + g = Abs_bcontfun (\<lambda>x. Rep_bcontfun f x + Rep_bcontfun g x)" + +definition "f - g = Abs_bcontfun (\<lambda>x. Rep_bcontfun f x - Rep_bcontfun g x)" + +definition "0 = Abs_bcontfun (\<lambda>x. 0)" + +definition "scaleR r f = Abs_bcontfun (\<lambda>x. r *\<^sub>R Rep_bcontfun f x)" + +lemma plus_cont: + fixes f g :: "'a \<Rightarrow> 'b" + assumes f: "f \<in> bcontfun" + and g: "g \<in> bcontfun" + shows "(\<lambda>x. f x + g x) \<in> bcontfun" +proof - + from bcontfunE'[OF f] obtain y where "continuous_on UNIV f" "\<And>x. dist (f x) undefined \<le> y" + by auto + moreover + from bcontfunE'[OF g] obtain z where "continuous_on UNIV g" "\<And>x. dist (g x) undefined \<le> z" + by auto + ultimately show ?thesis + proof (intro bcontfunI) + fix x + have "dist (f x + g x) 0 = dist (f x + g x) (0 + 0)" + by simp + also have "\<dots> \<le> dist (f x) 0 + dist (g x) 0" + by (rule dist_triangle_add) + also have "\<dots> \<le> dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0" + unfolding zero_bcontfun_def using assms + by (metis add_mono const_bcontfun dist_bounded_Abs) + finally show "dist (f x + g x) 0 \<le> dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0" . + qed (simp add: continuous_on_add) +qed + +lemma Rep_bcontfun_plus[simp]: "Rep_bcontfun (f + g) x = Rep_bcontfun f x + Rep_bcontfun g x" + by (simp add: plus_bcontfun_def Abs_bcontfun_inverse plus_cont Rep_bcontfun) + +lemma uminus_cont: + fixes f :: "'a \<Rightarrow> 'b" + assumes "f \<in> bcontfun" + shows "(\<lambda>x. - f x) \<in> bcontfun" +proof - + from bcontfunE[OF assms, of 0] obtain y + where "continuous_on UNIV f" "\<And>x. dist (f x) 0 \<le> y" + by auto + then show ?thesis + proof (intro bcontfunI) + fix x + assume "\<And>x. dist (f x) 0 \<le> y" + then show "dist (- f x) 0 \<le> y" + by (subst dist_minus[symmetric]) simp + qed (simp add: continuous_on_minus) +qed + +lemma Rep_bcontfun_uminus[simp]: "Rep_bcontfun (- f) x = - Rep_bcontfun f x" + by (simp add: uminus_bcontfun_def Abs_bcontfun_inverse uminus_cont Rep_bcontfun) + +lemma minus_cont: + fixes f g :: "'a \<Rightarrow> 'b" + assumes f: "f \<in> bcontfun" + and g: "g \<in> bcontfun" + shows "(\<lambda>x. f x - g x) \<in> bcontfun" + using plus_cont [of f "- g"] assms + by (simp add: uminus_cont fun_Compl_def) + +lemma Rep_bcontfun_minus[simp]: "Rep_bcontfun (f - g) x = Rep_bcontfun f x - Rep_bcontfun g x" + by (simp add: minus_bcontfun_def Abs_bcontfun_inverse minus_cont Rep_bcontfun) + +lemma scaleR_cont: + fixes a :: real + and f :: "'a \<Rightarrow> 'b" + assumes "f \<in> bcontfun" + shows " (\<lambda>x. a *\<^sub>R f x) \<in> bcontfun" +proof - + from bcontfunE[OF assms, of 0] obtain y + where "continuous_on UNIV f" "\<And>x. dist (f x) 0 \<le> y" + by auto + then show ?thesis + proof (intro bcontfunI) + fix x + assume "\<And>x. dist (f x) 0 \<le> y" + then show "dist (a *\<^sub>R f x) 0 \<le> \<bar>a\<bar> * y" + by (metis norm_cmul_rule_thm norm_conv_dist) + qed (simp add: continuous_intros) +qed + +lemma Rep_bcontfun_scaleR[simp]: "Rep_bcontfun (a *\<^sub>R g) x = a *\<^sub>R Rep_bcontfun g x" + by (simp add: scaleR_bcontfun_def Abs_bcontfun_inverse scaleR_cont Rep_bcontfun) + +instance + by standard + (simp_all add: plus_bcontfun_def zero_bcontfun_def minus_bcontfun_def scaleR_bcontfun_def + Abs_bcontfun_inverse Rep_bcontfun_inverse Rep_bcontfun algebra_simps + plus_cont const_bcontfun minus_cont scaleR_cont) + +end + +instantiation bcontfun :: (topological_space, real_normed_vector) real_normed_vector +begin + +definition norm_bcontfun :: "('a, 'b) bcontfun \<Rightarrow> real" + where "norm_bcontfun f = dist f 0" + +definition "sgn (f::('a,'b) bcontfun) = f /\<^sub>R norm f" + +instance +proof + fix a :: real + fix f g :: "('a, 'b) bcontfun" + show "dist f g = norm (f - g)" + by (simp add: norm_bcontfun_def dist_bcontfun_def zero_bcontfun_def + Abs_bcontfun_inverse const_bcontfun dist_norm) + show "norm (f + g) \<le> norm f + norm g" + unfolding norm_bcontfun_def + proof (subst dist_bcontfun_def, safe intro!: cSUP_least) + fix x + have "dist (Rep_bcontfun (f + g) x) (Rep_bcontfun 0 x) \<le> + dist (Rep_bcontfun f x) 0 + dist (Rep_bcontfun g x) 0" + by (metis (hide_lams, no_types) Rep_bcontfun_minus Rep_bcontfun_plus diff_0_right dist_norm + le_less_linear less_irrefl norm_triangle_lt) + also have "dist (Rep_bcontfun f x) 0 \<le> dist f 0" + using dist_bounded[of f x 0] + by (simp add: Abs_bcontfun_inverse const_bcontfun zero_bcontfun_def) + also have "dist (Rep_bcontfun g x) 0 \<le> dist g 0" using dist_bounded[of g x 0] + by (simp add: Abs_bcontfun_inverse const_bcontfun zero_bcontfun_def) + finally show "dist (Rep_bcontfun (f + g) x) (Rep_bcontfun 0 x) \<le> dist f 0 + dist g 0" by simp + qed + show "norm (a *\<^sub>R f) = \<bar>a\<bar> * norm f" + proof - + have "\<bar>a\<bar> * Sup (range (\<lambda>x. dist (Rep_bcontfun f x) 0)) = + (SUP i:range (\<lambda>x. dist (Rep_bcontfun f x) 0). \<bar>a\<bar> * i)" + proof (intro continuous_at_Sup_mono bdd_aboveI2) + fix x + show "dist (Rep_bcontfun f x) 0 \<le> norm f" using dist_bounded[of f x 0] + by (simp add: norm_bcontfun_def Abs_bcontfun_inverse zero_bcontfun_def + const_bcontfun) + qed (auto intro!: monoI mult_left_mono continuous_intros) + moreover + have "range (\<lambda>x. dist (Rep_bcontfun (a *\<^sub>R f) x) 0) = + (\<lambda>x. \<bar>a\<bar> * x) ` (range (\<lambda>x. dist (Rep_bcontfun f x) 0))" + by auto + ultimately + show "norm (a *\<^sub>R f) = \<bar>a\<bar> * norm f" + by (simp add: norm_bcontfun_def dist_bcontfun_def Abs_bcontfun_inverse + zero_bcontfun_def const_bcontfun image_image) + qed +qed (auto simp: norm_bcontfun_def sgn_bcontfun_def) + +end + +lemma bcontfun_normI: "continuous_on UNIV f \<Longrightarrow> (\<And>x. norm (f x) \<le> b) \<Longrightarrow> f \<in> bcontfun" + by (metis bcontfunI dist_0_norm dist_commute) + +lemma norm_bounded: + fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun" + shows "norm (Rep_bcontfun f x) \<le> norm f" + using dist_bounded[of f x 0] + by (simp add: norm_bcontfun_def Abs_bcontfun_inverse zero_bcontfun_def + const_bcontfun) + +lemma norm_bound: + fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun" + assumes "\<And>x. norm (Rep_bcontfun f x) \<le> b" + shows "norm f \<le> b" + using dist_bound[of f 0 b] assms + by (simp add: norm_bcontfun_def Abs_bcontfun_inverse zero_bcontfun_def const_bcontfun) + + +subsection \<open>Continuously Extended Functions\<close> + +definition clamp :: "'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where + "clamp a b x = (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i)" + +definition ext_cont :: "('a::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a, 'b) bcontfun" + where "ext_cont f a b = Abs_bcontfun ((\<lambda>x. f (clamp a b x)))" + +lemma ext_cont_def': + "ext_cont f a b = Abs_bcontfun (\<lambda>x. + f (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i))" + unfolding ext_cont_def clamp_def .. + +lemma clamp_in_interval: + assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" + shows "clamp a b x \<in> cbox a b" + unfolding clamp_def + using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) + +lemma dist_clamps_le_dist_args: + fixes x :: "'a::euclidean_space" + assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" + shows "dist (clamp a b y) (clamp a b x) \<le> dist y x" +proof - + from box_ne_empty(1)[of a b] assms have "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" + by (simp add: cbox_def) + then have "(\<Sum>i\<in>Basis. (dist (clamp a b y \<bullet> i) (clamp a b x \<bullet> i))\<^sup>2) \<le> + (\<Sum>i\<in>Basis. (dist (y \<bullet> i) (x \<bullet> i))\<^sup>2)" + by (auto intro!: setsum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) + then show ?thesis + by (auto intro: real_sqrt_le_mono + simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] setL2_def) +qed + +lemma clamp_continuous_at: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" + and x :: 'a + assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" + and f_cont: "continuous_on (cbox a b) f" + shows "continuous (at x) (\<lambda>x. f (clamp a b x))" + unfolding continuous_at_eps_delta +proof safe + fix x :: 'a + fix e :: real + assume "e > 0" + moreover have "clamp a b x \<in> cbox a b" + by (simp add: clamp_in_interval assms) + moreover note f_cont[simplified continuous_on_iff] + ultimately + obtain d where d: "0 < d" + "\<And>x'. x' \<in> cbox a b \<Longrightarrow> dist x' (clamp a b x) < d \<Longrightarrow> dist (f x') (f (clamp a b x)) < e" + by force + show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow> + dist (f (clamp a b x')) (f (clamp a b x)) < e" + by (auto intro!: d clamp_in_interval assms dist_clamps_le_dist_args[THEN le_less_trans]) +qed + +lemma clamp_continuous_on: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" + assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" + and f_cont: "continuous_on (cbox a b) f" + shows "continuous_on UNIV (\<lambda>x. f (clamp a b x))" + using assms + by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) + +lemma clamp_bcontfun: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" + and continuous: "continuous_on (cbox a b) f" + shows "(\<lambda>x. f (clamp a b x)) \<in> bcontfun" +proof - + have "bounded (f ` (cbox a b))" + by (rule compact_continuous_image[OF continuous compact_cbox[of a b], THEN compact_imp_bounded]) + then obtain c where f_bound: "\<forall>x\<in>f ` cbox a b. norm x \<le> c" + by (auto simp add: bounded_pos) + show "(\<lambda>x. f (clamp a b x)) \<in> bcontfun" + proof (intro bcontfun_normI) + fix x + show "norm (f (clamp a b x)) \<le> c" + using clamp_in_interval[OF assms(1), of x] f_bound + by (simp add: ext_cont_def) + qed (simp add: clamp_continuous_on assms) +qed + +lemma integral_clamp: + "integral {t0::real..clamp t0 t1 x} f = + (if x < t0 then 0 else if x \<le> t1 then integral {t0..x} f else integral {t0..t1} f)" + by (auto simp: clamp_def) + + +declare [[coercion Rep_bcontfun]] + +lemma ext_cont_cancel[simp]: + fixes x a b :: "'a::euclidean_space" + assumes x: "x \<in> cbox a b" + and "continuous_on (cbox a b) f" + shows "ext_cont f a b x = f x" + using assms + unfolding ext_cont_def +proof (subst Abs_bcontfun_inverse[OF clamp_bcontfun]) + show "f (clamp a b x) = f x" + using x unfolding clamp_def mem_box + by (intro arg_cong[where f=f] euclidean_eqI[where 'a='a]) (simp add: not_less) +qed (auto simp: cbox_def) + +lemma ext_cont_cong: + assumes "t0 = s0" + and "t1 = s1" + and "\<And>t. t \<in> (cbox t0 t1) \<Longrightarrow> f t = g t" + and "continuous_on (cbox t0 t1) f" + and "continuous_on (cbox s0 s1) g" + and ord: "\<And>i. i \<in> Basis \<Longrightarrow> t0 \<bullet> i \<le> t1 \<bullet> i" + shows "ext_cont f t0 t1 = ext_cont g s0 s1" + unfolding assms ext_cont_def + using assms clamp_in_interval[OF ord] + by (subst Rep_bcontfun_inject[symmetric]) simp + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Bounded_Linear_Function.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,695 @@ +(* Title: HOL/Analysis/Bounded_Linear_Function.thy + Author: Fabian Immler, TU München +*) + +section \<open>Bounded Linear Function\<close> + +theory Bounded_Linear_Function +imports + Topology_Euclidean_Space + Operator_Norm +begin + +subsection \<open>Intro rules for @{term bounded_linear}\<close> + +named_theorems bounded_linear_intros + +lemma onorm_inner_left: + assumes "bounded_linear r" + shows "onorm (\<lambda>x. r x \<bullet> f) \<le> onorm r * norm f" +proof (rule onorm_bound) + fix x + have "norm (r x \<bullet> f) \<le> norm (r x) * norm f" + by (simp add: Cauchy_Schwarz_ineq2) + also have "\<dots> \<le> onorm r * norm x * norm f" + by (intro mult_right_mono onorm assms norm_ge_zero) + finally show "norm (r x \<bullet> f) \<le> onorm r * norm f * norm x" + by (simp add: ac_simps) +qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le assms) + +lemma onorm_inner_right: + assumes "bounded_linear r" + shows "onorm (\<lambda>x. f \<bullet> r x) \<le> norm f * onorm r" + apply (subst inner_commute) + apply (rule onorm_inner_left[OF assms, THEN order_trans]) + apply simp + done + +lemmas [bounded_linear_intros] = + bounded_linear_zero + bounded_linear_add + bounded_linear_const_mult + bounded_linear_mult_const + bounded_linear_scaleR_const + bounded_linear_const_scaleR + bounded_linear_ident + bounded_linear_setsum + bounded_linear_Pair + bounded_linear_sub + bounded_linear_fst_comp + bounded_linear_snd_comp + bounded_linear_inner_left_comp + bounded_linear_inner_right_comp + + +subsection \<open>declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\<close> + +attribute_setup bounded_linear = + \<open>Scan.succeed (Thm.declaration_attribute (fn thm => + fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r)) + [ + (@{thm bounded_linear.has_derivative}, @{named_theorems derivative_intros}), + (@{thm bounded_linear.tendsto}, @{named_theorems tendsto_intros}), + (@{thm bounded_linear.continuous}, @{named_theorems continuous_intros}), + (@{thm bounded_linear.continuous_on}, @{named_theorems continuous_intros}), + (@{thm bounded_linear.uniformly_continuous_on}, @{named_theorems continuous_intros}), + (@{thm bounded_linear_compose}, @{named_theorems bounded_linear_intros}) + ]))\<close> + +attribute_setup bounded_bilinear = + \<open>Scan.succeed (Thm.declaration_attribute (fn thm => + fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r)) + [ + (@{thm bounded_bilinear.FDERIV}, @{named_theorems derivative_intros}), + (@{thm bounded_bilinear.tendsto}, @{named_theorems tendsto_intros}), + (@{thm bounded_bilinear.continuous}, @{named_theorems continuous_intros}), + (@{thm bounded_bilinear.continuous_on}, @{named_theorems continuous_intros}), + (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_left]}, + @{named_theorems bounded_linear_intros}), + (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_right]}, + @{named_theorems bounded_linear_intros}), + (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_left]}, + @{named_theorems continuous_intros}), + (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_right]}, + @{named_theorems continuous_intros}) + ]))\<close> + + +subsection \<open>type of bounded linear functions\<close> + +typedef (overloaded) ('a, 'b) blinfun ("(_ \<Rightarrow>\<^sub>L /_)" [22, 21] 21) = + "{f::'a::real_normed_vector\<Rightarrow>'b::real_normed_vector. bounded_linear f}" + morphisms blinfun_apply Blinfun + by (blast intro: bounded_linear_intros) + +declare [[coercion + "blinfun_apply :: ('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'b"]] + +lemma bounded_linear_blinfun_apply[bounded_linear_intros]: + "bounded_linear g \<Longrightarrow> bounded_linear (\<lambda>x. blinfun_apply f (g x))" + by (metis blinfun_apply mem_Collect_eq bounded_linear_compose) + +setup_lifting type_definition_blinfun + +lemma blinfun_eqI: "(\<And>i. blinfun_apply x i = blinfun_apply y i) \<Longrightarrow> x = y" + by transfer auto + +lemma bounded_linear_Blinfun_apply: "bounded_linear f \<Longrightarrow> blinfun_apply (Blinfun f) = f" + by (auto simp: Blinfun_inverse) + + +subsection \<open>type class instantiations\<close> + +instantiation blinfun :: (real_normed_vector, real_normed_vector) real_normed_vector +begin + +lift_definition norm_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> real" is onorm . + +lift_definition minus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + is "\<lambda>f g x. f x - g x" + by (rule bounded_linear_sub) + +definition dist_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> real" + where "dist_blinfun a b = norm (a - b)" + +definition [code del]: + "(uniformity :: (('a \<Rightarrow>\<^sub>L 'b) \<times> ('a \<Rightarrow>\<^sub>L 'b)) filter) = (INF e:{0 <..}. principal {(x, y). dist x y < e})" + +definition open_blinfun :: "('a \<Rightarrow>\<^sub>L 'b) set \<Rightarrow> bool" + where [code del]: "open_blinfun S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)" + +lift_definition uminus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>f x. - f x" + by (rule bounded_linear_minus) + +lift_definition zero_blinfun :: "'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>x. 0" + by (rule bounded_linear_zero) + +lift_definition plus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + is "\<lambda>f g x. f x + g x" + by (metis bounded_linear_add) + +lift_definition scaleR_blinfun::"real \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>r f x. r *\<^sub>R f x" + by (metis bounded_linear_compose bounded_linear_scaleR_right) + +definition sgn_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + where "sgn_blinfun x = scaleR (inverse (norm x)) x" + +instance + apply standard + unfolding dist_blinfun_def open_blinfun_def sgn_blinfun_def uniformity_blinfun_def + apply (rule refl | (transfer, force simp: onorm_triangle onorm_scaleR onorm_eq_0 algebra_simps))+ + done + +end + +declare uniformity_Abort[where 'a="('a :: real_normed_vector) \<Rightarrow>\<^sub>L ('b :: real_normed_vector)", code] + +lemma norm_blinfun_eqI: + assumes "n \<le> norm (blinfun_apply f x) / norm x" + assumes "\<And>x. norm (blinfun_apply f x) \<le> n * norm x" + assumes "0 \<le> n" + shows "norm f = n" + by (auto simp: norm_blinfun_def + intro!: antisym onorm_bound assms order_trans[OF _ le_onorm] + bounded_linear_intros) + +lemma norm_blinfun: "norm (blinfun_apply f x) \<le> norm f * norm x" + by transfer (rule onorm) + +lemma norm_blinfun_bound: "0 \<le> b \<Longrightarrow> (\<And>x. norm (blinfun_apply f x) \<le> b * norm x) \<Longrightarrow> norm f \<le> b" + by transfer (rule onorm_bound) + +lemma bounded_bilinear_blinfun_apply[bounded_bilinear]: "bounded_bilinear blinfun_apply" +proof + fix f g::"'a \<Rightarrow>\<^sub>L 'b" and a b::'a and r::real + show "(f + g) a = f a + g a" "(r *\<^sub>R f) a = r *\<^sub>R f a" + by (transfer, simp)+ + interpret bounded_linear f for f::"'a \<Rightarrow>\<^sub>L 'b" + by (auto intro!: bounded_linear_intros) + show "f (a + b) = f a + f b" "f (r *\<^sub>R a) = r *\<^sub>R f a" + by (simp_all add: add scaleR) + show "\<exists>K. \<forall>a b. norm (blinfun_apply a b) \<le> norm a * norm b * K" + by (auto intro!: exI[where x=1] norm_blinfun) +qed + +interpretation blinfun: bounded_bilinear blinfun_apply + by (rule bounded_bilinear_blinfun_apply) + +lemmas bounded_linear_apply_blinfun[intro, simp] = blinfun.bounded_linear_left + + +context bounded_bilinear +begin + +named_theorems bilinear_simps + +lemmas [bilinear_simps] = + add_left + add_right + diff_left + diff_right + minus_left + minus_right + scaleR_left + scaleR_right + zero_left + zero_right + setsum_left + setsum_right + +end + + +instance blinfun :: (banach, banach) banach +proof + fix X::"nat \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + assume "Cauchy X" + { + fix x::'a + { + fix x::'a + assume "norm x \<le> 1" + have "Cauchy (\<lambda>n. X n x)" + proof (rule CauchyI) + fix e::real + assume "0 < e" + from CauchyD[OF \<open>Cauchy X\<close> \<open>0 < e\<close>] obtain M + where M: "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> norm (X m - X n) < e" + by auto + show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m x - X n x) < e" + proof (safe intro!: exI[where x=M]) + fix m n + assume le: "M \<le> m" "M \<le> n" + have "norm (X m x - X n x) = norm ((X m - X n) x)" + by (simp add: blinfun.bilinear_simps) + also have "\<dots> \<le> norm (X m - X n) * norm x" + by (rule norm_blinfun) + also have "\<dots> \<le> norm (X m - X n) * 1" + using \<open>norm x \<le> 1\<close> norm_ge_zero by (rule mult_left_mono) + also have "\<dots> = norm (X m - X n)" by simp + also have "\<dots> < e" using le by fact + finally show "norm (X m x - X n x) < e" . + qed + qed + hence "convergent (\<lambda>n. X n x)" + by (metis Cauchy_convergent_iff) + } note convergent_norm1 = this + define y where "y = x /\<^sub>R norm x" + have y: "norm y \<le> 1" and xy: "x = norm x *\<^sub>R y" + by (simp_all add: y_def inverse_eq_divide) + have "convergent (\<lambda>n. norm x *\<^sub>R X n y)" + by (intro bounded_bilinear.convergent[OF bounded_bilinear_scaleR] convergent_const + convergent_norm1 y) + also have "(\<lambda>n. norm x *\<^sub>R X n y) = (\<lambda>n. X n x)" + by (subst xy) (simp add: blinfun.bilinear_simps) + finally have "convergent (\<lambda>n. X n x)" . + } + then obtain v where v: "\<And>x. (\<lambda>n. X n x) \<longlonglongrightarrow> v x" + unfolding convergent_def + by metis + + have "Cauchy (\<lambda>n. norm (X n))" + proof (rule CauchyI) + fix e::real + assume "e > 0" + from CauchyD[OF \<open>Cauchy X\<close> \<open>0 < e\<close>] obtain M + where M: "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> norm (X m - X n) < e" + by auto + show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm (norm (X m) - norm (X n)) < e" + proof (safe intro!: exI[where x=M]) + fix m n assume mn: "m \<ge> M" "n \<ge> M" + have "norm (norm (X m) - norm (X n)) \<le> norm (X m - X n)" + by (metis norm_triangle_ineq3 real_norm_def) + also have "\<dots> < e" using mn by fact + finally show "norm (norm (X m) - norm (X n)) < e" . + qed + qed + then obtain K where K: "(\<lambda>n. norm (X n)) \<longlonglongrightarrow> K" + unfolding Cauchy_convergent_iff convergent_def + by metis + + have "bounded_linear v" + proof + fix x y and r::real + from tendsto_add[OF v[of x] v [of y]] v[of "x + y", unfolded blinfun.bilinear_simps] + tendsto_scaleR[OF tendsto_const[of r] v[of x]] v[of "r *\<^sub>R x", unfolded blinfun.bilinear_simps] + show "v (x + y) = v x + v y" "v (r *\<^sub>R x) = r *\<^sub>R v x" + by (metis (poly_guards_query) LIMSEQ_unique)+ + show "\<exists>K. \<forall>x. norm (v x) \<le> norm x * K" + proof (safe intro!: exI[where x=K]) + fix x + have "norm (v x) \<le> K * norm x" + by (rule tendsto_le[OF _ tendsto_mult[OF K tendsto_const] tendsto_norm[OF v]]) + (auto simp: norm_blinfun) + thus "norm (v x) \<le> norm x * K" + by (simp add: ac_simps) + qed + qed + hence Bv: "\<And>x. (\<lambda>n. X n x) \<longlonglongrightarrow> Blinfun v x" + by (auto simp: bounded_linear_Blinfun_apply v) + + have "X \<longlonglongrightarrow> Blinfun v" + proof (rule LIMSEQ_I) + fix r::real assume "r > 0" + define r' where "r' = r / 2" + have "0 < r'" "r' < r" using \<open>r > 0\<close> by (simp_all add: r'_def) + from CauchyD[OF \<open>Cauchy X\<close> \<open>r' > 0\<close>] + obtain M where M: "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> norm (X m - X n) < r'" + by metis + show "\<exists>no. \<forall>n\<ge>no. norm (X n - Blinfun v) < r" + proof (safe intro!: exI[where x=M]) + fix n assume n: "M \<le> n" + have "norm (X n - Blinfun v) \<le> r'" + proof (rule norm_blinfun_bound) + fix x + have "eventually (\<lambda>m. m \<ge> M) sequentially" + by (metis eventually_ge_at_top) + hence ev_le: "eventually (\<lambda>m. norm (X n x - X m x) \<le> r' * norm x) sequentially" + proof eventually_elim + case (elim m) + have "norm (X n x - X m x) = norm ((X n - X m) x)" + by (simp add: blinfun.bilinear_simps) + also have "\<dots> \<le> norm ((X n - X m)) * norm x" + by (rule norm_blinfun) + also have "\<dots> \<le> r' * norm x" + using M[OF n elim] by (simp add: mult_right_mono) + finally show ?case . + qed + have tendsto_v: "(\<lambda>m. norm (X n x - X m x)) \<longlonglongrightarrow> norm (X n x - Blinfun v x)" + by (auto intro!: tendsto_intros Bv) + show "norm ((X n - Blinfun v) x) \<le> r' * norm x" + by (auto intro!: tendsto_ge_const tendsto_v ev_le simp: blinfun.bilinear_simps) + qed (simp add: \<open>0 < r'\<close> less_imp_le) + thus "norm (X n - Blinfun v) < r" + by (metis \<open>r' < r\<close> le_less_trans) + qed + qed + thus "convergent X" + by (rule convergentI) +qed + +subsection \<open>On Euclidean Space\<close> + +lemma Zfun_setsum: + assumes "finite s" + assumes f: "\<And>i. i \<in> s \<Longrightarrow> Zfun (f i) F" + shows "Zfun (\<lambda>x. setsum (\<lambda>i. f i x) s) F" + using assms by induct (auto intro!: Zfun_zero Zfun_add) + +lemma norm_blinfun_euclidean_le: + fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::real_normed_vector" + shows "norm a \<le> setsum (\<lambda>x. norm (a x)) Basis" + apply (rule norm_blinfun_bound) + apply (simp add: setsum_nonneg) + apply (subst euclidean_representation[symmetric, where 'a='a]) + apply (simp only: blinfun.bilinear_simps setsum_left_distrib) + apply (rule order.trans[OF norm_setsum setsum_mono]) + apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm) + done + +lemma tendsto_componentwise1: + fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::real_normed_vector" + and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + assumes "(\<And>j. j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n j) \<longlongrightarrow> a j) F)" + shows "(b \<longlongrightarrow> a) F" +proof - + have "\<And>j. j \<in> Basis \<Longrightarrow> Zfun (\<lambda>x. norm (b x j - a j)) F" + using assms unfolding tendsto_Zfun_iff Zfun_norm_iff . + hence "Zfun (\<lambda>x. \<Sum>j\<in>Basis. norm (b x j - a j)) F" + by (auto intro!: Zfun_setsum) + thus ?thesis + unfolding tendsto_Zfun_iff + by (rule Zfun_le) + (auto intro!: order_trans[OF norm_blinfun_euclidean_le] simp: blinfun.bilinear_simps) +qed + +lift_definition + blinfun_of_matrix::"('b::euclidean_space \<Rightarrow> 'a::euclidean_space \<Rightarrow> real) \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + is "\<lambda>a x. \<Sum>i\<in>Basis. \<Sum>j\<in>Basis. ((x \<bullet> j) * a i j) *\<^sub>R i" + by (intro bounded_linear_intros) + +lemma blinfun_of_matrix_works: + fixes f::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::euclidean_space" + shows "blinfun_of_matrix (\<lambda>i j. (f j) \<bullet> i) = f" +proof (transfer, rule, rule euclidean_eqI) + fix f::"'a \<Rightarrow> 'b" and x::'a and b::'b assume "bounded_linear f" and b: "b \<in> Basis" + then interpret bounded_linear f by simp + have "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b + = (\<Sum>j\<in>Basis. if j = b then (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j))) else 0)" + using b + by (auto simp add: algebra_simps inner_setsum_left inner_Basis split: if_split intro!: setsum.cong) + also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> b)))" + using b by (simp add: setsum.delta) + also have "\<dots> = f x \<bullet> b" + by (subst linear_componentwise[symmetric]) (unfold_locales, rule) + finally show "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b = f x \<bullet> b" . +qed + +lemma blinfun_of_matrix_apply: + "blinfun_of_matrix a x = (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. ((x \<bullet> j) * a i j) *\<^sub>R i)" + by transfer simp + +lemma blinfun_of_matrix_minus: "blinfun_of_matrix x - blinfun_of_matrix y = blinfun_of_matrix (x - y)" + by transfer (auto simp: algebra_simps setsum_subtractf) + +lemma norm_blinfun_of_matrix: + "norm (blinfun_of_matrix a) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>a i j\<bar>)" + apply (rule norm_blinfun_bound) + apply (simp add: setsum_nonneg) + apply (simp only: blinfun_of_matrix_apply setsum_left_distrib) + apply (rule order_trans[OF norm_setsum setsum_mono]) + apply (rule order_trans[OF norm_setsum setsum_mono]) + apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm) + done + +lemma tendsto_blinfun_of_matrix: + assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n i j) \<longlongrightarrow> a i j) F" + shows "((\<lambda>n. blinfun_of_matrix (b n)) \<longlongrightarrow> blinfun_of_matrix a) F" +proof - + have "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> Zfun (\<lambda>x. norm (b x i j - a i j)) F" + using assms unfolding tendsto_Zfun_iff Zfun_norm_iff . + hence "Zfun (\<lambda>x. (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>b x i j - a i j\<bar>)) F" + by (auto intro!: Zfun_setsum) + thus ?thesis + unfolding tendsto_Zfun_iff blinfun_of_matrix_minus + by (rule Zfun_le) (auto intro!: order_trans[OF norm_blinfun_of_matrix]) +qed + +lemma tendsto_componentwise: + fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::euclidean_space" + and b::"'c \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + shows "(\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> ((\<lambda>n. b n j \<bullet> i) \<longlongrightarrow> a j \<bullet> i) F) \<Longrightarrow> (b \<longlongrightarrow> a) F" + apply (subst blinfun_of_matrix_works[of a, symmetric]) + apply (subst blinfun_of_matrix_works[of "b x" for x, symmetric, abs_def]) + by (rule tendsto_blinfun_of_matrix) + +lemma + continuous_blinfun_componentwiseI: + fixes f:: "'b::t2_space \<Rightarrow> 'a::euclidean_space \<Rightarrow>\<^sub>L 'c::euclidean_space" + assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> continuous F (\<lambda>x. (f x) j \<bullet> i)" + shows "continuous F f" + using assms by (auto simp: continuous_def intro!: tendsto_componentwise) + +lemma + continuous_blinfun_componentwiseI1: + fixes f:: "'b::t2_space \<Rightarrow> 'a::euclidean_space \<Rightarrow>\<^sub>L 'c::real_normed_vector" + assumes "\<And>i. i \<in> Basis \<Longrightarrow> continuous F (\<lambda>x. f x i)" + shows "continuous F f" + using assms by (auto simp: continuous_def intro!: tendsto_componentwise1) + +lemma bounded_linear_blinfun_matrix: "bounded_linear (\<lambda>x. (x::_\<Rightarrow>\<^sub>L _) j \<bullet> i)" + by (auto intro!: bounded_linearI' bounded_linear_intros) + +lemma continuous_blinfun_matrix: + fixes f:: "'b::t2_space \<Rightarrow> 'a::real_normed_vector \<Rightarrow>\<^sub>L 'c::real_inner" + assumes "continuous F f" + shows "continuous F (\<lambda>x. (f x) j \<bullet> i)" + by (rule bounded_linear.continuous[OF bounded_linear_blinfun_matrix assms]) + +lemma continuous_on_blinfun_matrix: + fixes f::"'a::t2_space \<Rightarrow> 'b::real_normed_vector \<Rightarrow>\<^sub>L 'c::real_inner" + assumes "continuous_on S f" + shows "continuous_on S (\<lambda>x. (f x) j \<bullet> i)" + using assms + by (auto simp: continuous_on_eq_continuous_within continuous_blinfun_matrix) + +lemma continuous_on_blinfun_of_matrix[continuous_intros]: + assumes "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> continuous_on S (\<lambda>s. g s i j)" + shows "continuous_on S (\<lambda>s. blinfun_of_matrix (g s))" + using assms + by (auto simp: continuous_on intro!: tendsto_blinfun_of_matrix) + +lemma mult_if_delta: + "(if P then (1::'a::comm_semiring_1) else 0) * q = (if P then q else 0)" + by auto + +lemma compact_blinfun_lemma: + fixes f :: "nat \<Rightarrow> 'a::euclidean_space \<Rightarrow>\<^sub>L 'b::euclidean_space" + assumes "bounded (range f)" + shows "\<forall>d\<subseteq>Basis. \<exists>l::'a \<Rightarrow>\<^sub>L 'b. \<exists> r. + subseq r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) i) (l i) < e) sequentially)" + by (rule compact_lemma_general[where unproj = "\<lambda>e. blinfun_of_matrix (\<lambda>i j. e j \<bullet> i)"]) + (auto intro!: euclidean_eqI[where 'a='b] bounded_linear_image assms + simp: blinfun_of_matrix_works blinfun_of_matrix_apply inner_Basis mult_if_delta setsum.delta' + scaleR_setsum_left[symmetric]) + +lemma blinfun_euclidean_eqI: "(\<And>i. i \<in> Basis \<Longrightarrow> blinfun_apply x i = blinfun_apply y i) \<Longrightarrow> x = y" + apply (auto intro!: blinfun_eqI) + apply (subst (2) euclidean_representation[symmetric, where 'a='a]) + apply (subst (1) euclidean_representation[symmetric, where 'a='a]) + apply (simp add: blinfun.bilinear_simps) + done + +lemma Blinfun_eq_matrix: "bounded_linear f \<Longrightarrow> Blinfun f = blinfun_of_matrix (\<lambda>i j. f j \<bullet> i)" + by (intro blinfun_euclidean_eqI) + (auto simp: blinfun_of_matrix_apply bounded_linear_Blinfun_apply inner_Basis if_distrib + cond_application_beta setsum.delta' euclidean_representation + cong: if_cong) + +text \<open>TODO: generalize (via @{thm compact_cball})?\<close> +instance blinfun :: (euclidean_space, euclidean_space) heine_borel +proof + fix f :: "nat \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" + assume f: "bounded (range f)" + then obtain l::"'a \<Rightarrow>\<^sub>L 'b" and r where r: "subseq r" + and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) i) (l i) < e) sequentially" + using compact_blinfun_lemma [OF f] by blast + { + fix e::real + let ?d = "real_of_nat DIM('a) * real_of_nat DIM('b)" + assume "e > 0" + hence "e / ?d > 0" by (simp add: DIM_positive) + with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) i) (l i) < e / ?d) sequentially" + by simp + moreover + { + fix n + assume n: "\<forall>i\<in>Basis. dist (f (r n) i) (l i) < e / ?d" + have "norm (f (r n) - l) = norm (blinfun_of_matrix (\<lambda>i j. (f (r n) - l) j \<bullet> i))" + unfolding blinfun_of_matrix_works .. + also note norm_blinfun_of_matrix + also have "(\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < + (\<Sum>i\<in>(Basis::'b set). e / real_of_nat DIM('b))" + proof (rule setsum_strict_mono) + fix i::'b assume i: "i \<in> Basis" + have "(\<Sum>j::'a\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < (\<Sum>j::'a\<in>Basis. e / ?d)" + proof (rule setsum_strict_mono) + fix j::'a assume j: "j \<in> Basis" + have "\<bar>(f (r n) - l) j \<bullet> i\<bar> \<le> norm ((f (r n) - l) j)" + by (simp add: Basis_le_norm i) + also have "\<dots> < e / ?d" + using n i j by (auto simp: dist_norm blinfun.bilinear_simps) + finally show "\<bar>(f (r n) - l) j \<bullet> i\<bar> < e / ?d" by simp + qed simp_all + also have "\<dots> \<le> e / real_of_nat DIM('b)" + by simp + finally show "(\<Sum>j\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < e / real_of_nat DIM('b)" + by simp + qed simp_all + also have "\<dots> \<le> e" by simp + finally have "dist (f (r n)) l < e" + by (auto simp: dist_norm) + } + ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" + using eventually_elim2 by force + } + then have *: "((f \<circ> r) \<longlongrightarrow> l) sequentially" + unfolding o_def tendsto_iff by simp + with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" + by auto +qed + + +subsection \<open>concrete bounded linear functions\<close> + +lemma transfer_bounded_bilinear_bounded_linearI: + assumes "g = (\<lambda>i x. (blinfun_apply (f i) x))" + shows "bounded_bilinear g = bounded_linear f" +proof + assume "bounded_bilinear g" + then interpret bounded_bilinear f by (simp add: assms) + show "bounded_linear f" + proof (unfold_locales, safe intro!: blinfun_eqI) + fix i + show "f (x + y) i = (f x + f y) i" "f (r *\<^sub>R x) i = (r *\<^sub>R f x) i" for r x y + by (auto intro!: blinfun_eqI simp: blinfun.bilinear_simps) + from _ nonneg_bounded show "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K" + by (rule ex_reg) (auto intro!: onorm_bound simp: norm_blinfun.rep_eq ac_simps) + qed +qed (auto simp: assms intro!: blinfun.comp) + +lemma transfer_bounded_bilinear_bounded_linear[transfer_rule]: + "(rel_fun (rel_fun op = (pcr_blinfun op = op =)) op =) bounded_bilinear bounded_linear" + by (auto simp: pcr_blinfun_def cr_blinfun_def rel_fun_def OO_def + intro!: transfer_bounded_bilinear_bounded_linearI) + +context bounded_bilinear +begin + +lift_definition prod_left::"'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'c" is "(\<lambda>b a. prod a b)" + by (rule bounded_linear_left) +declare prod_left.rep_eq[simp] + +lemma bounded_linear_prod_left[bounded_linear]: "bounded_linear prod_left" + by transfer (rule flip) + +lift_definition prod_right::"'a \<Rightarrow> 'b \<Rightarrow>\<^sub>L 'c" is "(\<lambda>a b. prod a b)" + by (rule bounded_linear_right) +declare prod_right.rep_eq[simp] + +lemma bounded_linear_prod_right[bounded_linear]: "bounded_linear prod_right" + by transfer (rule bounded_bilinear_axioms) + +end + +lift_definition id_blinfun::"'a::real_normed_vector \<Rightarrow>\<^sub>L 'a" is "\<lambda>x. x" + by (rule bounded_linear_ident) + +lemmas blinfun_apply_id_blinfun[simp] = id_blinfun.rep_eq + +lemma norm_blinfun_id[simp]: + "norm (id_blinfun::'a::{real_normed_vector, perfect_space} \<Rightarrow>\<^sub>L 'a) = 1" + by transfer (auto simp: onorm_id) + +lemma norm_blinfun_id_le: + "norm (id_blinfun::'a::real_normed_vector \<Rightarrow>\<^sub>L 'a) \<le> 1" + by transfer (auto simp: onorm_id_le) + + +lift_definition fst_blinfun::"('a::real_normed_vector \<times> 'b::real_normed_vector) \<Rightarrow>\<^sub>L 'a" is fst + by (rule bounded_linear_fst) + +lemma blinfun_apply_fst_blinfun[simp]: "blinfun_apply fst_blinfun = fst" + by transfer (rule refl) + + +lift_definition snd_blinfun::"('a::real_normed_vector \<times> 'b::real_normed_vector) \<Rightarrow>\<^sub>L 'b" is snd + by (rule bounded_linear_snd) + +lemma blinfun_apply_snd_blinfun[simp]: "blinfun_apply snd_blinfun = snd" + by transfer (rule refl) + + +lift_definition blinfun_compose:: + "'a::real_normed_vector \<Rightarrow>\<^sub>L 'b::real_normed_vector \<Rightarrow> + 'c::real_normed_vector \<Rightarrow>\<^sub>L 'a \<Rightarrow> + 'c \<Rightarrow>\<^sub>L 'b" (infixl "o\<^sub>L" 55) is "op o" + parametric comp_transfer + unfolding o_def + by (rule bounded_linear_compose) + +lemma blinfun_apply_blinfun_compose[simp]: "(a o\<^sub>L b) c = a (b c)" + by (simp add: blinfun_compose.rep_eq) + +lemma norm_blinfun_compose: + "norm (f o\<^sub>L g) \<le> norm f * norm g" + by transfer (rule onorm_compose) + +lemma bounded_bilinear_blinfun_compose[bounded_bilinear]: "bounded_bilinear op o\<^sub>L" + by unfold_locales + (auto intro!: blinfun_eqI exI[where x=1] simp: blinfun.bilinear_simps norm_blinfun_compose) + +lemma blinfun_compose_zero[simp]: + "blinfun_compose 0 = (\<lambda>_. 0)" + "blinfun_compose x 0 = 0" + by (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI) + + +lift_definition blinfun_inner_right::"'a::real_inner \<Rightarrow> 'a \<Rightarrow>\<^sub>L real" is "op \<bullet>" + by (rule bounded_linear_inner_right) +declare blinfun_inner_right.rep_eq[simp] + +lemma bounded_linear_blinfun_inner_right[bounded_linear]: "bounded_linear blinfun_inner_right" + by transfer (rule bounded_bilinear_inner) + + +lift_definition blinfun_inner_left::"'a::real_inner \<Rightarrow> 'a \<Rightarrow>\<^sub>L real" is "\<lambda>x y. y \<bullet> x" + by (rule bounded_linear_inner_left) +declare blinfun_inner_left.rep_eq[simp] + +lemma bounded_linear_blinfun_inner_left[bounded_linear]: "bounded_linear blinfun_inner_left" + by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_inner]) + + +lift_definition blinfun_scaleR_right::"real \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a::real_normed_vector" is "op *\<^sub>R" + by (rule bounded_linear_scaleR_right) +declare blinfun_scaleR_right.rep_eq[simp] + +lemma bounded_linear_blinfun_scaleR_right[bounded_linear]: "bounded_linear blinfun_scaleR_right" + by transfer (rule bounded_bilinear_scaleR) + + +lift_definition blinfun_scaleR_left::"'a::real_normed_vector \<Rightarrow> real \<Rightarrow>\<^sub>L 'a" is "\<lambda>x y. y *\<^sub>R x" + by (rule bounded_linear_scaleR_left) +lemmas [simp] = blinfun_scaleR_left.rep_eq + +lemma bounded_linear_blinfun_scaleR_left[bounded_linear]: "bounded_linear blinfun_scaleR_left" + by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_scaleR]) + + +lift_definition blinfun_mult_right::"'a \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a::real_normed_algebra" is "op *" + by (rule bounded_linear_mult_right) +declare blinfun_mult_right.rep_eq[simp] + +lemma bounded_linear_blinfun_mult_right[bounded_linear]: "bounded_linear blinfun_mult_right" + by transfer (rule bounded_bilinear_mult) + + +lift_definition blinfun_mult_left::"'a::real_normed_algebra \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'a" is "\<lambda>x y. y * x" + by (rule bounded_linear_mult_left) +lemmas [simp] = blinfun_mult_left.rep_eq + +lemma bounded_linear_blinfun_mult_left[bounded_linear]: "bounded_linear blinfun_mult_left" + by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_mult]) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,4164 @@ +(* Author: John Harrison + Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP +*) + +(* ========================================================================= *) +(* Results connected with topological dimension. *) +(* *) +(* 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 \<open>Results connected with topological dimension.\<close> + +theory Brouwer_Fixpoint +imports Path_Connected Homeomorphism +begin + +lemma bij_betw_singleton_eq: + assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \<in> A" + assumes eq: "(\<And>x. x \<in> A \<Longrightarrow> x \<noteq> a \<Longrightarrow> 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 Diff_subset) + moreover have "f a \<in> B" "g a \<in> 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 \<in> A then (if j \<in> A then f ` A else f ` ((A - {i}) \<union> {j})) + else (if j \<in> A then f ` ((A - {j}) \<union> {i}) else f ` A))" + apply (auto simp: Fun.swap_def image_iff) + apply metis + apply (metis member_remove remove_def) + apply (metis member_remove remove_def) + done + +lemmas swap_apply1 = swap_apply(1) +lemmas swap_apply2 = swap_apply(2) +lemmas lessThan_empty_iff = Iio_eq_empty_iff_nat +lemmas Zero_notin_Suc = zero_notin_Suc_image +lemmas atMost_Suc_eq_insert_0 = Iic_Suc_eq_insert_0 + +lemma setsum_union_disjoint': + assumes "finite A" + and "finite B" + and "A \<inter> B = {}" + and "A \<union> B = C" + shows "setsum g C = setsum g A + setsum g B" + using setsum.union_disjoint[OF assms(1-3)] and assms(4) by auto + +lemma pointwise_minimal_pointwise_maximal: + fixes s :: "(nat \<Rightarrow> nat) set" + assumes "finite s" + and "s \<noteq> {}" + and "\<forall>x\<in>s. \<forall>y\<in>s. x \<le> y \<or> y \<le> x" + shows "\<exists>a\<in>s. \<forall>x\<in>s. a \<le> x" + and "\<exists>a\<in>s. \<forall>x\<in>s. x \<le> a" + using assms +proof (induct s rule: finite_ne_induct) + case (insert b s) + assume *: "\<forall>x\<in>insert b s. \<forall>y\<in>insert b s. x \<le> y \<or> y \<le> x" + then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u" + using insert by auto + with * show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a" + using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+ +qed auto + +lemma brouwer_compactness_lemma: + fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" + assumes "compact s" + and "continuous_on s f" + and "\<not> (\<exists>x\<in>s. f x = 0)" + obtains d where "0 < d" and "\<forall>x\<in>s. d \<le> norm (f x)" +proof (cases "s = {}") + case True + show thesis + by (rule that [of 1]) (auto simp: True) +next + case False + have "continuous_on s (norm \<circ> f)" + by (rule continuous_intros continuous_on_norm assms(2))+ + with False obtain x where x: "x \<in> s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y" + using continuous_attains_inf[OF assms(1), of "norm \<circ> f"] + unfolding o_def + by auto + have "(norm \<circ> f) x > 0" + using assms(3) and x(1) + by auto + then show ?thesis + by (rule that) (insert x(2), auto simp: o_def) +qed + +lemma kuhn_labelling_lemma: + fixes P Q :: "'a::euclidean_space \<Rightarrow> bool" + assumes "\<forall>x. P x \<longrightarrow> P (f x)" + and "\<forall>x. P x \<longrightarrow> (\<forall>i\<in>Basis. Q i \<longrightarrow> 0 \<le> x\<bullet>i \<and> x\<bullet>i \<le> 1)" + shows "\<exists>l. (\<forall>x.\<forall>i\<in>Basis. l x i \<le> (1::nat)) \<and> + (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 0) \<longrightarrow> (l x i = 0)) \<and> + (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 1) \<longrightarrow> (l x i = 1)) \<and> + (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x\<bullet>i \<le> f x\<bullet>i) \<and> + (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f x\<bullet>i \<le> x\<bullet>i)" +proof - + { fix x i + let ?R = "\<lambda>y. (P x \<and> Q i \<and> x \<bullet> i = 0 \<longrightarrow> y = (0::nat)) \<and> + (P x \<and> Q i \<and> x \<bullet> i = 1 \<longrightarrow> y = 1) \<and> + (P x \<and> Q i \<and> y = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i) \<and> + (P x \<and> Q i \<and> y = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i)" + { assume "P x" "Q i" "i \<in> Basis" with assms have "0 \<le> f x \<bullet> i \<and> f x \<bullet> i \<le> 1" by auto } + then have "i \<in> Basis \<Longrightarrow> ?R 0 \<or> ?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 + + +subsection \<open>The key "counting" observation, somewhat abstracted.\<close> + +lemma kuhn_counting_lemma: + fixes bnd compo compo' face S F + defines "nF s == card {f\<in>F. face f s \<and> compo' f}" + assumes [simp, intro]: "finite F" \<comment> "faces" and [simp, intro]: "finite S" \<comment> "simplices" + and "\<And>f. f \<in> F \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 1" + and "\<And>f. f \<in> F \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 2" + and "\<And>s. s \<in> S \<Longrightarrow> compo s \<Longrightarrow> nF s = 1" + and "\<And>s. s \<in> S \<Longrightarrow> \<not> compo s \<Longrightarrow> nF s = 0 \<or> nF s = 2" + and "odd (card {f\<in>F. compo' f \<and> bnd f})" + shows "odd (card {s\<in>S. compo s})" +proof - + have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + (\<Sum>s | s \<in> S \<and> compo s. nF s) = (\<Sum>s\<in>S. nF s)" + by (subst setsum.union_disjoint[symmetric]) (auto intro!: setsum.cong) + also have "\<dots> = (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> bnd f}. face f s}) + + (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})" + unfolding setsum.distrib[symmetric] + by (subst card_Un_disjoint[symmetric]) + (auto simp: nF_def intro!: setsum.cong arg_cong[where f=card]) + also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}" + using assms(4,5) by (fastforce intro!: arg_cong2[where f="op +"] setsum_multicount) + finally have "odd ((\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + card {s\<in>S. compo s})" + using assms(6,8) by simp + moreover have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) = + (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 0. nF s) + (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 2. nF s)" + using assms(7) by (subst setsum.union_disjoint[symmetric]) (fastforce intro!: setsum.cong)+ + ultimately show ?thesis + by auto +qed + +subsection \<open>The odd/even result for faces of complete vertices, generalized.\<close> + +lemma kuhn_complete_lemma: + assumes [simp]: "finite simplices" + and face: "\<And>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})" + and card_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> card s = n + 2" + and rl_bd: "\<And>s. s \<in> simplices \<Longrightarrow> rl ` s \<subseteq> {..Suc n}" + and bnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 1" + and nbnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 2" + and odd_card: "odd (card {f. (\<exists>s\<in>simplices. face f s) \<and> rl ` f = {..n} \<and> bnd f})" + shows "odd (card {s\<in>simplices. (rl ` s = {..Suc n})})" +proof (rule kuhn_counting_lemma) + have finite_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> finite s" + by (metis add_is_0 zero_neq_numeral card_infinite assms(3)) + + let ?F = "{f. \<exists>s\<in>simplices. face f s}" + have F_eq: "?F = (\<Union>s\<in>simplices. \<Union>a\<in>s. {s - {a}})" + by (auto simp: face) + show "finite ?F" + using \<open>finite simplices\<close> unfolding F_eq by auto + + show "card {s \<in> simplices. face f s} = 1" if "f \<in> ?F" "bnd f" for f + using bnd that by auto + + show "card {s \<in> simplices. face f s} = 2" if "f \<in> ?F" "\<not> bnd f" for f + using nbnd that by auto + + show "odd (card {f \<in> {f. \<exists>s\<in>simplices. face f s}. rl ` f = {..n} \<and> bnd f})" + using odd_card by simp + + fix s assume s[simp]: "s \<in> simplices" + let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {..n}}" + have "?S = (\<lambda>a. s - {a}) ` {a\<in>s. rl ` (s - {a}) = {..n}}" + using s by (fastforce simp: face) + then have card_S: "card ?S = card {a\<in>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 \<in> s" + by (metis atMost_iff image_iff le_Suc_eq rl) + ultimately have n: "{..n} = rl ` (s - {a})" + by (auto simp add: inj_on_image_set_diff Diff_subset rl) + have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}" + using inj_rl \<open>a \<in> s\<close> by (auto simp add: n inj_on_image_eq_iff[OF inj_rl] Diff_subset) + then show "card ?S = 1" + unfolding card_S by simp } + + { assume rl: "rl ` s \<noteq> {..Suc n}" + show "card ?S = 0 \<or> card ?S = 2" + proof cases + assume *: "{..n} \<subseteq> rl ` s" + with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}" + by (auto simp add: atMost_Suc subset_insert_iff split: if_split_asm) + then have "\<not> inj_on rl s" + by (intro pigeonhole) simp + then obtain a b where ab: "a \<in> s" "b \<in> s" "rl a = rl b" "a \<noteq> 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 add: rl_s card_Diff_singleton_if) + + { fix x assume "x \<in> s" "x \<notin> {a, b}" + then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})" + by (auto simp: eq Diff_subset inj_on_image_set_diff[OF inj]) + also have "\<dots> = rl ` (s - {x})" + using ab \<open>x \<notin> {a, b}\<close> by auto + also assume "\<dots> = rl ` s" + finally have False + using \<open>x\<in>s\<close> by auto } + moreover + { fix x assume "x \<in> {a, b}" with ab have "x \<in> s \<and> rl ` (s - {x}) = rl ` s" + by (simp add: set_eq_iff image_iff Bex_def) metis } + ultimately have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a, b}" + unfolding rl_s[symmetric] by fastforce + with \<open>a \<noteq> b\<close> show "card ?S = 0 \<or> card ?S = 2" + unfolding card_S by simp + next + assume "\<not> {..n} \<subseteq> rl ` s" + then have "\<And>x. rl ` (s - {x}) \<noteq> {..n}" + by auto + then show "card ?S = 0 \<or> card ?S = 2" + unfolding card_S by simp + qed } +qed fact + +locale kuhn_simplex = + fixes p n and base upd and s :: "(nat \<Rightarrow> nat) set" + assumes base: "base \<in> {..< n} \<rightarrow> {..< p}" + assumes base_out: "\<And>i. n \<le> i \<Longrightarrow> base i = p" + assumes upd: "bij_betw upd {..< n} {..< n}" + assumes s_pre: "s = (\<lambda>i j. if j \<in> upd`{..< i} then Suc (base j) else base j) ` {.. n}" +begin + +definition "enum i j = (if j \<in> 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 \<Longrightarrow> upd i < n" + using upd by (auto dest!: bij_betwE) + +lemma s_space: "s \<subseteq> {..< n} \<rightarrow> {.. p}" +proof - + { fix i assume "i \<le> n" then have "enum i \<in> {..< n} \<rightarrow> {.. 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 \<noteq> y" "x \<le> n" "y \<le> n" + with upd have "upd ` {..< x} \<noteq> upd ` {..< y}" + by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def) + then have "enum x \<noteq> enum y" + by (auto simp add: 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 \<in> s" + unfolding s_eq by (subst enum_0[symmetric]) auto + +lemma enum_in: "i \<le> n \<Longrightarrow> enum i \<in> s" + unfolding s_eq by auto + +lemma one_step: + assumes a: "a \<in> s" "j < n" + assumes *: "\<And>a'. a' \<in> s \<Longrightarrow> a' \<noteq> a \<Longrightarrow> a' j = p'" + shows "a j \<noteq> p'" +proof + assume "a j = p'" + with * a have "\<And>a'. a' \<in> s \<Longrightarrow> a' j = p'" + by auto + then have "\<And>i. i \<le> n \<Longrightarrow> enum i j = p'" + unfolding s_eq by auto + from this[of 0] this[of n] have "j \<notin> upd ` {..< n}" + by (auto simp: enum_def fun_eq_iff split: if_split_asm) + with upd \<open>j < n\<close> show False + by (auto simp: bij_betw_def) +qed + +lemma upd_inj: "i < n \<Longrightarrow> j < n \<Longrightarrow> upd i = upd j \<longleftrightarrow> 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 \<subseteq> {..< n} \<Longrightarrow> i < n \<Longrightarrow> upd i \<in> upd ` A \<longleftrightarrow> i \<in> A" + using inj_on_image_mem_iff[of upd "{..< n}"] upd + by (auto simp: bij_betw_def) + +lemma enum_inj: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i = enum j \<longleftrightarrow> i = j" + using inj_enum by (auto simp: inj_on_eq_iff) + +lemma in_enum_image: "A \<subseteq> {.. n} \<Longrightarrow> i \<le> n \<Longrightarrow> enum i \<in> enum ` A \<longleftrightarrow> i \<in> A" + using inj_on_image_mem_iff[OF inj_enum] by auto + +lemma enum_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i \<le> enum j \<longleftrightarrow> i \<le> j" + by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric]) + +lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j" + using enum_mono[of i j] enum_inj[of i j] by (auto simp add: le_less) + +lemma chain: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a \<le> b \<or> b \<le> a" + by (auto simp: s_eq enum_mono) + +lemma less: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a i < b i \<Longrightarrow> a < b" + using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric]) + +lemma enum_0_bot: "a \<in> s \<Longrightarrow> a = enum 0 \<longleftrightarrow> (\<forall>a'\<in>s. a \<le> a')" + unfolding s_eq by (auto simp: enum_mono Ball_def) + +lemma enum_n_top: "a \<in> s \<Longrightarrow> a = enum n \<longleftrightarrow> (\<forall>a'\<in>s. a' \<le> a)" + unfolding s_eq by (auto simp: enum_mono Ball_def) + +lemma enum_Suc: "i < n \<Longrightarrow> 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 \<le> n \<Longrightarrow> n \<le> j \<Longrightarrow> 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 \<in> s \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p" + unfolding s_eq by (auto simp add: enum_eq_p) + +lemma s_le_p: "a \<in> s \<Longrightarrow> a j \<le> p" + using out_eq_p[of a j] s_space by (cases "j < n") auto + +lemma le_Suc_base: "a \<in> s \<Longrightarrow> a j \<le> Suc (base j)" + unfolding s_eq by (auto simp: enum_def) + +lemma base_le: "a \<in> s \<Longrightarrow> base j \<le> a j" + unfolding s_eq by (auto simp: enum_def) + +lemma enum_le_p: "i \<le> n \<Longrightarrow> j < n \<Longrightarrow> enum i j \<le> p" + using enum_in[of i] s_space by auto + +lemma enum_less: "a \<in> s \<Longrightarrow> i < n \<Longrightarrow> enum i < a \<longleftrightarrow> enum (Suc i) \<le> a" + unfolding s_eq by (auto simp: enum_strict_mono enum_mono) + +lemma ksimplex_0: + "n = 0 \<Longrightarrow> s = {(\<lambda>x. p)}" + using s_eq enum_def base_out by auto + +lemma replace_0: + assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = 0" and "x \<in> s" + shows "x \<le> a" +proof cases + assume "x \<noteq> a" + have "a j \<noteq> 0" + using assms by (intro one_step[where a=a]) auto + with less[OF \<open>x\<in>s\<close> \<open>a\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close> + show ?thesis + by auto +qed simp + +lemma replace_1: + assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = p" and "x \<in> s" + shows "a \<le> x" +proof cases + assume "x \<noteq> a" + have "a j \<noteq> p" + using assms by (intro one_step[where a=a]) auto + with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>s\<close> + have "a j < p" + by (auto simp: less_le s_eq) + with less[OF \<open>a\<in>s\<close> \<open>x\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close> + 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 \<le> l" "l \<le> j" and "j + d \<le> 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 \<in> t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \<in> s.enum ` {i .. j}" + using eq by auto + from t \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "s.enum i \<le> t.enum (i + d)" + by (auto simp: s.enum_mono) + moreover from s \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "t.enum (i + d) \<le> s.enum i" + by (auto simp: t.enum_mono) + ultimately show ?case + by auto +next + case (step l) + moreover from step.prems \<open>j + d \<le> n\<close> 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) \<in> t.enum ` {i + d .. j + d}" + "t.enum (Suc l + d) \<in> s.enum ` {i .. j}" + using step \<open>j + d \<le> n\<close> eq by (auto simp: s.enum_inj t.enum_inj) + ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" + using \<open>j + d \<le> n\<close> + 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 \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a \<le> a'" + assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b \<le> 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 \<noteq> 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 \<open>n \<noteq> 0\<close> 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 \<le> 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 \<open>n \<noteq> 0\<close> 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 "\<And>j. 0 < j \<Longrightarrow> j < n \<Longrightarrow> u_s j = u_t j" + by (auto simp: gr0_conv_Suc) + with \<open>n \<noteq> 0\<close> 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 \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a' \<le> a" + assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b' \<le> 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 "\<And>j. j < n' \<Longrightarrow> 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 \<Longrightarrow> 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 \<in> 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 \<open>a \<in> s\<close> out_eq_p[OF \<open>a \<in> s\<close>] + have "a \<in> (\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^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} \<subseteq> Pow ((\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^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" "\<forall>x\<in>s'. x n = p" + shows "ksimplex p n s' \<longleftrightarrow> (\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a})" + using assms +proof safe + fix s a assume "ksimplex p (Suc n) s" and a: "a \<in> s" and na: "\<forall>x\<in>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 \<open>a\<in>s\<close> s_space by (auto simp: less_le) + then have "a = enum 0" + using \<open>a \<in> s\<close> 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 Zero_notin_Suc in_enum_image subset_eq) + then have "enum 1 \<in> s - {a}" + by auto + then have "upd 0 = n" + using \<open>a n < p\<close> \<open>a = enum 0\<close> 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\<circ>Suc) {..<n} {..<n}" + by (rule bij_betw_trans[rotated]) (auto simp: bij_betw_def) + + have "a n = p - 1" + using enum_Suc[of 0] na[rule_format, OF \<open>enum 1 \<in> s - {a}\<close>] \<open>a = enum 0\<close> by (auto simp: \<open>upd 0 = n\<close>) + + show ?thesis + proof (rule ksimplex.intros, standard) + show "bij_betw (upd\<circ>Suc) {..< n} {..< n}" by fact + show "base(n := p) \<in> {..<n} \<rightarrow> {..<p}" "\<And>i. n\<le>i \<Longrightarrow> (base(n := p)) i = p" + using base base_out by (auto simp: Pi_iff) + + have "\<And>i. Suc ` {..< i} = {..< Suc i} - {0}" + by (auto simp: image_iff Ball_def) arith + then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}" + using \<open>upd 0 = n\<close> upd_inj + by (auto simp add: image_comp[symmetric] inj_on_image_set_diff[OF inj_upd]) + have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}" + using \<open>upd 0 = n\<close> by auto + + define f' where "f' i j = + (if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j + { fix x i assume i[arith]: "i \<le> n" then have "enum (Suc i) x = f' i x" + unfolding f'_def enum_def using \<open>a n < p\<close> \<open>a = enum 0\<close> \<open>upd 0 = n\<close> \<open>a n = p - 1\<close> + by (simp add: upd_Suc enum_0 n_in_upd) } + 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 *: "\<forall>x\<in>s'. x n = p" + then show "\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> 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 \<Rightarrow> n | Suc i \<Rightarrow> upd i)" for i + + have "ksimplex p (Suc n) (s' \<union> {b})" + proof (rule ksimplex.intros, standard) + show "b \<in> {..<Suc n} \<rightarrow> {..<p}" + using base \<open>0 < p\<close> unfolding lessThan_Suc b_def by (auto simp: PiE_iff) + show "\<And>i. Suc n \<le> i \<Longrightarrow> b i = p" + using base_out by (auto simp: b_def) + + have "bij_betw u (Suc ` {..< n} \<union> {0}) ({..<n} \<union> {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 {..<Suc n} {..<Suc n}" + by (simp add: u_def lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) + + define f' where "f' i j = (if j \<in> u`{..< i} then Suc (b j) else b j)" for i j + + have u_eq: "\<And>i. i \<le> n \<Longrightarrow> u ` {..< Suc i} = upd ` {..< i} \<union> { n }" + by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith + + { fix x have "x \<le> n \<Longrightarrow> n \<notin> upd ` {..<x}" + using upd_space by (simp add: image_iff neq_iff) } + note n_not_upd = this + + have *: "f' ` {.. Suc n} = f' ` (Suc ` {.. n} \<union> {0})" + unfolding atMost_Suc_eq_insert_0 by simp + also have "\<dots> = (f' \<circ> Suc) ` {.. n} \<union> {b}" + by (auto simp: f'_def) + also have "(f' \<circ> Suc) ` {.. n} = s'" + using \<open>0 < p\<close> 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' \<union> {b} = f' ` {.. Suc n}" .. + qed + moreover have "b \<notin> s'" + using * \<open>0 < p\<close> 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 \<in> s" + assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = 0" + shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>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 \<in> 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' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}" + using s \<open>a \<in> s\<close> by auto + then show ?thesis + by simp +qed + +lemma ksimplex_replace_1: + assumes s: "ksimplex p n s" and a: "a \<in> s" + assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = p" + shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>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 \<in> 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' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}" + using s \<open>a \<in> s\<close> by auto + then show ?thesis + by simp +qed + +lemma card_2_exists: "card s = 2 \<longleftrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y))" + by (auto simp add: card_Suc_eq eval_nat_numeral) + +lemma ksimplex_replace_2: + assumes s: "ksimplex p n s" and "a \<in> s" and "n \<noteq> 0" + and lb: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> 0" + and ub: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> p" + shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2" + using s +proof cases + case (ksimplex base upd) + then interpret kuhn_simplex p n base upd s . + + from \<open>a \<in> s\<close> obtain i where "i \<le> n" "a = enum i" + unfolding s_eq by auto + + from \<open>i \<le> n\<close> have "i = 0 \<or> i = n \<or> (0 < i \<and> i < n)" + by linarith + then have "\<exists>!s'. s' \<noteq> s \<and> ksimplex p n s' \<and> (\<exists>b\<in>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 \<circ> 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 {..<n} {..<n}" + by (rule bij_betw_trans) + + define f' where [abs_def]: "f' i j = + (if j \<in> ?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 \<circ> rot" "f' ` {.. n}" + proof + from \<open>a = enum i\<close> ub \<open>n \<noteq> 0\<close> \<open>i = 0\<close> + obtain i' where "i' \<le> n" "enum i' \<noteq> enum 0" "enum i' (upd 0) \<noteq> p" + unfolding s_eq by (auto intro: upd_space simp: enum_inj) + then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p" + using enum_le_p[of i' "upd 0"] by (auto simp add: enum_inj enum_mono upd_space) + then have "enum 1 (upd 0) < p" + by (auto simp add: le_fun_def intro: le_less_trans) + then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}" + using base \<open>n \<noteq> 0\<close> by (auto simp add: enum_0 enum_Suc PiE_iff extensional_def upd_space) + + { fix i assume "n \<le> i" then show "enum (Suc 0) i = p" + using \<open>n \<noteq> 0\<close> by (auto simp: enum_eq_p) } + show "bij_betw ?upd {..<n} {..<n}" by fact + qed (simp add: f'_def) + have ks_f': "ksimplex p n (f' ` {.. n})" + by rule unfold_locales + + have b_enum: "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. + with b.inj_enum have inj_f': "inj_on f' {.. n}" by simp + + have [simp]: "\<And>j. j < n \<Longrightarrow> rot ` {..< j} = {0 <..< Suc j}" + by (auto simp: rot_def image_iff Ball_def) + arith + + { fix j assume j: "j < n" + from j \<open>n \<noteq> 0\<close> have "f' j = enum (Suc j)" + by (auto simp add: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) } + note f'_eq_enum = this + 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 \<open>a = enum i\<close> \<open>i = 0\<close> + by (simp add: Diff_subset inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f']) + + have "enum 0 < f' 0" + using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono f'_eq_enum) + also have "\<dots> < f' n" + using \<open>n \<noteq> 0\<close> b.enum_strict_mono[of 0 n] unfolding b_enum by simp + finally have "a \<noteq> f' n" + using \<open>a = enum i\<close> \<open>i = 0\<close> by auto + + { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}" + obtain b u where "kuhn_simplex p n b u t" + using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) + then interpret t: kuhn_simplex p n b u t . + + { fix x assume "x \<in> s" "x \<noteq> a" + then have "x (upd 0) = enum (Suc 0) (upd 0)" + by (auto simp: \<open>a = enum i\<close> \<open>i = 0\<close> s_eq enum_def enum_inj) } + then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd 0) = enum (Suc 0) (upd 0)" + unfolding eq_sma[symmetric] by auto + then have "c (upd 0) \<noteq> enum (Suc 0) (upd 0)" + using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: upd_space) + then have "c (upd 0) < enum (Suc 0) (upd 0) \<or> c (upd 0) > enum (Suc 0) (upd 0)" + by auto + then have "t = s \<or> 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 \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x" + by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } + note top = this + have "s = t" + using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close> + 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 \<circ> rot" "f' ` {.. n}" b u t .. + have eq: "f' ` {..n} - {f' n} = t - {c}" + using eq_sma eq by simp + { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c" + by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } + note top = this + have "f' ` {..n} = t" + using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close> + 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 \<open>a \<noteq> f' n\<close> \<open>n \<noteq> 0\<close> show ?thesis + apply (intro ex1I[of _ "f' ` {.. n}"]) + apply auto [] + apply metis + done + next + assume "i = n" + from \<open>n \<noteq> 0\<close> obtain n' where n': "n = Suc n'" + by (cases n) auto + + define rot where "rot i = (case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i)" for i + let ?upd = "upd \<circ> 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 {..<n} {..<n}" + by (rule bij_betw_trans) + + define b where "b = base (upd n' := base (upd n') - 1)" + define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (b j) else b j)" for i j + + interpret b: kuhn_simplex p n b "upd \<circ> rot" "f' ` {.. n}" + proof + { fix i assume "n \<le> i" then show "b i = p" + using base_out[of i] upd_space[of n'] by (auto simp: b_def n') } + show "b \<in> {..<n} \<rightarrow> {..<p}" + using base \<open>n \<noteq> 0\<close> 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} {..<n}" by fact + qed (simp add: f'_def) + have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. + have ks_f': "ksimplex p n (b.enum ` {.. n})" + unfolding f' by rule unfold_locales + + have "0 < n" + using \<open>n \<noteq> 0\<close> by auto + + { from \<open>a = enum i\<close> \<open>n \<noteq> 0\<close> \<open>i = n\<close> lb upd_space[of n'] + obtain i' where "i' \<le> n" "enum i' \<noteq> 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 \<open>i' \<le> n\<close> \<open>enum i' \<noteq> enum n\<close> 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 \<open>0<n\<close>] b.enum_0 by (auto simp: b_def rot_def) + + have [simp]: "\<And>j. Suc j < n \<Longrightarrow> rot ` {..< Suc j} = {n'} \<union> {..< j}" + by (auto simp: rot_def image_iff Ball_def split: nat.splits) + have rot_simps: "\<And>j. rot (Suc j) = j" "rot 0 = n'" + by (simp_all add: rot_def) + + { fix j assume j: "Suc j \<le> n" then have "b.enum (Suc j) = enum j" + by (induct j) (auto simp add: 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 add: 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 \<open>a = enum i\<close> \<open>i = n\<close> + 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 \<le> b.enum n" + by (simp add: b.enum_mono) + also have "b.enum n < enum n" + using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono b_enum_eq_enum n') + finally have "a \<noteq> b.enum 0" + using \<open>a = enum i\<close> \<open>i = n\<close> by auto + + { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}" + obtain b' u where "kuhn_simplex p n b' u t" + using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) + then interpret t: kuhn_simplex p n b' u t . + + { fix x assume "x \<in> s" "x \<noteq> a" + then have "x (upd n') = enum n' (upd n')" + by (auto simp: \<open>a = enum i\<close> n' \<open>i = n\<close> s_eq enum_def enum_inj in_upd_image) } + then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd n') = enum n' (upd n')" + unfolding eq_sma[symmetric] by auto + then have "c (upd n') \<noteq> enum n' (upd n')" + using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: n' upd_space[unfolded n']) + then have "c (upd n') < enum n' (upd n') \<or> c (upd n') > enum n' (upd n')" + by auto + then have "t = s \<or> 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 \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c" + by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } + note top = this + have "s = t" + using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close> + 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 \<circ> 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 \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x" + by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } + note bot = this + have "f' ` {..n} = t" + using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close> + 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 \<open>a \<noteq> b.enum 0\<close> \<open>n \<noteq> 0\<close> 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 \<in> ?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 \<in> {..<n} \<rightarrow> {..<p}" by fact + { fix i assume "n \<le> i" then show "base i = p" by fact } + show "bij_betw ?upd {..<n} {..<n}" by fact + qed (simp add: f'_def) + have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. + have ks_f': "ksimplex p n (b.enum ` {.. n})" + unfolding f' by rule unfold_locales + + have "{i} \<subseteq> {..n}" + using i by auto + { fix j assume "j \<le> n" + moreover have "j < i \<or> i = j \<or> i < j" by arith + moreover note i + ultimately have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i" + unfolding enum_def[abs_def] b.enum_def[abs_def] + by (auto simp add: 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 \<open>a = enum i\<close> + using inj_on_image_set_diff[OF inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>] + inj_on_image_set_diff[OF b.inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>] + by (simp add: comp_def) + + have "a \<noteq> b.enum i" + using \<open>a = enum i\<close> enum_eq_benum i by auto + + { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}" + obtain b' u where "kuhn_simplex p n b' u t" + using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) + then interpret t: kuhn_simplex p n b' u t . + have "enum i' \<in> s - {a}" "enum (i + 1) \<in> s - {a}" + using \<open>a = enum i\<close> i enum_in by (auto simp: enum_inj i'_def) + then obtain l k where + l: "t.enum l = enum i'" "l \<le> n" "t.enum l \<noteq> c" and + k: "t.enum k = enum (i + 1)" "k \<le> n" "t.enum k \<noteq> 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 \<open>l \<le> n\<close> \<open>k \<le> n\<close> 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 \<open>Suc l = k\<close> i'_def) + then have False + using \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> + 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 \<open>l < k\<close> have "Suc l < k" + by arith + have c_eq: "c = t.enum (Suc l)" + proof (rule ccontr) + assume "c \<noteq> t.enum (Suc l)" + then have "t.enum (Suc l) \<in> s - {a}" + using \<open>l < k\<close> \<open>k \<le> n\<close> by (simp add: t.s_eq eq_sma) + then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i" + unfolding s_eq \<open>a = enum i\<close> by auto + with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)" + by (auto simp add: i'_def enum_mono enum_inj l k) + with \<open>Suc l < k\<close> \<open>k \<le> n\<close> show False + by (simp add: t.enum_mono) + qed + + { have "t.enum (Suc (Suc l)) \<in> s - {a}" + unfolding eq_sma c_eq t.s_eq using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_inj) + then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \<le> n" "j \<noteq> i" + by (auto simp: s_eq \<open>a = enum i\<close>) + moreover have "enum i' < t.enum (Suc (Suc l))" + unfolding l(1)[symmetric] using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_strict_mono) + ultimately have "i' < j" + using i by (simp add: enum_strict_mono i'_def) + with \<open>j \<noteq> i\<close> \<open>j \<le> n\<close> have "t.enum k \<le> t.enum (Suc (Suc l))" + unfolding i'_def by (simp add: enum_mono k eq) + then have "k \<le> Suc (Suc l)" + using \<open>k \<le> n\<close> \<open>Suc l < k\<close> by (simp add: t.enum_mono) } + with \<open>Suc l < k\<close> 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 "\<dots> = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" + using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (simp add: t.enum_Suc l t.upd_inj) + finally have "(u l = upd i' \<and> u (Suc l) = upd (Suc i')) \<or> + (u l = upd (Suc i') \<and> u (Suc l) = upd i')" + using \<open>Suc i' < n\<close> by (auto simp: enum_Suc fun_eq_iff split: if_split_asm) + + then have "t = s \<or> 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 \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> by (simp add: enum_Suc t.enum_Suc l) + also have "\<dots> = a" + using \<open>a = enum i\<close> i by (simp add: i'_def) + finally show ?thesis + using eq_sma \<open>a \<in> s\<close> \<open>c \<in> t\<close> 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 add: 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 \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> + by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \<open>b.enum i' = enum i'\<close> swap_apply1) + (simp add: Suc_i') + also have "\<dots> = b.enum i" + using i by (simp add: i'_def) + finally have "c = b.enum i" . + then have "t - {c} = B - {c}" "c \<in> B" + unfolding eq_sma[symmetric] eq B_def using i by auto + with \<open>c \<in> t\<close> have "t = B" + by auto + then show ?thesis + by (simp add: B_def) + qed } + with ks_f' eq \<open>a \<noteq> b.enum i\<close> \<open>n \<noteq> 0\<close> \<open>i \<le> n\<close> show ?thesis + apply (intro ex1I[of _ "b.enum ` {.. n}"]) + apply auto [] + apply metis + done + qed + then show ?thesis + using s \<open>a \<in> s\<close> by (simp add: card_2_exists Ex1_def) metis +qed + +text \<open>Hence another step towards concreteness.\<close> + +lemma kuhn_simplex_lemma: + assumes "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> rl ` s \<subseteq> {.. Suc n}" + and "odd (card {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> (f = s - {a}) \<and> + rl ` f = {..n} \<and> ((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p))})" + shows "odd (card {s. ksimplex p (Suc n) s \<and> rl ` s = {..Suc n}})" +proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq, + where bnd="\<lambda>f. (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p)"], + safe del: notI) + + have *: "\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)" + by auto + show "odd (card {f. (\<exists>s\<in>{s. ksimplex p (Suc n) s}. \<exists>a\<in>s. f = s - {a}) \<and> + rl ` f = {..n} \<and> ((\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>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 \<in> s" then show "rl a \<le> Suc n" + using assms(1) s by (auto simp: subset_eq) + + let ?S = "{t. ksimplex p (Suc n) t \<and> (\<exists>b\<in>t. s - {a} = t - {b})}" + { fix j assume j: "j \<le> n" "\<forall>x\<in>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 \<le> n" "\<forall>x\<in>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 \<noteq> 2" "\<not> (\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = p)" + with s a show "\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = 0" + using ksimplex_replace_2[of p "n + 1" s a] + by (subst (asm) eq_commute) auto } +qed + +subsection \<open>Reduced labelling\<close> + +definition reduced :: "nat \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> nat" where "reduced n x = (LEAST k. k = n \<or> x k \<noteq> 0)" + +lemma reduced_labelling: + shows "reduced n x \<le> n" + and "\<forall>i<reduced n x. x i = 0" + and "reduced n x = n \<or> x (reduced n x) \<noteq> 0" +proof - + show "reduced n x \<le> n" + unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto + show "\<forall>i<reduced n x. x i = 0" + unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ + show "reduced n x = n \<or> x (reduced n x) \<noteq> 0" + unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ +qed + +lemma reduced_labelling_unique: + "r \<le> n \<Longrightarrow> \<forall>i<r. x i = 0 \<Longrightarrow> r = n \<or> x r \<noteq> 0 \<Longrightarrow> 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 \<Longrightarrow> x j = 0 \<Longrightarrow> reduced n x \<noteq> 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 \<Longrightarrow> x j \<noteq> 0 \<Longrightarrow> reduced n x \<le> j" + using reduced_labelling[of n x] by (elim allE[where x=j]) auto + +lemma reduced_labelling_Suc: "reduced (Suc n) x \<noteq> Suc n \<Longrightarrow> 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 "\<forall>x\<in>f. \<forall>j\<le>n. x j = 0 \<longrightarrow> lab x j = 0" + and "\<forall>x\<in>f. \<forall>j\<le>n. x j = p \<longrightarrow> lab x j = 1" + and eq: "(reduced (Suc n) \<circ> lab) ` f = {..n}" + shows "((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p)) \<longleftrightarrow> (\<forall>x\<in>f. x n = p)" +proof (safe del: disjCI) + fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = 0" + { fix x assume "x \<in> f" with assms j have "reduced (Suc n) (lab x) \<noteq> j" + by (intro reduced_labelling_zero) auto } + moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f" + using j eq by auto + ultimately show "x n = p" + by force +next + fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = p" and x: "x \<in> f" + have "j = n" + proof (rule ccontr) + assume "\<not> ?thesis" + { fix x assume "x \<in> f" + with assms j have "reduced (Suc n) (lab x) \<le> j" + by (intro reduced_labelling_nonzero) auto + then have "reduced (Suc n) (lab x) \<noteq> n" + using \<open>j \<noteq> n\<close> \<open>j \<le> n\<close> by simp } + moreover + have "n \<in> (reduced (Suc n) \<circ> lab) ` f" + using eq by auto + ultimately show False + by force + qed + moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f" + using j eq by auto + ultimately show "x n = p" + using j x by auto +qed auto + +text \<open>Hence we get just about the nice induction.\<close> + +lemma kuhn_induction: + assumes "0 < p" + and lab_0: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0" + and lab_1: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1" + and odd: "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})" + shows "odd (card {s. ksimplex p (Suc n) s \<and> (reduced (Suc n)\<circ>lab) ` s = {..Suc n}})" +proof - + let ?rl = "reduced (Suc n) \<circ> lab" and ?ext = "\<lambda>f v. \<exists>j\<le>n. \<forall>x\<in>f. x j = v" + let ?ext = "\<lambda>s. (\<exists>j\<le>n. \<forall>x\<in>s. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>s. x j = p)" + have "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> ?rl ` s \<subseteq> {..Suc n}" + by (simp add: reduced_labelling subset_eq) + moreover + have "{s. ksimplex p n s \<and> (reduced n \<circ> lab) ` s = {..n}} = + {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> f = s - {a} \<and> ?rl ` f = {..n} \<and> ?ext f}" + proof (intro set_eqI, safe del: disjCI equalityI disjE) + fix s assume s: "ksimplex p n s" and rl: "(reduced n \<circ> 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: "\<forall>x\<in>s. x n = p" + by (auto simp: out_eq_p) + moreover + { fix x assume "x \<in> s" + with lab_1[rule_format, of n x] all_eq_p s_le_p[of x] + have "?rl x \<le> 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 \<in> t" "s = t - {a}" + using s unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] by auto + ultimately + show "\<exists>t a. ksimplex p (Suc n) t \<and> a \<in> t \<and> s = t - {a} \<and> ?rl ` s = {..n} \<and> ?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 \<in> 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: "\<forall>x\<in>s. x (Suc n) = p" + by (auto simp: out_eq_p) + + { fix x assume "x \<in> s - {a}" + then have "?rl x \<in> ?rl ` (s - {a})" + by auto + then have "?rl x \<le> n" + unfolding rl by auto + then have "?rl x = reduced n (lab x)" + by (auto intro!: reduced_labelling_Suc) } + then show rl': "(reduced n\<circ>lab) ` (s - {a}) = {..n}" + unfolding rl[symmetric] by (intro image_cong) auto + + from \<open>?ext (s - {a})\<close> + have all_eq_p: "\<forall>x\<in>s - {a}. x n = p" + proof (elim disjE exE conjE) + fix j assume "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0" + with lab_0[rule_format, of j] all_eq_p s_le_p + have "\<And>x. x \<in> s - {a} \<Longrightarrow> reduced (Suc n) (lab x) \<noteq> j" + by (intro reduced_labelling_zero) auto + moreover have "j \<in> ?rl ` (s - {a})" + using \<open>j \<le> n\<close> unfolding rl by auto + ultimately show ?thesis + by force + next + fix j assume "j \<le> n" and eq_p: "\<forall>x\<in>s - {a}. x j = p" + show ?thesis + proof cases + assume "j = n" with eq_p show ?thesis by simp + next + assume "j \<noteq> n" + { fix x assume x: "x \<in> s - {a}" + have "reduced n (lab x) \<le> j" + proof (rule reduced_labelling_nonzero) + show "lab x j \<noteq> 0" + using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \<open>j \<le> n\<close> by auto + show "j < n" + using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp + qed + then have "reduced n (lab x) \<noteq> n" + using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp } + moreover have "n \<in> (reduced n\<circ>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 \<open>0 < p\<close> all_eq_p] using s a by auto + qed + ultimately show ?thesis + using assms by (intro kuhn_simplex_lemma) auto +qed + +text \<open>And so we get the final combinatorial result.\<close> + +lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}" +proof + assume "ksimplex p 0 s" then show "s = {(\<lambda>x. p)}" + by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases) +next + assume s: "s = {(\<lambda>x. p)}" + show "ksimplex p 0 s" + proof (intro ksimplex, unfold_locales) + show "(\<lambda>_. p) \<in> {..<0::nat} \<rightarrow> {..<p}" by auto + show "bij_betw id {..<0} {..<0}" + by simp + qed (auto simp: s) +qed + +lemma kuhn_combinatorial: + assumes "0 < p" + and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = 0 \<longrightarrow> lab x j = 0" + and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = p \<longrightarrow> lab x j = 1" + shows "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>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 "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. label x i = (0::nat) \<or> label x i = 1)" + and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> label x i = 0)" + and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> label x i = 1)" + obtains q where "\<forall>i<n. q i < p" + and "\<forall>i<n. \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and> (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> label r i \<noteq> label s i" +proof - + let ?rl = "reduced n \<circ> label" + let ?A = "{s. ksimplex p n s \<and> ?rl ` s = {..n}}" + have "odd (card ?A)" + using assms by (intro kuhn_combinatorial[of p n label]) auto + then have "?A \<noteq> {}" + by fastforce + 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 \<in> {.. n}" "Suc i \<in> {.. n}" + by auto + then obtain u v where u: "u \<in> s" "Suc i = ?rl u" and v: "v \<in> s" "i = ?rl v" + unfolding rl[symmetric] by blast + + have "label u i \<noteq> label v i" + using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"] + u(2)[symmetric] v(2)[symmetric] \<open>i < n\<close> + by auto + moreover + have "b j \<le> u j" "u j \<le> b j + 1" "b j \<le> v j" "v j \<le> b j + 1" if "j < n" for j + using that base_le[OF \<open>u\<in>s\<close>] le_Suc_base[OF \<open>u\<in>s\<close>] base_le[OF \<open>v\<in>s\<close>] le_Suc_base[OF \<open>v\<in>s\<close>] + by auto + ultimately show "\<exists>r s. (\<forall>j<n. b j \<le> r j \<and> r j \<le> b j + 1) \<and> + (\<forall>j<n. b j \<le> s j \<and> s j \<le> b j + 1) \<and> label r i \<noteq> label s i" + by blast + qed +qed + +subsection \<open>The main result for the unit cube\<close> + +lemma kuhn_labelling_lemma': + assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))" + and "\<forall>x. P x \<longrightarrow> (\<forall>i::nat. Q i \<longrightarrow> 0 \<le> x i \<and> x i \<le> 1)" + shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and> + (\<forall>x i. P x \<and> Q i \<and> x i = 0 \<longrightarrow> l x i = 0) \<and> + (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and> + (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and> + (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)" +proof - + have and_forall_thm: "\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" + by auto + have *: "\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x" + by auto + show ?thesis + unfolding and_forall_thm + apply (subst choice_iff[symmetric])+ + apply rule + apply rule + proof - + fix x x' + let ?R = "\<lambda>y::nat. + (P x \<and> Q x' \<and> x x' = 0 \<longrightarrow> y = 0) \<and> + (P x \<and> Q x' \<and> x x' = 1 \<longrightarrow> y = 1) \<and> + (P x \<and> Q x' \<and> y = 0 \<longrightarrow> x x' \<le> (f x) x') \<and> + (P x \<and> Q x' \<and> y = 1 \<longrightarrow> (f x) x' \<le> x x')" + have "0 \<le> f x x' \<and> f x x' \<le> 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 \<or> ?R 1" + by auto + then show "\<exists>y\<le>1. ?R y" + by auto + qed +qed + +definition unit_cube :: "'a::euclidean_space set" + where "unit_cube = {x. \<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1}" + +lemma mem_unit_cube: "x \<in> unit_cube \<longleftrightarrow> (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)" + unfolding unit_cube_def by simp + +lemma bounded_unit_cube: "bounded unit_cube" + unfolding bounded_def +proof (intro exI ballI) + fix y :: 'a assume y: "y \<in> unit_cube" + have "dist 0 y = norm y" by (rule dist_0_norm) + also have "\<dots> = norm (\<Sum>i\<in>Basis. (y \<bullet> i) *\<^sub>R i)" unfolding euclidean_representation .. + also have "\<dots> \<le> (\<Sum>i\<in>Basis. norm ((y \<bullet> i) *\<^sub>R i))" by (rule norm_setsum) + also have "\<dots> \<le> (\<Sum>i::'a\<in>Basis. 1)" + by (rule setsum_mono, simp add: y [unfolded mem_unit_cube]) + finally show "dist 0 y \<le> (\<Sum>i::'a\<in>Basis. 1)" . +qed + +lemma closed_unit_cube: "closed unit_cube" + unfolding unit_cube_def Collect_ball_eq Collect_conj_eq + by (rule closed_INT, auto intro!: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + +lemma compact_unit_cube: "compact unit_cube" (is "compact ?C") + unfolding compact_eq_seq_compact_metric + using bounded_unit_cube closed_unit_cube + by (rule bounded_closed_imp_seq_compact) + +lemma brouwer_cube: + fixes f :: "'a::euclidean_space \<Rightarrow> 'a" + assumes "continuous_on unit_cube f" + and "f ` unit_cube \<subseteq> unit_cube" + shows "\<exists>x\<in>unit_cube. f x = x" +proof (rule ccontr) + define n where "n = DIM('a)" + have n: "1 \<le> n" "0 < n" "n \<noteq> 0" + unfolding n_def by (auto simp add: Suc_le_eq DIM_positive) + assume "\<not> ?thesis" + then have *: "\<not> (\<exists>x\<in>unit_cube. f x - x = 0)" + by auto + obtain d where + d: "d > 0" "\<And>x. x \<in> unit_cube \<Longrightarrow> d \<le> norm (f x - x)" + apply (rule brouwer_compactness_lemma[OF compact_unit_cube _ *]) + apply (rule continuous_intros assms)+ + apply blast + done + have *: "\<forall>x. x \<in> unit_cube \<longrightarrow> f x \<in> unit_cube" + "\<forall>x. x \<in> (unit_cube::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)" + using assms(2)[unfolded image_subset_iff Ball_def] + unfolding mem_unit_cube + by auto + obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where + "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1" + "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0" + "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1" + "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i" + "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i" + using kuhn_labelling_lemma[OF *] by blast + note label = this [rule_format] + have lem1: "\<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow> + \<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" + proof safe + fix x y :: 'a + assume x: "x \<in> unit_cube" + assume y: "y \<in> unit_cube" + fix i + assume i: "label x i \<noteq> label y i" "i \<in> Basis" + have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow> + \<bar>fx - x\<bar> \<le> \<bar>fy - fx\<bar> + \<bar>y - x\<bar>" by auto + have "\<bar>(f x - x) \<bullet> i\<bar> \<le> \<bar>(f y - f x)\<bullet>i\<bar> + \<bar>(y - x)\<bullet>i\<bar>" + unfolding inner_simps + apply (rule *) + apply (cases "label x i = 0") + apply (rule disjI1) + apply rule + prefer 3 + apply (rule disjI2) + apply rule + proof - + assume lx: "label x i = 0" + then have ly: "label y i = 1" + using i label(1)[of i y] + by auto + show "x \<bullet> i \<le> f x \<bullet> i" + apply (rule label(4)[rule_format]) + using x y lx i(2) + apply auto + done + show "f y \<bullet> i \<le> y \<bullet> i" + apply (rule label(5)[rule_format]) + using x y ly i(2) + apply auto + done + next + assume "label x i \<noteq> 0" + then have l: "label x i = 1" "label y i = 0" + using i label(1)[of i x] label(1)[of i y] + by auto + show "f x \<bullet> i \<le> x \<bullet> i" + apply (rule label(5)[rule_format]) + using x y l i(2) + apply auto + done + show "y \<bullet> i \<le> f y \<bullet> i" + apply (rule label(4)[rule_format]) + using x y l i(2) + apply auto + done + qed + also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" + apply (rule add_mono) + apply (rule Basis_le_norm[OF i(2)])+ + done + finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" + unfolding inner_simps . + qed + have "\<exists>e>0. \<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>z\<in>unit_cube. \<forall>i\<in>Basis. + norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> + \<bar>(f(z) - z)\<bullet>i\<bar> < d / (real n)" + proof - + have d': "d / real n / 8 > 0" + using d(1) by (simp add: n_def DIM_positive) + have *: "uniformly_continuous_on unit_cube f" + by (rule compact_uniformly_continuous[OF assms(1) compact_unit_cube]) + obtain e where e: + "e > 0" + "\<And>x x'. x \<in> unit_cube \<Longrightarrow> + x' \<in> unit_cube \<Longrightarrow> + norm (x' - x) < e \<Longrightarrow> + 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 + apply (rule_tac x="min (e/2) (d/real n/8)" in exI) + apply safe + proof - + show "0 < min (e / 2) (d / real n / 8)" + using d' e by auto + fix x y z i + assume as: + "x \<in> unit_cube" "y \<in> unit_cube" "z \<in> unit_cube" + "norm (x - z) < min (e / 2) (d / real n / 8)" + "norm (y - z) < min (e / 2) (d / real n / 8)" + "label x i \<noteq> label y i" + assume i: "i \<in> Basis" + have *: "\<And>z fz x fx n1 n2 n3 n4 d4 d :: real. \<bar>fx - x\<bar> \<le> n1 + n2 \<Longrightarrow> + \<bar>fx - fz\<bar> \<le> n3 \<Longrightarrow> \<bar>x - z\<bar> \<le> n4 \<Longrightarrow> + n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> + (8 * d4 = d) \<Longrightarrow> \<bar>fz - z\<bar> < d" + by auto + show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" + unfolding inner_simps + proof (rule *) + show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)" + apply (rule lem1[rule_format]) + using as i + apply auto + done + show "\<bar>f x \<bullet> i - f z \<bullet> i\<bar> \<le> norm (f x - f z)" "\<bar>x \<bullet> i - z \<bullet> i\<bar> \<le> norm (x - z)" + unfolding inner_diff_left[symmetric] + by (rule Basis_le_norm[OF i])+ + have tria: "norm (y - x) \<le> norm (y - z) + norm (x - z)" + using dist_triangle[of y x z, unfolded dist_norm] + unfolding norm_minus_commute + by auto + also have "\<dots> < e / 2 + e / 2" + apply (rule add_strict_mono) + using as(4,5) + apply auto + done + finally show "norm (f y - f x) < d / real n / 8" + apply - + apply (rule e(2)) + using as + apply auto + done + have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" + apply (rule add_strict_mono) + using as + apply auto + done + then show "norm (y - x) < 2 * (d / real n / 8)" + using tria + by auto + show "norm (f x - f z) < d / real n / 8" + apply (rule e(2)) + using as e(1) + apply auto + done + qed (insert as, auto) + qed + qed + then + obtain e where e: + "e > 0" + "\<And>x y z i. x \<in> unit_cube \<Longrightarrow> + y \<in> unit_cube \<Longrightarrow> + z \<in> unit_cube \<Longrightarrow> + i \<in> Basis \<Longrightarrow> + norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<Longrightarrow> + \<bar>(f z - z) \<bullet> i\<bar> < d / real n" + by blast + obtain p :: nat where p: "1 + real n / e \<le> 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 \<Rightarrow> '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: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {..< n}" + unfolding bij_betw_def by (auto simp: set_eq_iff) + have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i" + unfolding b'_def + using b + by (auto simp: f_inv_into_f bij_betw_def) + have b'b[simp]:"\<And>i. i < n \<Longrightarrow> b' (b i) = i" + unfolding b'_def + using b + by (auto simp: inv_into_f_eq bij_betw_def) + have *: "\<And>x :: nat. x = 0 \<or> x = 1 \<longleftrightarrow> x \<le> 1" + by auto + have b'': "\<And>j. j < n \<Longrightarrow> b j \<in> Basis" + using b unfolding bij_betw_def by auto + have q1: "0 < p" "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> + (\<forall>i<n. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or> + (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" + unfolding * + using \<open>p > 0\<close> \<open>n > 0\<close> + using label(1)[OF b''] + by auto + { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0" + then have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (unit_cube::'a set)" + using b'_Basis + by (auto simp add: mem_unit_cube inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) } + note cube = this + have q2: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> + (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)" + unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp add: b'') + have q3: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> + (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" + using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp add: b'') + obtain q where q: + "\<forall>i<n. q i < p" + "\<forall>i<n. + \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and> + (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> + (label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) i \<noteq> + (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) i" + by (rule kuhn_lemma[OF q1 q2 q3]) + define z :: 'a where "z = (\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)" + have "\<exists>i\<in>Basis. d / real n \<le> \<bar>(f z - z)\<bullet>i\<bar>" + proof (rule ccontr) + have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}" + using q(1) b' + by (auto intro: less_imp_le simp: bij_betw_def) + then have "z \<in> unit_cube" + unfolding z_def mem_unit_cube + using b'_Basis + by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1) + then have d_fz_z: "d \<le> norm (f z - z)" + by (rule d) + assume "\<not> ?thesis" + then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n" + using \<open>n > 0\<close> + by (auto simp add: not_le inner_diff) + have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)" + unfolding inner_diff_left[symmetric] + by (rule norm_le_l1) + also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)" + apply (rule setsum_strict_mono) + using as + apply auto + done + also have "\<dots> = 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 \<in> Basis" "d / real n \<le> \<bar>(f z - z) \<bullet> i\<bar>" .. + have *: "b' i < n" + using i and b'[unfolded bij_betw_def] + by auto + obtain r s where rs: + "\<And>j. j < n \<Longrightarrow> q j \<le> r j \<and> r j \<le> q j + 1" + "\<And>j. j < n \<Longrightarrow> q j \<le> s j \<and> s j \<le> q j + 1" + "(label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i) \<noteq> + (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i)" + using q(2)[rule_format,OF *] by blast + have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow> b' i < n" + using b' unfolding bij_betw_def by auto + define r' ::'a where "r' = (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)" + have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> 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 add: Suc_le_eq) + done + then have "r' \<in> unit_cube" + unfolding r'_def mem_unit_cube + using b'_Basis + by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1) + define s' :: 'a where "s' = (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)" + have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p" + apply (rule order_trans) + apply (rule rs(2)[OF b'_im, THEN conjunct2]) + using q(1)[rule_format,OF b'_im] + apply (auto simp add: Suc_le_eq) + done + then have "s' \<in> unit_cube" + unfolding s'_def mem_unit_cube + using b'_Basis + by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1) + have "z \<in> unit_cube" + unfolding z_def mem_unit_cube + using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close> + by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) + have *: "\<And>x. 1 + real x = real (Suc x)" + by auto + { + have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" + apply (rule setsum_mono) + using rs(1)[OF b'_im] + apply (auto simp add:* field_simps simp del: of_nat_Suc) + done + also have "\<dots> < e * real p" + using p \<open>e > 0\<close> \<open>p > 0\<close> + by (auto simp add: field_simps n_def) + finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . + } + moreover + { + have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" + apply (rule setsum_mono) + using rs(2)[OF b'_im] + apply (auto simp add:* field_simps simp del: of_nat_Suc) + done + also have "\<dots> < e * real p" + using p \<open>e > 0\<close> \<open>p > 0\<close> + by (auto simp add: field_simps n_def) + finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . + } + ultimately + have "norm (r' - z) < e" and "norm (s' - z) < e" + unfolding r'_def s'_def z_def + using \<open>p > 0\<close> + apply (rule_tac[!] le_less_trans[OF norm_le_l1]) + apply (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left) + done + then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" + using rs(3) i + unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' + by (intro e(2)[OF \<open>r'\<in>unit_cube\<close> \<open>s'\<in>unit_cube\<close> \<open>z\<in>unit_cube\<close>]) auto + then show False + using i by auto +qed + + +subsection \<open>Retractions\<close> + +definition "retraction s t r \<longleftrightarrow> t \<subseteq> s \<and> continuous_on s r \<and> r ` s \<subseteq> t \<and> (\<forall>x\<in>t. r x = x)" + +definition retract_of (infixl "retract'_of" 50) + where "(t retract_of s) \<longleftrightarrow> (\<exists>r. retraction s t r)" + +lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow> r (r x) = r x" + unfolding retraction_def by auto + +subsection \<open>Preservation of fixpoints under (more general notion of) retraction\<close> + +lemma invertible_fixpoint_property: + fixes s :: "'a::euclidean_space set" + and t :: "'b::euclidean_space set" + assumes "continuous_on t i" + and "i ` t \<subseteq> s" + and "continuous_on s r" + and "r ` s \<subseteq> t" + and "\<forall>y\<in>t. r (i y) = y" + and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" + and "continuous_on t g" + and "g ` t \<subseteq> t" + obtains y where "y \<in> t" and "g y = y" +proof - + have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x" + apply (rule assms(6)[rule_format]) + apply rule + apply (rule continuous_on_compose assms)+ + apply ((rule continuous_on_subset)?, rule assms)+ + using assms(2,4,8) + apply auto + apply blast + done + then obtain x where x: "x \<in> s" "(i \<circ> g \<circ> r) x = x" .. + then have *: "g (r x) \<in> t" + using assms(4,8) by auto + have "r ((i \<circ> g \<circ> r) x) = r x" + using x by auto + then show ?thesis + apply (rule_tac that[of "r x"]) + using x + unfolding o_def + unfolding assms(5)[rule_format,OF *] + using assms(4) + apply auto + done +qed + +lemma homeomorphic_fixpoint_property: + fixes s :: "'a::euclidean_space set" + and t :: "'b::euclidean_space set" + assumes "s homeomorphic t" + shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow> + (\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))" +proof - + obtain r i where + "\<forall>x\<in>s. i (r x) = x" + "r ` s = t" + "continuous_on s r" + "\<forall>y\<in>t. r (i y) = y" + "i ` t = s" + "continuous_on t i" + using assms + unfolding homeomorphic_def homeomorphism_def + by blast + then show ?thesis + apply - + apply rule + apply (rule_tac[!] allI impI)+ + apply (rule_tac g=g in invertible_fixpoint_property[of t i s r]) + prefer 10 + apply (rule_tac g=f in invertible_fixpoint_property[of s r t i]) + apply auto + done +qed + +lemma retract_fixpoint_property: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + and s :: "'a set" + assumes "t retract_of s" + and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" + and "continuous_on t g" + and "g ` t \<subseteq> t" + obtains y where "y \<in> t" and "g y = y" +proof - + obtain h where "retraction s t h" + using assms(1) unfolding retract_of_def .. + then show ?thesis + unfolding retraction_def + apply - + apply (rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g]) + prefer 7 + apply (rule_tac y = y in that) + using assms + apply auto + done +qed + + +subsection \<open>The Brouwer theorem for any set with nonempty interior\<close> + +lemma convex_unit_cube: "convex unit_cube" + apply (rule is_interval_convex) + apply (clarsimp simp add: is_interval_def mem_unit_cube) + apply (drule (1) bspec)+ + apply auto + done + +lemma brouwer_weak: + fixes f :: "'a::euclidean_space \<Rightarrow> 'a" + assumes "compact s" + and "convex s" + and "interior s \<noteq> {}" + and "continuous_on s f" + and "f ` s \<subseteq> s" + obtains x where "x \<in> s" and "f x = x" +proof - + let ?U = "unit_cube :: 'a set" + have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U" + proof (rule interiorI) + let ?I = "(\<Inter>i\<in>Basis. {x::'a. 0 < x \<bullet> i} \<inter> {x. x \<bullet> i < 1})" + show "open ?I" + by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner continuous_on_const continuous_on_id) + show "\<Sum>Basis /\<^sub>R 2 \<in> ?I" + by simp + show "?I \<subseteq> unit_cube" + unfolding unit_cube_def by force + qed + then have *: "interior ?U \<noteq> {}" by fast + have *: "?U homeomorphic s" + using homeomorphic_convex_compact[OF convex_unit_cube compact_unit_cube * assms(2,1,3)] . + have "\<forall>f. continuous_on ?U f \<and> f ` ?U \<subseteq> ?U \<longrightarrow> + (\<exists>x\<in>?U. f x = x)" + using brouwer_cube by auto + then show ?thesis + unfolding homeomorphic_fixpoint_property[OF *] + using assms + by (auto simp: intro: that) +qed + + +text \<open>And in particular for a closed ball.\<close> + +lemma brouwer_ball: + fixes f :: "'a::euclidean_space \<Rightarrow> 'a" + assumes "e > 0" + and "continuous_on (cball a e) f" + and "f ` cball a e \<subseteq> cball a e" + obtains x where "x \<in> 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 \<open>Still more general form; could derive this directly without using the + rather involved \<open>HOMEOMORPHIC_CONVEX_COMPACT\<close> theorem, just using + a scaling and translation to put the set inside the unit cube.\<close> + +lemma brouwer: + fixes f :: "'a::euclidean_space \<Rightarrow> 'a" + assumes "compact s" + and "convex s" + and "s \<noteq> {}" + and "continuous_on s f" + and "f ` s \<subseteq> s" + obtains x where "x \<in> s" and "f x = x" +proof - + have "\<exists>e>0. s \<subseteq> cball 0 e" + using compact_imp_bounded[OF assms(1)] + unfolding bounded_pos + apply (erule_tac exE) + apply (rule_tac x=b in exI) + apply (auto simp add: dist_norm) + done + then obtain e where e: "e > 0" "s \<subseteq> cball 0 e" + by blast + have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x" + apply (rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"]) + apply (rule continuous_on_compose ) + apply (rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)]) + apply (rule continuous_on_subset[OF assms(4)]) + apply (insert closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) + using assms(5)[unfolded subset_eq] + using e(2)[unfolded subset_eq mem_cball] + apply (auto simp add: dist_norm) + done + then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point s) x = x" .. + have *: "closest_point s x = x" + apply (rule closest_point_self) + apply (rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"], unfolded image_iff]) + apply (rule_tac x="closest_point s x" in bexI) + using x + unfolding o_def + using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x] + apply auto + done + show thesis + apply (rule_tac x="closest_point s x" in that) + unfolding x(2)[unfolded o_def] + apply (rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) + using * + apply auto + done +qed + +text \<open>So we get the no-retraction theorem.\<close> + +lemma no_retraction_cball: + fixes a :: "'a::euclidean_space" + assumes "e > 0" + shows "\<not> (frontier (cball a e) retract_of (cball a e))" +proof + assume *: "frontier (cball a e) retract_of (cball a e)" + have **: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" + using scaleR_left_distrib[of 1 1 a] by auto + obtain x where x: + "x \<in> {x. norm (a - x) = e}" + "2 *\<^sub>R a - x = x" + apply (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"]) + apply (blast intro: brouwer_ball[OF assms]) + apply (intro continuous_intros) + unfolding frontier_cball subset_eq Ball_def image_iff dist_norm sphere_def + apply (auto simp add: ** norm_minus_commute) + done + then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" + by (auto simp add: algebra_simps) + then have "a = x" + unfolding scaleR_left_distrib[symmetric] + by auto + then show False + using x assms by auto +qed + +subsection\<open>Retractions\<close> + +lemma retraction: + "retraction s t r \<longleftrightarrow> + t \<subseteq> s \<and> continuous_on s r \<and> r ` s = t \<and> (\<forall>x \<in> t. r x = x)" +by (force simp: retraction_def) + +lemma retract_of_imp_extensible: + assumes "s retract_of t" and "continuous_on s f" and "f ` s \<subseteq> u" + obtains g where "continuous_on t g" "g ` t \<subseteq> u" "\<And>x. x \<in> s \<Longrightarrow> g x = f x" +using assms +apply (clarsimp simp add: retract_of_def retraction) +apply (rule_tac g = "f o r" in that) +apply (auto simp: continuous_on_compose2) +done + +lemma idempotent_imp_retraction: + assumes "continuous_on s f" and "f ` s \<subseteq> s" and "\<And>x. x \<in> s \<Longrightarrow> f(f x) = f x" + shows "retraction s (f ` s) f" +by (simp add: assms retraction) + +lemma retraction_subset: + assumes "retraction s t r" and "t \<subseteq> s'" and "s' \<subseteq> s" + shows "retraction s' t r" +apply (simp add: retraction_def) +by (metis assms continuous_on_subset image_mono retraction) + +lemma retract_of_subset: + assumes "t retract_of s" and "t \<subseteq> s'" and "s' \<subseteq> s" + shows "t retract_of s'" +by (meson assms retract_of_def retraction_subset) + +lemma retraction_refl [simp]: "retraction s s (\<lambda>x. x)" +by (simp add: continuous_on_id retraction) + +lemma retract_of_refl [iff]: "s retract_of s" + using continuous_on_id retract_of_def retraction_def by fastforce + +lemma retract_of_imp_subset: + "s retract_of t \<Longrightarrow> s \<subseteq> t" +by (simp add: retract_of_def retraction_def) + +lemma retract_of_empty [simp]: + "({} retract_of s) \<longleftrightarrow> s = {}" "(s retract_of {}) \<longleftrightarrow> s = {}" +by (auto simp: retract_of_def retraction_def) + +lemma retract_of_singleton [iff]: "({x} retract_of s) \<longleftrightarrow> x \<in> s" + using continuous_on_const + by (auto simp: retract_of_def retraction_def) + +lemma retraction_comp: + "\<lbrakk>retraction s t f; retraction t u g\<rbrakk> + \<Longrightarrow> retraction s u (g o f)" +apply (auto simp: retraction_def intro: continuous_on_compose2) +by blast + +lemma retract_of_trans [trans]: + assumes "s retract_of t" and "t retract_of u" + shows "s retract_of u" +using assms by (auto simp: retract_of_def intro: retraction_comp) + +lemma closedin_retract: + fixes s :: "'a :: real_normed_vector set" + assumes "s retract_of t" + shows "closedin (subtopology euclidean t) s" +proof - + obtain r where "s \<subseteq> t" "continuous_on t r" "r ` t \<subseteq> s" "\<And>x. x \<in> s \<Longrightarrow> r x = x" + using assms by (auto simp: retract_of_def retraction_def) + then have s: "s = {x \<in> t. (norm(r x - x)) = 0}" by auto + show ?thesis + apply (subst s) + apply (rule continuous_closedin_preimage_constant) + by (simp add: \<open>continuous_on t r\<close> continuous_on_diff continuous_on_id continuous_on_norm) +qed + +lemma closedin_self [simp]: + fixes S :: "'a :: real_normed_vector set" + shows "closedin (subtopology euclidean S) S" + by (simp add: closedin_retract) + +lemma retract_of_contractible: + assumes "contractible t" "s retract_of t" + shows "contractible s" +using assms +apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with) +apply (rule_tac x="r a" in exI) +apply (rule_tac x="r o h" in exI) +apply (intro conjI continuous_intros continuous_on_compose) +apply (erule continuous_on_subset | force)+ +done + +lemma retract_of_compact: + "\<lbrakk>compact t; s retract_of t\<rbrakk> \<Longrightarrow> compact s" + by (metis compact_continuous_image retract_of_def retraction) + +lemma retract_of_closed: + fixes s :: "'a :: real_normed_vector set" + shows "\<lbrakk>closed t; s retract_of t\<rbrakk> \<Longrightarrow> closed s" + by (metis closedin_retract closedin_closed_eq) + +lemma retract_of_connected: + "\<lbrakk>connected t; s retract_of t\<rbrakk> \<Longrightarrow> connected s" + by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) + +lemma retract_of_path_connected: + "\<lbrakk>path_connected t; s retract_of t\<rbrakk> \<Longrightarrow> path_connected s" + by (metis path_connected_continuous_image retract_of_def retraction) + +lemma retract_of_simply_connected: + "\<lbrakk>simply_connected t; s retract_of t\<rbrakk> \<Longrightarrow> simply_connected s" +apply (simp add: retract_of_def retraction_def, clarify) +apply (rule simply_connected_retraction_gen) +apply (force simp: continuous_on_id elim!: continuous_on_subset)+ +done + +lemma retract_of_homotopically_trivial: + assumes ts: "t retract_of s" + and hom: "\<And>f g. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; + continuous_on u g; g ` u \<subseteq> s\<rbrakk> + \<Longrightarrow> homotopic_with (\<lambda>x. True) u s f g" + and "continuous_on u f" "f ` u \<subseteq> t" + and "continuous_on u g" "g ` u \<subseteq> t" + shows "homotopic_with (\<lambda>x. True) u t f g" +proof - + obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>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: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s\<rbrakk> + \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) u s f (\<lambda>x. c)" + and "continuous_on u f" "f ` u \<subseteq> t" + obtains c where "homotopic_with (\<lambda>x. True) u t f (\<lambda>x. c)" +proof - + obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>s. r (r x) = r x" "t = r ` s" + using ts by (auto simp: retract_of_def retraction) + then obtain k where "Retracts s r t k" + unfolding Retracts_def + by (metis continuous_on_subset dual_order.trans image_iff image_mono) + then show ?thesis + apply (rule Retracts.homotopically_trivial_retraction_null_gen) + apply (rule TrueI refl assms that | assumption)+ + done +qed + +lemma retraction_imp_quotient_map: + "retraction s t r + \<Longrightarrow> u \<subseteq> t + \<Longrightarrow> (openin (subtopology euclidean s) {x. x \<in> s \<and> r x \<in> u} \<longleftrightarrow> + openin (subtopology euclidean t) u)" +apply (clarsimp simp add: retraction) +apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) +apply (auto simp: elim: continuous_on_subset) +done + +lemma retract_of_locally_compact: + fixes s :: "'a :: {heine_borel,real_normed_vector} set" + shows "\<lbrakk> locally compact s; t retract_of s\<rbrakk> \<Longrightarrow> locally compact t" + by (metis locally_compact_closedin closedin_retract) + +lemma retract_of_Times: + "\<lbrakk>s retract_of s'; t retract_of t'\<rbrakk> \<Longrightarrow> (s \<times> t) retract_of (s' \<times> t')" +apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) +apply (rename_tac f g) +apply (rule_tac x="\<lambda>z. ((f o fst) z, (g o snd) z)" in exI) +apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ +done + +lemma homotopic_into_retract: + "\<lbrakk>f ` s \<subseteq> t; g ` s \<subseteq> t; t retract_of u; + homotopic_with (\<lambda>x. True) s u f g\<rbrakk> + \<Longrightarrow> homotopic_with (\<lambda>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 o h" in exI) +apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ +done + +lemma retract_of_locally_connected: + assumes "locally connected T" "S retract_of T" + shows "locally connected S" + using assms + by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_connected_quotient_image) + +lemma retract_of_locally_path_connected: + assumes "locally path_connected T" "S retract_of T" + shows "locally path_connected S" + using assms + by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image) + +subsection\<open>Punctured affine hulls, etc.\<close> + +lemma continuous_on_compact_surface_projection_aux: + fixes S :: "'a::t2_space set" + assumes "compact S" "S \<subseteq> T" "image q T \<subseteq> S" + and contp: "continuous_on T p" + and "\<And>x. x \<in> S \<Longrightarrow> q x = x" + and [simp]: "\<And>x. x \<in> T \<Longrightarrow> q(p x) = q x" + and "\<And>x. x \<in> T \<Longrightarrow> 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 \<open>S \<subseteq> T\<close>]) + have "continuous_on T (q \<circ> p)" + apply (rule continuous_on_compose [OF contp]) + apply (simp add: *) + apply (rule continuous_on_inv [OF contp' \<open>compact S\<close>]) + using assms by auto + then show ?thesis + apply (rule continuous_on_eq [of _ "q o p"]) + apply (simp add: o_def) + done +qed + +lemma continuous_on_compact_surface_projection: + fixes S :: "'a::real_normed_vector set" + assumes "compact S" + and S: "S \<subseteq> V - {0}" and "cone V" + and iff: "\<And>x k. x \<in> V - {0} \<Longrightarrow> 0 < k \<and> (k *\<^sub>R x) \<in> S \<longleftrightarrow> d x = k" + shows "continuous_on (V - {0}) (\<lambda>x. d x *\<^sub>R x)" +proof (rule continuous_on_compact_surface_projection_aux [OF \<open>compact S\<close> S]) + show "(\<lambda>x. d x *\<^sub>R x) ` (V - {0}) \<subseteq> S" + using iff by auto + show "continuous_on (V - {0}) (\<lambda>x. inverse(norm x) *\<^sub>R x)" + by (intro continuous_intros) force + show "\<And>x. x \<in> S \<Longrightarrow> 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 \<in> V - {0}" for x + using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \<open>cone V\<close> + 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 \<in> V - {0}" for x + proof - + have "0 < d x" + using local.iff that by blast + then show ?thesis + by simp + qed +qed + +proposition rel_frontier_deformation_retract_of_punctured_convex: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "convex T" "bounded S" + and arelS: "a \<in> rel_interior S" + and relS: "rel_frontier S \<subseteq> T" + and affS: "T \<subseteq> affine hull S" + obtains r where "homotopic_with (\<lambda>x. True) (T - {a}) (T - {a}) id r" + "retraction (T - {a}) (rel_frontier S) r" +proof - + have "\<exists>d. 0 < d \<and> (a + d *\<^sub>R l) \<in> rel_frontier S \<and> + (\<forall>e. 0 \<le> e \<and> e < d \<longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S)" + if "(a + l) \<in> affine hull S" "l \<noteq> 0" for l + apply (rule ray_to_rel_frontier [OF \<open>bounded S\<close> arelS]) + apply (rule that)+ + by metis + then obtain dd + where dd1: "\<And>l. \<lbrakk>(a + l) \<in> affine hull S; l \<noteq> 0\<rbrakk> \<Longrightarrow> 0 < dd l \<and> (a + dd l *\<^sub>R l) \<in> rel_frontier S" + and dd2: "\<And>l e. \<lbrakk>(a + l) \<in> affine hull S; e < dd l; 0 \<le> e; l \<noteq> 0\<rbrakk> + \<Longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S" + by metis+ + have aaffS: "a \<in> affine hull S" + by (meson arelS subsetD hull_inc rel_interior_subset) + have "((\<lambda>z. z - a) ` (affine hull S - {a})) = ((\<lambda>z. z - a) ` (affine hull S)) - {0}" + by (auto simp: ) + moreover have "continuous_on (((\<lambda>z. z - a) ` (affine hull S)) - {0}) (\<lambda>x. dd x *\<^sub>R x)" + proof (rule continuous_on_compact_surface_projection) + show "compact (rel_frontier ((\<lambda>z. z - a) ` S))" + by (simp add: \<open>bounded S\<close> bounded_translation_minus compact_rel_frontier_bounded) + have releq: "rel_frontier ((\<lambda>z. z - a) ` S) = (\<lambda>z. z - a) ` rel_frontier S" + using rel_frontier_translation [of "-a"] add.commute by simp + also have "... \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}" + using rel_frontier_affine_hull arelS rel_frontier_def by fastforce + finally show "rel_frontier ((\<lambda>z. z - a) ` S) \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}" . + show "cone ((\<lambda>z. z - a) ` (affine hull S))" + apply (rule subspace_imp_cone) + using aaffS + apply (simp add: subspace_affine image_comp o_def affine_translation_aux [of a]) + done + show "(0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)) \<longleftrightarrow> (dd x = k)" + if x: "x \<in> (\<lambda>z. z - a) ` (affine hull S) - {0}" for k x + proof + show "dd x = k \<Longrightarrow> 0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)" + using dd1 [of x] that image_iff by (fastforce simp add: releq) + next + assume k: "0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)" + have False if "dd x < k" + proof - + have "k \<noteq> 0" "a + k *\<^sub>R x \<in> closure S" + using k closure_translation [of "-a"] + by (auto simp: rel_frontier_def) + then have segsub: "open_segment a (a + k *\<^sub>R x) \<subseteq> rel_interior S" + by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS]) + have "x \<noteq> 0" and xaffS: "a + x \<in> affine hull S" + using x by (auto simp: ) + then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \<in> rel_frontier S" + using dd1 by auto + moreover have "a + dd x *\<^sub>R x \<in> open_segment a (a + k *\<^sub>R x)" + using k \<open>x \<noteq> 0\<close> \<open>0 < dd x\<close> + 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) + apply (metis (no_types) \<open>k \<noteq> 0\<close> divide_inverse_commute inverse_eq_divide mult.left_commute right_inverse) + done + ultimately show ?thesis + using segsub by (auto simp add: 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 ((\<lambda>z. z - a) ` (affine hull S - {a})) (\<lambda>x. dd x *\<^sub>R x)" + by auto + have "continuous_on (affine hull S - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))" + by (intro * continuous_intros continuous_on_compose) + with affS have contdd: "continuous_on (T - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))" + by (blast intro: continuous_on_subset elim: ) + show ?thesis + proof + show "homotopic_with (\<lambda>x. True) (T - {a}) (T - {a}) id (\<lambda>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}) (\<lambda>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)) \<subseteq> T - {a}" + if "x \<in> T - {a}" for x + proof (clarsimp simp: in_segment, intro conjI) + fix u::real assume u: "0 \<le> u" "u \<le> 1" + show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<in> T" + apply (rule convexD [OF \<open>convex T\<close>]) + using that u apply (auto simp add: ) + apply (metis add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 relS subsetD) + done + have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \<longleftrightarrow> + (1 - u + u * d) *\<^sub>R (x - a) = 0" for d + by (auto simp: algebra_simps) + have "x \<in> T" "x \<noteq> a" using that by auto + then have axa: "a + (x - a) \<in> affine hull S" + by (metis (no_types) add.commute affS diff_add_cancel set_rev_mp) + then have "\<not> dd (x - a) \<le> 0 \<and> a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S" + using \<open>x \<noteq> a\<close> dd1 by fastforce + with \<open>x \<noteq> a\<close> show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<noteq> 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) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))" + proof (simp add: retraction_def, intro conjI ballI) + show "rel_frontier S \<subseteq> T - {a}" + using arelS relS rel_frontier_def by fastforce + show "continuous_on (T - {a}) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))" + using contdd by (simp add: o_def) + show "(\<lambda>x. a + dd (x - a) *\<^sub>R (x - a)) ` (T - {a}) \<subseteq> 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 \<in> rel_frontier S" for x + proof - + have "x \<noteq> a" + using that arelS by (auto simp add: rel_frontier_def) + have False if "dd (x - a) < 1" + proof - + have "x \<in> closure S" + using x by (auto simp: rel_frontier_def) + then have segsub: "open_segment a x \<subseteq> rel_interior S" + by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS]) + have xaffS: "x \<in> 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) \<in> rel_frontier S" + using dd1 by (auto simp add: \<open>x \<noteq> a\<close>) + moreover have "a + dd (x - a) *\<^sub>R (x - a) \<in> open_segment a x" + using \<open>x \<noteq> a\<close> \<open>0 < dd (x - a)\<close> + 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 add: rel_frontier_def) + qed + moreover have False if "1 < dd (x - a)" + using x that dd2 [of "x - a" 1] \<open>x \<noteq> a\<close> closure_affine_hull + by (auto simp: rel_frontier_def) + ultimately have "dd (x - a) = 1" --\<open>similar to another proof above\<close> + 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 \<in> 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 add: 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 \<in> 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) + +subsection\<open>Borsuk-style characterization of separation\<close> + +lemma continuous_on_Borsuk_map: + "a \<notin> s \<Longrightarrow> continuous_on s (\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a))" +by (rule continuous_intros | force)+ + +lemma Borsuk_map_into_sphere: + "(\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a)) ` s \<subseteq> sphere 0 1 \<longleftrightarrow> (a \<notin> 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 (\<lambda>x. True) s (sphere 0 1) + (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) + (\<lambda>x. inverse(norm(x - b)) *\<^sub>R (x - b))" +proof - + obtain g where "path g" "path_image g \<subseteq> -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 = "\<lambda>z. inverse(norm(snd z - (g o fst)z)) *\<^sub>R (snd z - (g o 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 \<in> components(- s)" and boc: "bounded c" and "a \<in> c" + shows "~ (\<exists>g. continuous_on (s \<union> c) g \<and> + g ` (s \<union> c) \<subseteq> sphere 0 1 \<and> + (\<forall>x \<in> 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 \<subseteq> -s" + using assms by (simp add: in_components_subset) + with \<open>a \<in> c\<close> have "a \<notin> s" by blast + then have ceq: "c = connected_component_set (- s) a" + by (metis \<open>a \<in> c\<close> cin components_iff connected_component_eq) + then have "bounded (s \<union> connected_component_set (- s) a)" + using \<open>compact s\<close> boc compact_imp_bounded by auto + with bounded_subset_ballD obtain r where "0 < r" and r: "(s \<union> connected_component_set (- s) a) \<subseteq> ball a r" + by blast + { fix g + assume "continuous_on (s \<union> c) g" + "g ` (s \<union> c) \<subseteq> sphere 0 1" + and [simp]: "\<And>x. x \<in> s \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)" + then have [simp]: "\<And>x. x \<in> s \<union> c \<Longrightarrow> norm (g x) = 1" + by force + have cb_eq: "cball a r = (s \<union> connected_component_set (- s) a) \<union> + (cball a r - connected_component_set (- s) a)" + using ball_subset_cball [of a r] r by auto + have cont1: "continuous_on (s \<union> connected_component_set (- s) a) + (\<lambda>x. a + r *\<^sub>R g x)" + apply (rule continuous_intros)+ + using \<open>continuous_on (s \<union> c) g\<close> ceq by blast + have cont2: "continuous_on (cball a r - connected_component_set (- s) a) + (\<lambda>x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" + by (rule continuous_intros | force simp: \<open>a \<notin> s\<close>)+ + have 1: "continuous_on (cball a r) + (\<lambda>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: \<open>closed s\<close> open_Compl open_connected_component) + done + have 2: "(\<lambda>x. a + r *\<^sub>R g x) ` (cball a r \<inter> connected_component_set (- s) a) + \<subseteq> sphere a r " + using \<open>0 < r\<close> by (force simp: dist_norm ceq) + have "retraction (cball a r) (sphere a r) + (\<lambda>x. if x \<in> 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 \<open>0 < r\<close> + apply (simp add: retraction_def dist_norm 1 2, safe) + apply (force simp: dist_norm abs_if mult_less_0_iff divide_simps \<open>a \<notin> s\<close>) + using r + by (auto simp: dist_norm norm_minus_commute) + then have False + using no_retraction_cball + [OF \<open>0 < r\<close>, of a, unfolded retract_of_def, simplified, rule_format, + of "\<lambda>x. if x \<in> 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 + +subsection\<open>Absolute retracts, Etc.\<close> + +text\<open>Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also + Euclidean neighbourhood retracts (ENR). We define AR and ANR by + specializing the standard definitions for a set to embedding in +spaces of higher dimension. \<close> + +(*This turns out to be sufficient (since any set in +R^n can be embedded as a closed subset of a convex subset of R^{n+1}) to +derive the usual definitions, but we need to split them into two +implications because of the lack of type quantifiers. Then ENR turns out +to be equivalent to ANR plus local compactness. -- JRH*) + +definition AR :: "'a::topological_space set => bool" + where + "AR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S' + \<longrightarrow> S' retract_of U" + +definition ANR :: "'a::topological_space set => bool" + where + "ANR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S' + \<longrightarrow> (\<exists>T. openin (subtopology euclidean U) T \<and> + S' retract_of T)" + +definition ENR :: "'a::topological_space set => bool" + where "ENR S \<equiv> \<exists>U. open U \<and> S retract_of U" + +text\<open> First, show that we do indeed get the "usual" properties of ARs and ANRs.\<close> + +proposition AR_imp_absolute_extensor: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "AR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S" + and cloUT: "closedin (subtopology euclidean U) T" + obtains g where "continuous_on U g" "g ` U \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x" +proof - + have "aff_dim S < int (DIM('b \<times> real))" + using aff_dim_le_DIM [of S] by simp + then obtain C and S' :: "('b * real) set" + where C: "convex C" "C \<noteq> {}" + and cloCS: "closedin (subtopology euclidean C) S'" + and hom: "S homeomorphic S'" + by (metis that homeomorphic_closedin_convex) + then have "S' retract_of C" + using \<open>AR S\<close> by (simp add: AR_def) + then obtain r where "S' \<subseteq> C" and contr: "continuous_on C r" + and "r ` C \<subseteq> S'" and rid: "\<And>x. x\<in>S' \<Longrightarrow> r x = x" + by (auto simp: retraction_def retract_of_def) + obtain g h where "homeomorphism S S' g h" + using hom by (force simp: homeomorphic_def) + then have "continuous_on (f ` T) g" + by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def) + then have contgf: "continuous_on T (g o f)" + by (metis continuous_on_compose contf) + have gfTC: "(g \<circ> f) ` T \<subseteq> C" + proof - + have "g ` S = S'" + by (metis (no_types) \<open>homeomorphism S S' g h\<close> homeomorphism_def) + with \<open>S' \<subseteq> C\<close> \<open>f ` T \<subseteq> S\<close> show ?thesis by force + qed + obtain f' where f': "continuous_on U f'" "f' ` U \<subseteq> C" + "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x" + by (metis Dugundji [OF C cloUT contgf gfTC]) + show ?thesis + proof (rule_tac g = "h o r o f'" in that) + show "continuous_on U (h \<circ> r \<circ> f')" + apply (intro continuous_on_compose f') + using continuous_on_subset contr f' apply blast + by (meson \<open>homeomorphism S S' g h\<close> \<open>r ` C \<subseteq> S'\<close> continuous_on_subset \<open>f' ` U \<subseteq> C\<close> homeomorphism_def image_mono) + show "(h \<circ> r \<circ> f') ` U \<subseteq> S" + using \<open>homeomorphism S S' g h\<close> \<open>r ` C \<subseteq> S'\<close> \<open>f' ` U \<subseteq> C\<close> + by (fastforce simp: homeomorphism_def) + show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x" + using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> f' + by (auto simp: rid homeomorphism_def) + qed +qed + +lemma AR_imp_absolute_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "AR S" "S homeomorphic S'" + and clo: "closedin (subtopology euclidean U) S'" + shows "S' retract_of U" +proof - + obtain g h where hom: "homeomorphism S S' g h" + using assms by (force simp: homeomorphic_def) + have h: "continuous_on S' h" " h ` S' \<subseteq> S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S" + and h'h: "\<And>x. x \<in> S' \<Longrightarrow> h' x = h x" + by (blast intro: AR_imp_absolute_extensor [OF \<open>AR S\<close> h clo]) + have [simp]: "S' \<subseteq> U" using clo closedin_limpt by blast + show ?thesis + proof (simp add: retraction_def retract_of_def, intro exI conjI) + show "continuous_on U (g o h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \<circ> h') ` U \<subseteq> S'" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\<forall>x\<in>S'. (g \<circ> h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed +qed + +lemma AR_imp_absolute_retract_UNIV: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "AR S" and hom: "S homeomorphic S'" + and clo: "closed S'" + shows "S' retract_of UNIV" +apply (rule AR_imp_absolute_retract [OF \<open>AR S\<close> hom]) +using clo closed_closedin by auto + +lemma absolute_extensor_imp_AR: + fixes S :: "'a::euclidean_space set" + assumes "\<And>f :: 'a * real \<Rightarrow> 'a. + \<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S; + closedin (subtopology euclidean U) T\<rbrakk> + \<Longrightarrow> \<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)" + shows "AR S" +proof (clarsimp simp: AR_def) + fix U and T :: "('a * real) set" + assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" + then obtain g h where hom: "homeomorphism S T g h" + by (force simp: homeomorphic_def) + have h: "continuous_on T h" " h ` T \<subseteq> S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S" + and h'h: "\<forall>x\<in>T. h' x = h x" + using assms [OF h clo] by blast + have [simp]: "T \<subseteq> U" + using clo closedin_imp_subset by auto + show "T retract_of U" + proof (simp add: retraction_def retract_of_def, intro exI conjI) + show "continuous_on U (g o h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \<circ> h') ` U \<subseteq> T" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\<forall>x\<in>T. (g \<circ> h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed +qed + +lemma AR_eq_absolute_extensor: + fixes S :: "'a::euclidean_space set" + shows "AR S \<longleftrightarrow> + (\<forall>f :: 'a * real \<Rightarrow> 'a. + \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow> + closedin (subtopology euclidean U) T \<longrightarrow> + (\<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))" +apply (rule iffI) + apply (metis AR_imp_absolute_extensor) +apply (simp add: absolute_extensor_imp_AR) +done + +lemma AR_imp_retract: + fixes S :: "'a::euclidean_space set" + assumes "AR S \<and> closedin (subtopology euclidean U) S" + shows "S retract_of U" +using AR_imp_absolute_retract assms homeomorphic_refl by blast + +lemma AR_homeomorphic_AR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "AR T" "S homeomorphic T" + shows "AR S" +unfolding AR_def +by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym) + +lemma homeomorphic_AR_iff_AR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + shows "S homeomorphic T \<Longrightarrow> AR S \<longleftrightarrow> AR T" +by (metis AR_homeomorphic_AR homeomorphic_sym) + + +proposition ANR_imp_absolute_neighbourhood_extensor: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "ANR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S" + and cloUT: "closedin (subtopology euclidean U) T" + obtains V g where "T \<subseteq> V" "openin (subtopology euclidean U) V" + "continuous_on V g" + "g ` V \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x" +proof - + have "aff_dim S < int (DIM('b \<times> real))" + using aff_dim_le_DIM [of S] by simp + then obtain C and S' :: "('b * real) set" + where C: "convex C" "C \<noteq> {}" + and cloCS: "closedin (subtopology euclidean C) S'" + and hom: "S homeomorphic S'" + by (metis that homeomorphic_closedin_convex) + then obtain D where opD: "openin (subtopology euclidean C) D" and "S' retract_of D" + using \<open>ANR S\<close> by (auto simp: ANR_def) + then obtain r where "S' \<subseteq> D" and contr: "continuous_on D r" + and "r ` D \<subseteq> S'" and rid: "\<And>x. x \<in> S' \<Longrightarrow> r x = x" + by (auto simp: retraction_def retract_of_def) + obtain g h where homgh: "homeomorphism S S' g h" + using hom by (force simp: homeomorphic_def) + have "continuous_on (f ` T) g" + by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def homgh) + then have contgf: "continuous_on T (g o f)" + by (intro continuous_on_compose contf) + have gfTC: "(g \<circ> f) ` T \<subseteq> C" + proof - + have "g ` S = S'" + by (metis (no_types) homeomorphism_def homgh) + then show ?thesis + by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology) + qed + obtain f' where contf': "continuous_on U f'" + and "f' ` U \<subseteq> C" + and eq: "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x" + by (metis Dugundji [OF C cloUT contgf gfTC]) + show ?thesis + proof (rule_tac V = "{x \<in> U. f' x \<in> D}" and g = "h o r o f'" in that) + show "T \<subseteq> {x \<in> U. f' x \<in> D}" + using cloUT closedin_imp_subset \<open>S' \<subseteq> D\<close> \<open>f ` T \<subseteq> S\<close> eq homeomorphism_image1 homgh + by fastforce + show ope: "openin (subtopology euclidean U) {x \<in> U. f' x \<in> D}" + using \<open>f' ` U \<subseteq> C\<close> by (auto simp: opD contf' continuous_openin_preimage) + have conth: "continuous_on (r ` f' ` {x \<in> U. f' x \<in> D}) h" + apply (rule continuous_on_subset [of S']) + using homeomorphism_def homgh apply blast + using \<open>r ` D \<subseteq> S'\<close> by blast + show "continuous_on {x \<in> U. f' x \<in> D} (h \<circ> r \<circ> f')" + apply (intro continuous_on_compose conth + continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto) + done + show "(h \<circ> r \<circ> f') ` {x \<in> U. f' x \<in> D} \<subseteq> S" + using \<open>homeomorphism S S' g h\<close> \<open>f' ` U \<subseteq> C\<close> \<open>r ` D \<subseteq> S'\<close> + by (auto simp: homeomorphism_def) + show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x" + using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> eq + by (auto simp: rid homeomorphism_def) + qed +qed + + +corollary ANR_imp_absolute_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ANR S" "S homeomorphic S'" + and clo: "closedin (subtopology euclidean U) S'" + obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" +proof - + obtain g h where hom: "homeomorphism S S' g h" + using assms by (force simp: homeomorphic_def) + have h: "continuous_on S' h" " h ` S' \<subseteq> S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + from ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo] + obtain V h' where "S' \<subseteq> V" and opUV: "openin (subtopology euclidean U) V" + and h': "continuous_on V h'" "h' ` V \<subseteq> S" + and h'h:"\<And>x. x \<in> S' \<Longrightarrow> h' x = h x" + by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo]) + have "S' retract_of V" + proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>S' \<subseteq> V\<close>) + show "continuous_on V (g o h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \<circ> h') ` V \<subseteq> S'" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\<forall>x\<in>S'. (g \<circ> h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed + then show ?thesis + by (rule that [OF opUV]) +qed + +corollary ANR_imp_absolute_neighbourhood_retract_UNIV: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'" + obtains V where "open V" "S' retract_of V" + using ANR_imp_absolute_neighbourhood_retract [OF \<open>ANR S\<close> hom] +by (metis clo closed_closedin open_openin subtopology_UNIV) + +lemma absolute_neighbourhood_extensor_imp_ANR: + fixes S :: "'a::euclidean_space set" + assumes "\<And>f :: 'a * real \<Rightarrow> 'a. + \<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S; + closedin (subtopology euclidean U) T\<rbrakk> + \<Longrightarrow> \<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> + continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)" + shows "ANR S" +proof (clarsimp simp: ANR_def) + fix U and T :: "('a * real) set" + assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" + then obtain g h where hom: "homeomorphism S T g h" + by (force simp: homeomorphic_def) + have h: "continuous_on T h" " h ` T \<subseteq> S" + using hom homeomorphism_def apply blast + apply (metis hom equalityE homeomorphism_def) + done + obtain V h' where "T \<subseteq> V" and opV: "openin (subtopology euclidean U) V" + and h': "continuous_on V h'" "h' ` V \<subseteq> S" + and h'h: "\<forall>x\<in>T. h' x = h x" + using assms [OF h clo] by blast + have [simp]: "T \<subseteq> U" + using clo closedin_imp_subset by auto + have "T retract_of V" + proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>T \<subseteq> V\<close>) + show "continuous_on V (g o h')" + apply (intro continuous_on_compose h') + apply (meson hom continuous_on_subset h' homeomorphism_cont1) + done + show "(g \<circ> h') ` V \<subseteq> T" + using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) + show "\<forall>x\<in>T. (g \<circ> h') x = x" + by clarsimp (metis h'h hom homeomorphism_def) + qed + then show "\<exists>V. openin (subtopology euclidean U) V \<and> T retract_of V" + using opV by blast +qed + +lemma ANR_eq_absolute_neighbourhood_extensor: + fixes S :: "'a::euclidean_space set" + shows "ANR S \<longleftrightarrow> + (\<forall>f :: 'a * real \<Rightarrow> 'a. + \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow> + closedin (subtopology euclidean U) T \<longrightarrow> + (\<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> + continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))" +apply (rule iffI) + apply (metis ANR_imp_absolute_neighbourhood_extensor) +apply (simp add: absolute_neighbourhood_extensor_imp_ANR) +done + +lemma ANR_imp_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" "closedin (subtopology euclidean U) S" + obtains V where "openin (subtopology euclidean U) V" "S retract_of V" +using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast + +lemma ANR_imp_absolute_closed_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ANR S" "S homeomorphic S'" and US': "closedin (subtopology euclidean U) S'" + obtains V W + where "openin (subtopology euclidean U) V" + "closedin (subtopology euclidean U) W" + "S' \<subseteq> V" "V \<subseteq> W" "S' retract_of W" +proof - + obtain Z where "openin (subtopology euclidean U) Z" and S'Z: "S' retract_of Z" + by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) + then have UUZ: "closedin (subtopology euclidean U) (U - Z)" + by auto + have "S' \<inter> (U - Z) = {}" + using \<open>S' retract_of Z\<close> closedin_retract closedin_subtopology by fastforce + then obtain V W + where "openin (subtopology euclidean U) V" + and "openin (subtopology euclidean U) W" + and "S' \<subseteq> V" "U - Z \<subseteq> W" "V \<inter> W = {}" + using separation_normal_local [OF US' UUZ] by auto + moreover have "S' retract_of U - W" + apply (rule retract_of_subset [OF S'Z]) + using US' \<open>S' \<subseteq> V\<close> \<open>V \<inter> W = {}\<close> closedin_subset apply fastforce + using Diff_subset_conv \<open>U - Z \<subseteq> W\<close> by blast + ultimately show ?thesis + apply (rule_tac V=V and W = "U-W" in that) + using openin_imp_subset apply (force simp:)+ + done +qed + +lemma ANR_imp_closed_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" "closedin (subtopology euclidean U) S" + obtains V W where "openin (subtopology euclidean U) V" + "closedin (subtopology euclidean U) W" + "S \<subseteq> V" "V \<subseteq> W" "S retract_of W" +by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) + +lemma ANR_homeomorphic_ANR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "ANR T" "S homeomorphic T" + shows "ANR S" +unfolding ANR_def +by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym) + +lemma homeomorphic_ANR_iff_ANR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + shows "S homeomorphic T \<Longrightarrow> ANR S \<longleftrightarrow> ANR T" +by (metis ANR_homeomorphic_ANR homeomorphic_sym) + +subsection\<open> Analogous properties of ENRs.\<close> + +proposition ENR_imp_absolute_neighbourhood_retract: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ENR S" and hom: "S homeomorphic S'" + and "S' \<subseteq> U" + obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" +proof - + obtain X where "open X" "S retract_of X" + using \<open>ENR S\<close> by (auto simp: ENR_def) + then obtain r where "retraction X S r" + by (auto simp: retract_of_def) + have "locally compact S'" + using retract_of_locally_compact open_imp_locally_compact + homeomorphic_local_compactness \<open>S retract_of X\<close> \<open>open X\<close> hom by blast + then obtain W where UW: "openin (subtopology euclidean U) W" + and WS': "closedin (subtopology euclidean W) S'" + apply (rule locally_compact_closedin_open) + apply (rename_tac W) + apply (rule_tac W = "U \<inter> W" in that, blast) + by (simp add: \<open>S' \<subseteq> U\<close> closedin_limpt) + obtain f g where hom: "homeomorphism S S' f g" + using assms by (force simp: homeomorphic_def) + have contg: "continuous_on S' g" + using hom homeomorphism_def by blast + moreover have "g ` S' \<subseteq> S" by (metis hom equalityE homeomorphism_def) + ultimately obtain h where conth: "continuous_on W h" and hg: "\<And>x. x \<in> S' \<Longrightarrow> h x = g x" + using Tietze_unbounded [of S' g W] WS' by blast + have "W \<subseteq> U" using UW openin_open by auto + have "S' \<subseteq> W" using WS' closedin_closed by auto + have him: "\<And>x. x \<in> S' \<Longrightarrow> h x \<in> X" + by (metis (no_types) \<open>S retract_of X\<close> hg hom homeomorphism_def image_insert insert_absorb insert_iff retract_of_imp_subset subset_eq) + have "S' retract_of {x \<in> W. h x \<in> X}" + proof (simp add: retraction_def retract_of_def, intro exI conjI) + show "S' \<subseteq> {x \<in> W. h x \<in> X}" + using him WS' closedin_imp_subset by blast + show "continuous_on {x \<in> W. h x \<in> X} (f o r o h)" + proof (intro continuous_on_compose) + show "continuous_on {x \<in> W. h x \<in> X} h" + by (metis (no_types) Collect_restrict conth continuous_on_subset) + show "continuous_on (h ` {x \<in> W. h x \<in> X}) r" + proof - + have "h ` {b \<in> W. h b \<in> X} \<subseteq> X" + by blast + then show "continuous_on (h ` {b \<in> W. h b \<in> X}) r" + by (meson \<open>retraction X S r\<close> continuous_on_subset retraction) + qed + show "continuous_on (r ` h ` {x \<in> W. h x \<in> X}) f" + apply (rule continuous_on_subset [of S]) + using hom homeomorphism_def apply blast + apply clarify + apply (meson \<open>retraction X S r\<close> subsetD imageI retraction_def) + done + qed + show "(f \<circ> r \<circ> h) ` {x \<in> W. h x \<in> X} \<subseteq> S'" + using \<open>retraction X S r\<close> hom + by (auto simp: retraction_def homeomorphism_def) + show "\<forall>x\<in>S'. (f \<circ> r \<circ> h) x = x" + using \<open>retraction X S r\<close> hom by (auto simp: retraction_def homeomorphism_def hg) + qed + then show ?thesis + apply (rule_tac V = "{x. x \<in> W \<and> h x \<in> X}" in that) + apply (rule openin_trans [OF _ UW]) + using \<open>continuous_on W h\<close> \<open>open X\<close> continuous_openin_preimage_eq apply blast+ + done +qed + +corollary ENR_imp_absolute_neighbourhood_retract_UNIV: + fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" + assumes "ENR S" "S homeomorphic S'" + obtains T' where "open T'" "S' retract_of T'" +by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV) + +lemma ENR_homeomorphic_ENR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "ENR T" "S homeomorphic T" + shows "ENR S" +unfolding ENR_def +by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym) + +lemma homeomorphic_ENR_iff_ENR: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "S homeomorphic T" + shows "ENR S \<longleftrightarrow> ENR T" +by (meson ENR_homeomorphic_ENR assms homeomorphic_sym) + +lemma ENR_translation: + fixes S :: "'a::euclidean_space set" + shows "ENR(image (\<lambda>x. a + x) S) \<longleftrightarrow> ENR S" +by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR) + +lemma ENR_linear_image_eq: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "linear f" "inj f" + shows "ENR (image f S) \<longleftrightarrow> ENR S" +apply (rule homeomorphic_ENR_iff_ENR) +using assms homeomorphic_sym linear_homeomorphic_image by auto + +subsection\<open>Some relations among the concepts\<close> + +text\<open>We also relate AR to being a retract of UNIV, which is often a more convenient proxy in the closed case.\<close> + +lemma AR_imp_ANR: "AR S \<Longrightarrow> ANR S" + using ANR_def AR_def by fastforce + +lemma ENR_imp_ANR: + fixes S :: "'a::euclidean_space set" + shows "ENR S \<Longrightarrow> ANR S" +apply (simp add: ANR_def) +by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset) + +lemma ENR_ANR: + fixes S :: "'a::euclidean_space set" + shows "ENR S \<longleftrightarrow> ANR S \<and> locally compact S" +proof + assume "ENR S" + then have "locally compact S" + using ENR_def open_imp_locally_compact retract_of_locally_compact by auto + then show "ANR S \<and> locally compact S" + using ENR_imp_ANR \<open>ENR S\<close> by blast +next + assume "ANR S \<and> locally compact S" + then have "ANR S" "locally compact S" by auto + then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T" + using locally_compact_homeomorphic_closed + by (metis DIM_prod DIM_real Suc_eq_plus1 lessI) + then show "ENR S" + using \<open>ANR S\<close> + apply (simp add: ANR_def) + apply (drule_tac x=UNIV in spec) + apply (drule_tac x=T in spec) + apply (auto simp: closed_closedin) + apply (meson ENR_def ENR_homeomorphic_ENR open_openin) + done +qed + + +proposition AR_ANR: + fixes S :: "'a::euclidean_space set" + shows "AR S \<longleftrightarrow> ANR S \<and> contractible S \<and> S \<noteq> {}" + (is "?lhs = ?rhs") +proof + assume ?lhs + obtain C and S' :: "('a * real) set" + where "convex C" "C \<noteq> {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'" + apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) + using aff_dim_le_DIM [of S] by auto + with \<open>AR S\<close> have "contractible S" + apply (simp add: AR_def) + apply (drule_tac x=C in spec) + apply (drule_tac x="S'" in spec, simp) + using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce + with \<open>AR S\<close> show ?rhs + apply (auto simp: AR_imp_ANR) + apply (force simp: AR_def) + done +next + assume ?rhs + then obtain a and h:: "real \<times> 'a \<Rightarrow> 'a" + where conth: "continuous_on ({0..1} \<times> S) h" + and hS: "h ` ({0..1} \<times> S) \<subseteq> S" + and [simp]: "\<And>x. h(0, x) = x" + and [simp]: "\<And>x. h(1, x) = a" + and "ANR S" "S \<noteq> {}" + by (auto simp: contractible_def homotopic_with_def) + then have "a \<in> S" + by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) + have "\<exists>g. continuous_on W g \<and> g ` W \<subseteq> S \<and> (\<forall>x\<in>T. g x = f x)" + if f: "continuous_on T f" "f ` T \<subseteq> S" + and WT: "closedin (subtopology euclidean W) T" + for W T and f :: "'a \<times> real \<Rightarrow> 'a" + proof - + obtain U g + where "T \<subseteq> U" and WU: "openin (subtopology euclidean W) U" + and contg: "continuous_on U g" + and "g ` U \<subseteq> S" and gf: "\<And>x. x \<in> T \<Longrightarrow> g x = f x" + using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \<open>ANR S\<close>, rule_format, OF f WT] + by auto + have WWU: "closedin (subtopology euclidean W) (W - U)" + using WU closedin_diff by fastforce + moreover have "(W - U) \<inter> T = {}" + using \<open>T \<subseteq> U\<close> by auto + ultimately obtain V V' + where WV': "openin (subtopology euclidean W) V'" + and WV: "openin (subtopology euclidean W) V" + and "W - U \<subseteq> V'" "T \<subseteq> V" "V' \<inter> V = {}" + using separation_normal_local [of W "W-U" T] WT by blast + then have WVT: "T \<inter> (W - V) = {}" + by auto + have WWV: "closedin (subtopology euclidean W) (W - V)" + using WV closedin_diff by fastforce + obtain j :: " 'a \<times> real \<Rightarrow> real" + where contj: "continuous_on W j" + and j: "\<And>x. x \<in> W \<Longrightarrow> j x \<in> {0..1}" + and j0: "\<And>x. x \<in> W - V \<Longrightarrow> j x = 1" + and j1: "\<And>x. x \<in> T \<Longrightarrow> j x = 0" + by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment) + have Weq: "W = (W - V) \<union> (W - V')" + using \<open>V' \<inter> V = {}\<close> by force + show ?thesis + proof (intro conjI exI) + have *: "continuous_on (W - V') (\<lambda>x. h (j x, g x))" + apply (rule continuous_on_compose2 [OF conth continuous_on_Pair]) + apply (rule continuous_on_subset [OF contj Diff_subset]) + apply (rule continuous_on_subset [OF contg]) + apply (metis Diff_subset_conv Un_commute \<open>W - U \<subseteq> V'\<close>) + using j \<open>g ` U \<subseteq> S\<close> \<open>W - U \<subseteq> V'\<close> apply fastforce + done + show "continuous_on W (\<lambda>x. if x \<in> W - V then a else h (j x, g x))" + apply (subst Weq) + apply (rule continuous_on_cases_local) + apply (simp_all add: Weq [symmetric] WWV continuous_on_const *) + using WV' closedin_diff apply fastforce + apply (auto simp: j0 j1) + done + next + have "h (j (x, y), g (x, y)) \<in> S" if "(x, y) \<in> W" "(x, y) \<in> V" for x y + proof - + have "j(x, y) \<in> {0..1}" + using j that by blast + moreover have "g(x, y) \<in> S" + using \<open>V' \<inter> V = {}\<close> \<open>W - U \<subseteq> V'\<close> \<open>g ` U \<subseteq> S\<close> that by fastforce + ultimately show ?thesis + using hS by blast + qed + with \<open>a \<in> S\<close> \<open>g ` U \<subseteq> S\<close> + show "(\<lambda>x. if x \<in> W - V then a else h (j x, g x)) ` W \<subseteq> S" + by auto + next + show "\<forall>x\<in>T. (if x \<in> W - V then a else h (j x, g x)) = f x" + using \<open>T \<subseteq> V\<close> by (auto simp: j0 j1 gf) + qed + qed + then show ?lhs + by (simp add: AR_eq_absolute_extensor) +qed + + +lemma ANR_retract_of_ANR: + fixes S :: "'a::euclidean_space set" + assumes "ANR T" "S retract_of T" + shows "ANR S" +using assms +apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def) +apply (clarsimp elim!: all_forward) +apply (erule impCE, metis subset_trans) +apply (clarsimp elim!: ex_forward) +apply (rule_tac x="r o g" in exI) +by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans) + +lemma AR_retract_of_AR: + fixes S :: "'a::euclidean_space set" + shows "\<lbrakk>AR T; S retract_of T\<rbrakk> \<Longrightarrow> AR S" +using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce + +lemma ENR_retract_of_ENR: + "\<lbrakk>ENR T; S retract_of T\<rbrakk> \<Longrightarrow> ENR S" +by (meson ENR_def retract_of_trans) + +lemma retract_of_UNIV: + fixes S :: "'a::euclidean_space set" + shows "S retract_of UNIV \<longleftrightarrow> AR S \<and> closed S" +by (metis AR_ANR AR_imp_retract ENR_def ENR_imp_ANR closed_UNIV closed_closedin contractible_UNIV empty_not_UNIV open_UNIV retract_of_closed retract_of_contractible retract_of_empty(1) subtopology_UNIV) + +lemma compact_AR [simp]: + fixes S :: "'a::euclidean_space set" + shows "compact S \<and> AR S \<longleftrightarrow> compact S \<and> S retract_of UNIV" +using compact_imp_closed retract_of_UNIV by blast + +subsection\<open>More properties of ARs, ANRs and ENRs\<close> + +lemma not_AR_empty [simp]: "~ AR({})" + by (auto simp: AR_def) + +lemma ENR_empty [simp]: "ENR {}" + by (simp add: ENR_def) + +lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)" + by (simp add: ENR_imp_ANR) + +lemma convex_imp_AR: + fixes S :: "'a::euclidean_space set" + shows "\<lbrakk>convex S; S \<noteq> {}\<rbrakk> \<Longrightarrow> AR S" +apply (rule absolute_extensor_imp_AR) +apply (rule Dugundji, assumption+) +by blast + +lemma convex_imp_ANR: + fixes S :: "'a::euclidean_space set" + shows "convex S \<Longrightarrow> ANR S" +using ANR_empty AR_imp_ANR convex_imp_AR by blast + +lemma ENR_convex_closed: + fixes S :: "'a::euclidean_space set" + shows "\<lbrakk>closed S; convex S\<rbrakk> \<Longrightarrow> ENR S" +using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast + +lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)" + using retract_of_UNIV by auto + +lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)" + by (simp add: AR_imp_ANR) + +lemma ENR_UNIV [simp]:"ENR UNIV" + using ENR_def by blast + +lemma AR_singleton: + fixes a :: "'a::euclidean_space" + shows "AR {a}" + using retract_of_UNIV by blast + +lemma ANR_singleton: + fixes a :: "'a::euclidean_space" + shows "ANR {a}" + by (simp add: AR_imp_ANR AR_singleton) + +lemma ENR_singleton: "ENR {a}" + using ENR_def by blast + +subsection\<open>ARs closed under union\<close> + +lemma AR_closed_Un_local_aux: + fixes U :: "'a::euclidean_space set" + assumes "closedin (subtopology euclidean U) S" + "closedin (subtopology euclidean U) T" + "AR S" "AR T" "AR(S \<inter> T)" + shows "(S \<union> T) retract_of U" +proof - + have "S \<inter> T \<noteq> {}" + using assms AR_def by fastforce + have "S \<subseteq> U" "T \<subseteq> U" + using assms by (auto simp: closedin_imp_subset) + define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}" + define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}" + define W where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}" + have US': "closedin (subtopology euclidean U) S'" + using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"] + by (simp add: S'_def continuous_intros) + have UT': "closedin (subtopology euclidean U) T'" + using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"] + by (simp add: T'_def continuous_intros) + have "S \<subseteq> S'" + using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce + have "T \<subseteq> T'" + using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce + have "S \<inter> T \<subseteq> W" "W \<subseteq> U" + using \<open>S \<subseteq> U\<close> by (auto simp: W_def setdist_sing_in_set) + have "(S \<inter> T) retract_of W" + apply (rule AR_imp_absolute_retract [OF \<open>AR(S \<inter> T)\<close>]) + apply (simp add: homeomorphic_refl) + apply (rule closedin_subset_trans [of U]) + apply (simp_all add: assms closedin_Int \<open>S \<inter> T \<subseteq> W\<close> \<open>W \<subseteq> U\<close>) + done + then obtain r0 + where "S \<inter> T \<subseteq> W" and contr0: "continuous_on W r0" + and "r0 ` W \<subseteq> S \<inter> T" + and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x" + by (auto simp: retract_of_def retraction_def) + have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x + using setdist_eq_0_closedin \<open>S \<inter> T \<noteq> {}\<close> assms + by (force simp: W_def setdist_sing_in_set) + have "S' \<inter> T' = W" + by (auto simp: S'_def T'_def W_def) + then have cloUW: "closedin (subtopology euclidean U) W" + using closedin_Int US' UT' by blast + define r where "r \<equiv> \<lambda>x. if x \<in> W then r0 x else x" + have "r ` (W \<union> S) \<subseteq> S" "r ` (W \<union> T) \<subseteq> T" + using \<open>r0 ` W \<subseteq> S \<inter> T\<close> r_def by auto + have contr: "continuous_on (W \<union> (S \<union> T)) r" + unfolding r_def + proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) + show "closedin (subtopology euclidean (W \<union> (S \<union> T))) W" + using \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> \<open>closedin (subtopology euclidean U) W\<close> closedin_subset_trans by fastforce + show "closedin (subtopology euclidean (W \<union> (S \<union> T))) (S \<union> T)" + by (meson \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) + show "\<And>x. x \<in> W \<and> x \<notin> W \<or> x \<in> S \<union> T \<and> x \<in> W \<Longrightarrow> r0 x = x" + by (auto simp: ST) + qed + have cloUWS: "closedin (subtopology euclidean U) (W \<union> S)" + by (simp add: cloUW assms closedin_Un) + obtain g where contg: "continuous_on U g" + and "g ` U \<subseteq> S" and geqr: "\<And>x. x \<in> W \<union> S \<Longrightarrow> g x = r x" + apply (rule AR_imp_absolute_extensor [OF \<open>AR S\<close> _ _ cloUWS]) + apply (rule continuous_on_subset [OF contr]) + using \<open>r ` (W \<union> S) \<subseteq> S\<close> apply auto + done + have cloUWT: "closedin (subtopology euclidean U) (W \<union> T)" + by (simp add: cloUW assms closedin_Un) + obtain h where conth: "continuous_on U h" + and "h ` U \<subseteq> T" and heqr: "\<And>x. x \<in> W \<union> T \<Longrightarrow> h x = r x" + apply (rule AR_imp_absolute_extensor [OF \<open>AR T\<close> _ _ cloUWT]) + apply (rule continuous_on_subset [OF contr]) + using \<open>r ` (W \<union> T) \<subseteq> T\<close> apply auto + done + have "U = S' \<union> T'" + by (force simp: S'_def T'_def) + then have cont: "continuous_on U (\<lambda>x. if x \<in> S' then g x else h x)" + apply (rule ssubst) + apply (rule continuous_on_cases_local) + using US' UT' \<open>S' \<inter> T' = W\<close> \<open>U = S' \<union> T'\<close> + contg conth continuous_on_subset geqr heqr apply auto + done + have UST: "(\<lambda>x. if x \<in> S' then g x else h x) ` U \<subseteq> S \<union> T" + using \<open>g ` U \<subseteq> S\<close> \<open>h ` U \<subseteq> T\<close> by auto + show ?thesis + apply (simp add: retract_of_def retraction_def \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close>) + apply (rule_tac x="\<lambda>x. if x \<in> S' then g x else h x" in exI) + apply (intro conjI cont UST) + by (metis IntI ST Un_iff \<open>S \<subseteq> S'\<close> \<open>S' \<inter> T' = W\<close> \<open>T \<subseteq> T'\<close> subsetD geqr heqr r0 r_def) +qed + + +proposition AR_closed_Un_local: + fixes S :: "'a::euclidean_space set" + assumes STS: "closedin (subtopology euclidean (S \<union> T)) S" + and STT: "closedin (subtopology euclidean (S \<union> T)) T" + and "AR S" "AR T" "AR(S \<inter> T)" + shows "AR(S \<union> T)" +proof - + have "C retract_of U" + if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" + for U and C :: "('a * real) set" + proof - + obtain f g where hom: "homeomorphism (S \<union> T) C f g" + using hom by (force simp: homeomorphic_def) + have US: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> S}" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) + using hom homeomorphism_def apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have UT: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> T}" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) + using hom homeomorphism_def apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have ARS: "AR {x \<in> C. g x \<in> S}" + apply (rule AR_homeomorphic_AR [OF \<open>AR S\<close>]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ART: "AR {x \<in> C. g x \<in> T}" + apply (rule AR_homeomorphic_AR [OF \<open>AR T\<close>]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ARI: "AR ({x \<in> C. g x \<in> S} \<inter> {x \<in> C. g x \<in> T})" + apply (rule AR_homeomorphic_AR [OF \<open>AR (S \<inter> T)\<close>]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom + apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have "C = {x \<in> C. g x \<in> S} \<union> {x \<in> C. g x \<in> T}" + using hom by (auto simp: homeomorphism_def) + then show ?thesis + by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI]) + qed + then show ?thesis + by (force simp: AR_def) +qed + +corollary AR_closed_Un: + fixes S :: "'a::euclidean_space set" + shows "\<lbrakk>closed S; closed T; AR S; AR T; AR (S \<inter> T)\<rbrakk> \<Longrightarrow> AR (S \<union> T)" +by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV) + +subsection\<open>ANRs closed under union\<close> + +lemma ANR_closed_Un_local_aux: + fixes U :: "'a::euclidean_space set" + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "ANR S" "ANR T" "ANR(S \<inter> T)" + obtains V where "openin (subtopology euclidean U) V" "(S \<union> T) retract_of V" +proof (cases "S = {} \<or> T = {}") + case True with assms that show ?thesis + by (auto simp: intro: ANR_imp_neighbourhood_retract) +next + case False + then have [simp]: "S \<noteq> {}" "T \<noteq> {}" by auto + have "S \<subseteq> U" "T \<subseteq> U" + using assms by (auto simp: closedin_imp_subset) + define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}" + define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}" + define W where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}" + have cloUS': "closedin (subtopology euclidean U) S'" + using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"] + by (simp add: S'_def continuous_intros) + have cloUT': "closedin (subtopology euclidean U) T'" + using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"] + by (simp add: T'_def continuous_intros) + have "S \<subseteq> S'" + using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce + have "T \<subseteq> T'" + using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce + have "S' \<union> T' = U" + by (auto simp: S'_def T'_def) + have "W \<subseteq> S'" + by (simp add: Collect_mono S'_def W_def) + have "W \<subseteq> T'" + by (simp add: Collect_mono T'_def W_def) + have ST_W: "S \<inter> T \<subseteq> W" and "W \<subseteq> U" + using \<open>S \<subseteq> U\<close> by (force simp: W_def setdist_sing_in_set)+ + have "S' \<inter> T' = W" + by (auto simp: S'_def T'_def W_def) + then have cloUW: "closedin (subtopology euclidean U) W" + using closedin_Int cloUS' cloUT' by blast + obtain W' W0 where "openin (subtopology euclidean W) W'" + and cloWW0: "closedin (subtopology euclidean W) W0" + and "S \<inter> T \<subseteq> W'" "W' \<subseteq> W0" + and ret: "(S \<inter> T) retract_of W0" + apply (rule ANR_imp_closed_neighbourhood_retract [OF \<open>ANR(S \<inter> T)\<close>]) + apply (rule closedin_subset_trans [of U, OF _ ST_W \<open>W \<subseteq> U\<close>]) + apply (blast intro: assms)+ + done + then obtain U0 where opeUU0: "openin (subtopology euclidean U) U0" + and U0: "S \<inter> T \<subseteq> U0" "U0 \<inter> W \<subseteq> W0" + unfolding openin_open using \<open>W \<subseteq> U\<close> by blast + have "W0 \<subseteq> U" + using \<open>W \<subseteq> U\<close> cloWW0 closedin_subset by fastforce + obtain r0 + where "S \<inter> T \<subseteq> W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \<subseteq> S \<inter> T" + and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x" + using ret by (force simp add: retract_of_def retraction_def) + have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x + using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin) + define r where "r \<equiv> \<lambda>x. if x \<in> W0 then r0 x else x" + have "r ` (W0 \<union> S) \<subseteq> S" "r ` (W0 \<union> T) \<subseteq> T" + using \<open>r0 ` W0 \<subseteq> S \<inter> T\<close> r_def by auto + have contr: "continuous_on (W0 \<union> (S \<union> T)) r" + unfolding r_def + proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) + show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) W0" + apply (rule closedin_subset_trans [of U]) + using cloWW0 cloUW closedin_trans \<open>W0 \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> apply blast+ + done + show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) (S \<union> T)" + by (meson \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W0 \<subseteq> U\<close> assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) + show "\<And>x. x \<in> W0 \<and> x \<notin> W0 \<or> x \<in> S \<union> T \<and> x \<in> W0 \<Longrightarrow> r0 x = x" + using ST cloWW0 closedin_subset by fastforce + qed + have cloS'WS: "closedin (subtopology euclidean S') (W0 \<union> S)" + by (meson closedin_subset_trans US cloUS' \<open>S \<subseteq> S'\<close> \<open>W \<subseteq> S'\<close> cloUW cloWW0 + closedin_Un closedin_imp_subset closedin_trans) + obtain W1 g where "W0 \<union> S \<subseteq> W1" and contg: "continuous_on W1 g" + and opeSW1: "openin (subtopology euclidean S') W1" + and "g ` W1 \<subseteq> S" and geqr: "\<And>x. x \<in> W0 \<union> S \<Longrightarrow> g x = r x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> _ \<open>r ` (W0 \<union> S) \<subseteq> S\<close> cloS'WS]) + apply (rule continuous_on_subset [OF contr]) + apply (blast intro: elim: )+ + done + have cloT'WT: "closedin (subtopology euclidean T') (W0 \<union> T)" + by (meson closedin_subset_trans UT cloUT' \<open>T \<subseteq> T'\<close> \<open>W \<subseteq> T'\<close> cloUW cloWW0 + closedin_Un closedin_imp_subset closedin_trans) + obtain W2 h where "W0 \<union> T \<subseteq> W2" and conth: "continuous_on W2 h" + and opeSW2: "openin (subtopology euclidean T') W2" + and "h ` W2 \<subseteq> T" and heqr: "\<And>x. x \<in> W0 \<union> T \<Longrightarrow> h x = r x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> _ \<open>r ` (W0 \<union> T) \<subseteq> T\<close> cloT'WT]) + apply (rule continuous_on_subset [OF contr]) + apply (blast intro: elim: )+ + done + have "S' \<inter> T' = W" + by (force simp: S'_def T'_def W_def) + obtain O1 O2 where "open O1" "W1 = S' \<inter> O1" "open O2" "W2 = T' \<inter> O2" + using opeSW1 opeSW2 by (force simp add: openin_open) + show ?thesis + proof + have eq: "W1 - (W - U0) \<union> (W2 - (W - U0)) = + ((U - T') \<inter> O1 \<union> (U - S') \<inter> O2 \<union> U \<inter> O1 \<inter> O2) - (W - U0)" + using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> + by (auto simp: \<open>S' \<union> T' = U\<close> [symmetric] \<open>S' \<inter> T' = W\<close> [symmetric] \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close>) + show "openin (subtopology euclidean U) (W1 - (W - U0) \<union> (W2 - (W - U0)))" + apply (subst eq) + apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>) + apply simp_all + done + have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))" + using cloUS' apply (simp add: closedin_closed) + apply (erule ex_forward) + using U0 \<open>W0 \<union> S \<subseteq> W1\<close> + apply (auto simp add: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric]) + done + have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))" + using cloUT' apply (simp add: closedin_closed) + apply (erule ex_forward) + using U0 \<open>W0 \<union> T \<subseteq> W2\<close> + apply (auto simp add: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric]) + done + have *: "\<forall>x\<in>S \<union> T. (if x \<in> S' then g x else h x) = x" + using ST \<open>S' \<inter> T' = W\<close> cloT'WT closedin_subset geqr heqr + apply (auto simp: r_def) + apply fastforce + using \<open>S \<subseteq> S'\<close> \<open>T \<subseteq> T'\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W1 = S' \<inter> O1\<close> by auto + have "\<exists>r. continuous_on (W1 - (W - U0) \<union> (W2 - (W - U0))) r \<and> + r ` (W1 - (W - U0) \<union> (W2 - (W - U0))) \<subseteq> S \<union> T \<and> + (\<forall>x\<in>S \<union> T. r x = x)" + apply (rule_tac x = "\<lambda>x. if x \<in> S' then g x else h x" in exI) + apply (intro conjI *) + apply (rule continuous_on_cases_local + [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]]) + using \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<inter> T' = W\<close> + \<open>g ` W1 \<subseteq> S\<close> \<open>h ` W2 \<subseteq> T\<close> apply auto + using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> apply (fastforce simp add: geqr heqr)+ + done + then show "S \<union> T retract_of W1 - (W - U0) \<union> (W2 - (W - U0))" + using \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> ST opeUU0 U0 + by (auto simp add: retract_of_def retraction_def) + qed +qed + + +proposition ANR_closed_Un_local: + fixes S :: "'a::euclidean_space set" + assumes STS: "closedin (subtopology euclidean (S \<union> T)) S" + and STT: "closedin (subtopology euclidean (S \<union> T)) T" + and "ANR S" "ANR T" "ANR(S \<inter> T)" + shows "ANR(S \<union> T)" +proof - + have "\<exists>T. openin (subtopology euclidean U) T \<and> C retract_of T" + if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" + for U and C :: "('a * real) set" + proof - + obtain f g where hom: "homeomorphism (S \<union> T) C f g" + using hom by (force simp: homeomorphic_def) + have US: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> S}" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) + using hom [unfolded homeomorphism_def] apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have UT: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> T}" + apply (rule closedin_trans [OF _ UC]) + apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) + using hom [unfolded homeomorphism_def] apply blast + apply (metis hom homeomorphism_def set_eq_subset) + done + have ANRS: "ANR {x \<in> C. g x \<in> S}" + apply (rule ANR_homeomorphic_ANR [OF \<open>ANR S\<close>]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ANRT: "ANR {x \<in> C. g x \<in> T}" + apply (rule ANR_homeomorphic_ANR [OF \<open>ANR T\<close>]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have ANRI: "ANR ({x \<in> C. g x \<in> S} \<inter> {x \<in> C. g x \<in> T})" + apply (rule ANR_homeomorphic_ANR [OF \<open>ANR (S \<inter> T)\<close>]) + apply (simp add: homeomorphic_def) + apply (rule_tac x=g in exI) + apply (rule_tac x=f in exI) + using hom + apply (auto simp: homeomorphism_def elim!: continuous_on_subset) + apply (rule_tac x="f x" in image_eqI, auto) + done + have "C = {x. x \<in> C \<and> g x \<in> S} \<union> {x. x \<in> C \<and> g x \<in> T}" + by auto (metis Un_iff hom homeomorphism_def imageI) + then show ?thesis + by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI]) + qed + then show ?thesis + by (auto simp: ANR_def) +qed + +corollary ANR_closed_Un: + fixes S :: "'a::euclidean_space set" + shows "\<lbrakk>closed S; closed T; ANR S; ANR T; ANR (S \<inter> T)\<rbrakk> \<Longrightarrow> ANR (S \<union> T)" +by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int) + +lemma ANR_openin: + fixes S :: "'a::euclidean_space set" + assumes "ANR T" and opeTS: "openin (subtopology euclidean T) S" + shows "ANR S" +proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) + fix f :: "'a \<times> real \<Rightarrow> 'a" and U C + assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> S" + and cloUC: "closedin (subtopology euclidean U) C" + have "f ` C \<subseteq> T" + using fim opeTS openin_imp_subset by blast + obtain W g where "C \<subseteq> W" + and UW: "openin (subtopology euclidean U) W" + and contg: "continuous_on W g" + and gim: "g ` W \<subseteq> T" + and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = f x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf \<open>f ` C \<subseteq> T\<close> cloUC]) + using fim by auto + show "\<exists>V g. C \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x\<in>C. g x = f x)" + proof (intro exI conjI) + show "C \<subseteq> {x \<in> W. g x \<in> S}" + using \<open>C \<subseteq> W\<close> fim geq by blast + show "openin (subtopology euclidean U) {x \<in> W. g x \<in> S}" + by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) + show "continuous_on {x \<in> W. g x \<in> S} g" + by (blast intro: continuous_on_subset [OF contg]) + show "g ` {x \<in> W. g x \<in> S} \<subseteq> S" + using gim by blast + show "\<forall>x\<in>C. g x = f x" + using geq by blast + qed +qed + +lemma ENR_openin: + fixes S :: "'a::euclidean_space set" + assumes "ENR T" and opeTS: "openin (subtopology euclidean T) S" + shows "ENR S" + using assms apply (simp add: ENR_ANR) + using ANR_openin locally_open_subset by blast + +lemma ANR_neighborhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ANR U" "S retract_of T" "openin (subtopology euclidean U) T" + shows "ANR S" + using ANR_openin ANR_retract_of_ANR assms by blast + +lemma ENR_neighborhood_retract: + fixes S :: "'a::euclidean_space set" + assumes "ENR U" "S retract_of T" "openin (subtopology euclidean U) T" + shows "ENR S" + using ENR_openin ENR_retract_of_ENR assms by blast + +lemma ANR_rel_interior: + fixes S :: "'a::euclidean_space set" + shows "ANR S \<Longrightarrow> ANR(rel_interior S)" + by (blast intro: ANR_openin openin_set_rel_interior) + +lemma ANR_delete: + fixes S :: "'a::euclidean_space set" + shows "ANR S \<Longrightarrow> ANR(S - {a})" + by (blast intro: ANR_openin openin_delete openin_subtopology_self) + +lemma ENR_rel_interior: + fixes S :: "'a::euclidean_space set" + shows "ENR S \<Longrightarrow> ENR(rel_interior S)" + by (blast intro: ENR_openin openin_set_rel_interior) + +lemma ENR_delete: + fixes S :: "'a::euclidean_space set" + shows "ENR S \<Longrightarrow> ENR(S - {a})" + by (blast intro: ENR_openin openin_delete openin_subtopology_self) + +lemma open_imp_ENR: "open S \<Longrightarrow> ENR S" + using ENR_def by blast + +lemma open_imp_ANR: + fixes S :: "'a::euclidean_space set" + shows "open S \<Longrightarrow> ANR S" + by (simp add: ENR_imp_ANR open_imp_ENR) + +lemma ANR_ball [iff]: + fixes a :: "'a::euclidean_space" + shows "ANR(ball a r)" + by (simp add: convex_imp_ANR) + +lemma ENR_ball [iff]: "ENR(ball a r)" + by (simp add: open_imp_ENR) + +lemma AR_ball [simp]: + fixes a :: "'a::euclidean_space" + shows "AR(ball a r) \<longleftrightarrow> 0 < r" + by (auto simp: AR_ANR convex_imp_contractible) + +lemma ANR_cball [iff]: + fixes a :: "'a::euclidean_space" + shows "ANR(cball a r)" + by (simp add: convex_imp_ANR) + +lemma ENR_cball: + fixes a :: "'a::euclidean_space" + shows "ENR(cball a r)" + using ENR_convex_closed by blast + +lemma AR_cball [simp]: + fixes a :: "'a::euclidean_space" + shows "AR(cball a r) \<longleftrightarrow> 0 \<le> r" + by (auto simp: AR_ANR convex_imp_contractible) + +lemma ANR_box [iff]: + fixes a :: "'a::euclidean_space" + shows "ANR(cbox a b)" "ANR(box a b)" + by (auto simp: convex_imp_ANR open_imp_ANR) + +lemma ENR_box [iff]: + fixes a :: "'a::euclidean_space" + shows "ENR(cbox a b)" "ENR(box a b)" +apply (simp add: ENR_convex_closed closed_cbox) +by (simp add: open_box open_imp_ENR) + +lemma AR_box [simp]: + "AR(cbox a b) \<longleftrightarrow> cbox a b \<noteq> {}" "AR(box a b) \<longleftrightarrow> box a b \<noteq> {}" + by (auto simp: AR_ANR convex_imp_contractible) + +lemma ANR_interior: + fixes S :: "'a::euclidean_space set" + shows "ANR(interior S)" + by (simp add: open_imp_ANR) + +lemma ENR_interior: + fixes S :: "'a::euclidean_space set" + shows "ENR(interior S)" + by (simp add: open_imp_ENR) + +lemma AR_imp_contractible: + fixes S :: "'a::euclidean_space set" + shows "AR S \<Longrightarrow> contractible S" + by (simp add: AR_ANR) + +lemma ENR_imp_locally_compact: + fixes S :: "'a::euclidean_space set" + shows "ENR S \<Longrightarrow> locally compact S" + by (simp add: ENR_ANR) + +lemma ANR_imp_locally_path_connected: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" + shows "locally path_connected S" +proof - + obtain U and T :: "('a \<times> real) set" + where "convex U" "U \<noteq> {}" + and UT: "closedin (subtopology euclidean U) T" + and "S homeomorphic T" + apply (rule homeomorphic_closedin_convex [of S]) + using aff_dim_le_DIM [of S] apply auto + done + have "locally path_connected T" + by (meson ANR_imp_absolute_neighbourhood_retract \<open>S homeomorphic T\<close> \<open>closedin (subtopology euclidean U) T\<close> \<open>convex U\<close> assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) + then have S: "locally path_connected S" + if "openin (subtopology euclidean U) V" "T retract_of V" "U \<noteq> {}" for V + using \<open>S homeomorphic T\<close> homeomorphic_locally homeomorphic_path_connectedness by blast + show ?thesis + using assms + apply (clarsimp simp: ANR_def) + apply (drule_tac x=U in spec) + apply (drule_tac x=T in spec) + using \<open>S homeomorphic T\<close> \<open>U \<noteq> {}\<close> UT apply (blast intro: S) + done +qed + +lemma ANR_imp_locally_connected: + fixes S :: "'a::euclidean_space set" + assumes "ANR S" + shows "locally connected S" +using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto + +lemma AR_imp_locally_path_connected: + fixes S :: "'a::euclidean_space set" + assumes "AR S" + shows "locally path_connected S" +by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms) + +lemma AR_imp_locally_connected: + fixes S :: "'a::euclidean_space set" + assumes "AR S" + shows "locally connected S" +using ANR_imp_locally_connected AR_ANR assms by blast + +lemma ENR_imp_locally_path_connected: + fixes S :: "'a::euclidean_space set" + assumes "ENR S" + shows "locally path_connected S" +by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms) + +lemma ENR_imp_locally_connected: + fixes S :: "'a::euclidean_space set" + assumes "ENR S" + shows "locally connected S" +using ANR_imp_locally_connected ENR_ANR assms by blast + +lemma ANR_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "ANR S" "ANR T" shows "ANR(S \<times> T)" +proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) + fix f :: " ('a \<times> 'b) \<times> real \<Rightarrow> 'a \<times> 'b" and U C + assume "continuous_on C f" and fim: "f ` C \<subseteq> S \<times> T" + and cloUC: "closedin (subtopology euclidean U) C" + have contf1: "continuous_on C (fst \<circ> f)" + by (simp add: \<open>continuous_on C f\<close> continuous_on_fst) + obtain W1 g where "C \<subseteq> W1" + and UW1: "openin (subtopology euclidean U) W1" + and contg: "continuous_on W1 g" + and gim: "g ` W1 \<subseteq> S" + and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = (fst \<circ> f) x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> contf1 _ cloUC]) + using fim apply auto + done + have contf2: "continuous_on C (snd \<circ> f)" + by (simp add: \<open>continuous_on C f\<close> continuous_on_snd) + obtain W2 h where "C \<subseteq> W2" + and UW2: "openin (subtopology euclidean U) W2" + and conth: "continuous_on W2 h" + and him: "h ` W2 \<subseteq> T" + and heq: "\<And>x. x \<in> C \<Longrightarrow> h x = (snd \<circ> f) x" + apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf2 _ cloUC]) + using fim apply auto + done + show "\<exists>V g. C \<subseteq> V \<and> + openin (subtopology euclidean U) V \<and> + continuous_on V g \<and> g ` V \<subseteq> S \<times> T \<and> (\<forall>x\<in>C. g x = f x)" + proof (intro exI conjI) + show "C \<subseteq> W1 \<inter> W2" + by (simp add: \<open>C \<subseteq> W1\<close> \<open>C \<subseteq> W2\<close>) + show "openin (subtopology euclidean U) (W1 \<inter> W2)" + by (simp add: UW1 UW2 openin_Int) + show "continuous_on (W1 \<inter> W2) (\<lambda>x. (g x, h x))" + by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) + show "(\<lambda>x. (g x, h x)) ` (W1 \<inter> W2) \<subseteq> S \<times> T" + using gim him by blast + show "(\<forall>x\<in>C. (g x, h x) = f x)" + using geq heq by auto + qed +qed + +lemma AR_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "AR S" "AR T" shows "AR(S \<times> T)" +using assms by (simp add: AR_ANR ANR_Times contractible_Times) + + +lemma ENR_rel_frontier_convex: + fixes S :: "'a::euclidean_space set" + assumes "bounded S" "convex S" + shows "ENR(rel_frontier S)" +proof (cases "S = {}") + case True then show ?thesis + by simp +next + case False + with assms have "rel_interior S \<noteq> {}" + by (simp add: rel_interior_eq_empty) + then obtain a where a: "a \<in> rel_interior S" + by auto + have ahS: "affine hull S - {a} \<subseteq> {x. closest_point (affine hull S) x \<noteq> a}" + by (auto simp: closest_point_self) + have "rel_frontier S retract_of affine hull S - {a}" + by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull) + also have "... retract_of {x. closest_point (affine hull S) x \<noteq> a}" + apply (simp add: retract_of_def retraction_def ahS) + apply (rule_tac x="closest_point (affine hull S)" in exI) + apply (auto simp add: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point) + done + finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \<noteq> a}" . + moreover have "openin (subtopology euclidean UNIV) {x \<in> UNIV. closest_point (affine hull S) x \<in> - {a}}" + apply (rule continuous_openin_preimage_gen) + apply (auto simp add: False affine_imp_convex continuous_on_closest_point) + done + ultimately show ?thesis + apply (simp add: ENR_def) + apply (rule_tac x = "{x. x \<in> UNIV \<and> + closest_point (affine hull S) x \<in> (- {a})}" in exI) + apply (simp add: open_openin) + done +qed + +lemma ANR_rel_frontier_convex: + fixes S :: "'a::euclidean_space set" + assumes "bounded S" "convex S" + shows "ANR(rel_frontier S)" +by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms) + +(*UNUSED +lemma ENR_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "ENR S" "ENR T" shows "ENR(S \<times> T)" +using assms apply (simp add: ENR_ANR ANR_Times) +thm locally_compact_Times +oops + SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);; +*) + +subsection\<open>Borsuk homotopy extension theorem\<close> + +text\<open>It's only this late so we can use the concept of retraction, + saying that the domain sets or range set are ENRs.\<close> + +theorem Borsuk_homotopy_extension_homotopic: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes cloTS: "closedin (subtopology euclidean T) S" + and anr: "(ANR S \<and> ANR T) \<or> ANR U" + and contf: "continuous_on T f" + and "f ` T \<subseteq> U" + and "homotopic_with (\<lambda>x. True) S U f g" + obtains g' where "homotopic_with (\<lambda>x. True) T U f g'" + "continuous_on T g'" "image g' T \<subseteq> U" + "\<And>x. x \<in> S \<Longrightarrow> g' x = g x" +proof - + have "S \<subseteq> T" using assms closedin_imp_subset by blast + obtain h where conth: "continuous_on ({0..1} \<times> S) h" + and him: "h ` ({0..1} \<times> S) \<subseteq> U" + and [simp]: "\<And>x. h(0, x) = f x" "\<And>x. h(1::real, x) = g x" + using assms by (auto simp: homotopic_with_def) + define h' where "h' \<equiv> \<lambda>z. if snd z \<in> S then h z else (f o snd) z" + define B where "B \<equiv> {0::real} \<times> T \<union> {0..1} \<times> S" + have clo0T: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0::real} \<times> T)" + by (simp add: closedin_subtopology_refl closedin_Times) + moreover have cloT1S: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0..1} \<times> S)" + by (simp add: closedin_subtopology_refl closedin_Times assms) + ultimately have clo0TB:"closedin (subtopology euclidean ({0..1} \<times> T)) B" + by (auto simp: B_def) + have cloBS: "closedin (subtopology euclidean B) ({0..1} \<times> S)" + by (metis (no_types) Un_subset_iff B_def closedin_subset_trans [OF cloT1S] clo0TB closedin_imp_subset closedin_self) + moreover have cloBT: "closedin (subtopology euclidean B) ({0} \<times> T)" + using \<open>S \<subseteq> T\<close> closedin_subset_trans [OF clo0T] + by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1) + moreover have "continuous_on ({0} \<times> T) (f \<circ> snd)" + apply (rule continuous_intros)+ + apply (simp add: contf) + done + ultimately have conth': "continuous_on B h'" + apply (simp add: h'_def B_def Un_commute [of "{0} \<times> T"]) + apply (auto intro!: continuous_on_cases_local conth) + done + have "image h' B \<subseteq> U" + using \<open>f ` T \<subseteq> U\<close> him by (auto simp: h'_def B_def) + obtain V k where "B \<subseteq> V" and opeTV: "openin (subtopology euclidean ({0..1} \<times> T)) V" + and contk: "continuous_on V k" and kim: "k ` V \<subseteq> U" + and keq: "\<And>x. x \<in> B \<Longrightarrow> k x = h' x" + using anr + proof + assume ST: "ANR S \<and> ANR T" + have eq: "({0} \<times> T \<inter> {0..1} \<times> S) = {0::real} \<times> S" + using \<open>S \<subseteq> T\<close> by auto + have "ANR B" + apply (simp add: B_def) + apply (rule ANR_closed_Un_local) + apply (metis cloBT B_def) + apply (metis Un_commute cloBS B_def) + apply (simp_all add: ANR_Times convex_imp_ANR ANR_singleton ST eq) + done + note Vk = that + have *: thesis if "openin (subtopology euclidean ({0..1::real} \<times> T)) V" + "retraction V B r" for V r + using that + apply (clarsimp simp add: retraction_def) + apply (rule Vk [of V "h' o r"], assumption+) + apply (metis continuous_on_compose conth' continuous_on_subset) + using \<open>h' ` B \<subseteq> U\<close> apply force+ + done + show thesis + apply (rule ANR_imp_neighbourhood_retract [OF \<open>ANR B\<close> clo0TB]) + apply (auto simp: ANR_Times ANR_singleton ST retract_of_def *) + done + next + assume "ANR U" + with ANR_imp_absolute_neighbourhood_extensor \<open>h' ` B \<subseteq> U\<close> clo0TB conth' that + show ?thesis by blast + qed + define S' where "S' \<equiv> {x. \<exists>u::real. u \<in> {0..1} \<and> (u, x::'a) \<in> {0..1} \<times> T - V}" + have "closedin (subtopology euclidean T) S'" + unfolding S'_def + apply (rule closedin_compact_projection, blast) + using closedin_self opeTV by blast + have S'_def: "S' = {x. \<exists>u::real. (u, x::'a) \<in> {0..1} \<times> T - V}" + by (auto simp: S'_def) + have cloTS': "closedin (subtopology euclidean T) S'" + using S'_def \<open>closedin (subtopology euclidean T) S'\<close> by blast + have "S \<inter> S' = {}" + using S'_def B_def \<open>B \<subseteq> V\<close> by force + obtain a :: "'a \<Rightarrow> real" where conta: "continuous_on T a" + and "\<And>x. x \<in> T \<Longrightarrow> a x \<in> closed_segment 1 0" + and a1: "\<And>x. x \<in> S \<Longrightarrow> a x = 1" + and a0: "\<And>x. x \<in> S' \<Longrightarrow> a x = 0" + apply (rule Urysohn_local [OF cloTS cloTS' \<open>S \<inter> S' = {}\<close>, of 1 0], blast) + done + then have ain: "\<And>x. x \<in> T \<Longrightarrow> a x \<in> {0..1}" + using closed_segment_eq_real_ivl by auto + have inV: "(u * a t, t) \<in> V" if "t \<in> T" "0 \<le> u" "u \<le> 1" for t u + proof (rule ccontr) + assume "(u * a t, t) \<notin> V" + with ain [OF \<open>t \<in> T\<close>] have "a t = 0" + apply simp + apply (rule a0) + by (metis (no_types, lifting) Diff_iff S'_def SigmaI atLeastAtMost_iff mem_Collect_eq mult_le_one mult_nonneg_nonneg that) + show False + using B_def \<open>(u * a t, t) \<notin> V\<close> \<open>B \<subseteq> V\<close> \<open>a t = 0\<close> that by auto + qed + show ?thesis + proof + show hom: "homotopic_with (\<lambda>x. True) T U f (\<lambda>x. k (a x, x))" + proof (simp add: homotopic_with, intro exI conjI) + show "continuous_on ({0..1} \<times> T) (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z)))" + apply (intro continuous_on_compose continuous_intros) + apply (rule continuous_on_subset [OF conta], force) + apply (rule continuous_on_subset [OF contk]) + apply (force intro: inV) + done + show "(k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) ` ({0..1} \<times> T) \<subseteq> U" + using inV kim by auto + show "\<forall>x\<in>T. (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) (0, x) = f x" + by (simp add: B_def h'_def keq) + show "\<forall>x\<in>T. (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) (1, x) = k (a x, x)" + by auto + qed + show "continuous_on T (\<lambda>x. k (a x, x))" + using hom homotopic_with_imp_continuous by blast + show "(\<lambda>x. k (a x, x)) ` T \<subseteq> U" + proof clarify + fix t + assume "t \<in> T" + show "k (a t, t) \<in> U" + by (metis \<open>t \<in> T\<close> image_subset_iff inV kim not_one_le_zero linear mult_cancel_right1) + qed + show "\<And>x. x \<in> S \<Longrightarrow> k (a x, x) = g x" + by (simp add: B_def a1 h'_def keq) + qed +qed + + +corollary nullhomotopic_into_ANR_extension: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "closed S" + and contf: "continuous_on S f" + and "ANR T" + and fim: "f ` S \<subseteq> T" + and "S \<noteq> {}" + shows "(\<exists>c. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)) \<longleftrightarrow> + (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> T \<and> (\<forall>x \<in> S. g x = f x))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain c where c: "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) f" + by (blast intro: homotopic_with_symD elim: ) + have "closedin (subtopology euclidean UNIV) S" + using \<open>closed S\<close> closed_closedin by fastforce + then obtain g where "continuous_on UNIV g" "range g \<subseteq> T" + "\<And>x. x \<in> S \<Longrightarrow> g x = f x" + apply (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ c, where T=UNIV]) + using \<open>ANR T\<close> \<open>S \<noteq> {}\<close> c homotopic_with_imp_subset1 apply fastforce+ + done + then show ?rhs by blast +next + assume ?rhs + then obtain g where "continuous_on UNIV g" "range g \<subseteq> T" "\<And>x. x\<in>S \<Longrightarrow> g x = f x" + by blast + then obtain c where "homotopic_with (\<lambda>h. True) UNIV T g (\<lambda>x. c)" + using nullhomotopic_from_contractible [of UNIV g T] contractible_UNIV by blast + then show ?lhs + apply (rule_tac x="c" in exI) + apply (rule homotopic_with_eq [of _ _ _ g "\<lambda>x. c"]) + apply (rule homotopic_with_subset_left) + apply (auto simp add: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>) + done +qed + +corollary nullhomotopic_into_rel_frontier_extension: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "closed S" + and contf: "continuous_on S f" + and "convex T" "bounded T" + and fim: "f ` S \<subseteq> rel_frontier T" + and "S \<noteq> {}" + shows "(\<exists>c. homotopic_with (\<lambda>x. True) S (rel_frontier T) f (\<lambda>x. c)) \<longleftrightarrow> + (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> rel_frontier T \<and> (\<forall>x \<in> S. g x = f x))" +by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex) + +corollary nullhomotopic_into_sphere_extension: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: euclidean_space" + assumes "closed S" and contf: "continuous_on S f" + and "S \<noteq> {}" and fim: "f ` S \<subseteq> sphere a r" + shows "((\<exists>c. homotopic_with (\<lambda>x. True) S (sphere a r) f (\<lambda>x. c)) \<longleftrightarrow> + (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> sphere a r \<and> (\<forall>x \<in> S. g x = f x)))" + (is "?lhs = ?rhs") +proof (cases "r = 0") + case True with fim show ?thesis + apply (auto simp: ) + using fim continuous_on_const apply fastforce + by (metis contf contractible_sing nullhomotopic_into_contractible) +next + case False + then have eq: "sphere a r = rel_frontier (cball a r)" by simp + show ?thesis + using fim unfolding eq + apply (rule nullhomotopic_into_rel_frontier_extension [OF \<open>closed S\<close> contf convex_cball bounded_cball]) + apply (rule \<open>S \<noteq> {}\<close>) + done +qed + +proposition Borsuk_map_essential_bounded_component: + fixes a :: "'a :: euclidean_space" + assumes "compact S" and "a \<notin> S" + shows "bounded (connected_component_set (- S) a) \<longleftrightarrow> + ~(\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) + (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\<lambda>x. c))" + (is "?lhs = ?rhs") +proof (cases "S = {}") + case True then show ?thesis + by simp +next + case False + have "closed S" "bounded S" + using \<open>compact S\<close> compact_eq_bounded_closed by auto + have s01: "(\<lambda>x. (x - a) /\<^sub>R norm (x - a)) ` S \<subseteq> sphere 0 1" + using \<open>a \<notin> S\<close> by clarsimp (metis dist_eq_0_iff dist_norm mult.commute right_inverse) + have aincc: "a \<in> connected_component_set (- S) a" + by (simp add: \<open>a \<notin> S\<close>) + obtain r where "r>0" and r: "S \<subseteq> ball 0 r" + using bounded_subset_ballD \<open>bounded S\<close> by blast + have "~ ?rhs \<longleftrightarrow> ~ ?lhs" + proof + assume notr: "~ ?rhs" + have nog: "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and> + g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and> + (\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))" + if "bounded (connected_component_set (- S) a)" + apply (rule non_extensible_Borsuk_map [OF \<open>compact S\<close> componentsI _ aincc]) + using \<open>a \<notin> S\<close> that by auto + obtain g where "range g \<subseteq> sphere 0 1" "continuous_on UNIV g" + "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)" + using notr + by (auto simp add: nullhomotopic_into_sphere_extension + [OF \<open>closed S\<close> continuous_on_Borsuk_map [OF \<open>a \<notin> S\<close>] False s01]) + with \<open>a \<notin> S\<close> show "~ ?lhs" + apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog) + apply (drule_tac x="g" in spec) + using continuous_on_subset by fastforce + next + assume "~ ?lhs" + then obtain b where b: "b \<in> connected_component_set (- S) a" and "r \<le> norm b" + using bounded_iff linear by blast + then have bnot: "b \<notin> ball 0 r" + by simp + have "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) + (\<lambda>x. (x - b) /\<^sub>R norm (x - b))" + apply (rule Borsuk_maps_homotopic_in_path_component) + using \<open>closed S\<close> b open_Compl open_path_connected_component apply fastforce + done + moreover + obtain c where "homotopic_with (\<lambda>x. True) (ball 0 r) (sphere 0 1) + (\<lambda>x. inverse (norm (x - b)) *\<^sub>R (x - b)) (\<lambda>x. c)" + proof (rule nullhomotopic_from_contractible) + show "contractible (ball (0::'a) r)" + by (metis convex_imp_contractible convex_ball) + show "continuous_on (ball 0 r) (\<lambda>x. inverse(norm (x - b)) *\<^sub>R (x - b))" + by (rule continuous_on_Borsuk_map [OF bnot]) + show "(\<lambda>x. (x - b) /\<^sub>R norm (x - b)) ` ball 0 r \<subseteq> sphere 0 1" + using bnot Borsuk_map_into_sphere by blast + qed blast + ultimately have "homotopic_with (\<lambda>x. True) S (sphere 0 1) + (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) (\<lambda>x. c)" + by (meson homotopic_with_subset_left homotopic_with_trans r) + then show "~ ?rhs" + by blast + qed + then show ?thesis by blast +qed + + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Caratheodory.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,891 @@ +(* Title: HOL/Analysis/Caratheodory.thy + Author: Lawrence C Paulson + Author: Johannes Hölzl, TU München +*) + +section \<open>Caratheodory Extension Theorem\<close> + +theory Caratheodory + imports Measure_Space +begin + +text \<open> + Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson. +\<close> + +lemma suminf_ennreal_2dimen: + fixes f:: "nat \<times> nat \<Rightarrow> ennreal" + assumes "\<And>m. g m = (\<Sum>n. f (m,n))" + shows "(\<Sum>i. f (prod_decode i)) = suminf g" +proof - + have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))" + using assms by (simp add: fun_eq_iff) + have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)" + by (simp add: setsum.reindex[OF inj_prod_decode] comp_def) + have "(SUP n. \<Sum>i<n. f (prod_decode i)) = (SUP p : UNIV \<times> UNIV. \<Sum>i<fst p. \<Sum>n<snd p. f (i, n))" + proof (intro SUP_eq; clarsimp simp: setsum.cartesian_product reindex) + fix n + let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))" + { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x" + then have "a < ?M fst" "b < ?M snd" + by (auto intro!: Max_ge le_imp_less_Suc image_eqI) } + then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})" + by (auto intro!: setsum_mono3) + then show "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto + next + fix a b + let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}" + { fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M" + by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) } + then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M" + by (auto intro!: setsum_mono3) + then show "\<exists>n. setsum f ({..<a} \<times> {..<b}) \<le> setsum f (prod_decode ` {..<n})" + by auto + qed + also have "\<dots> = (SUP p. \<Sum>i<p. \<Sum>n. f (i, n))" + unfolding suminf_setsum[OF summableI, symmetric] + by (simp add: suminf_eq_SUP SUP_pair setsum.commute[of _ "{..< fst _}"]) + finally show ?thesis unfolding g_def + by (simp add: suminf_eq_SUP) +qed + +subsection \<open>Characterizations of Measures\<close> + +definition outer_measure_space where + "outer_measure_space M f \<longleftrightarrow> positive M f \<and> increasing M f \<and> countably_subadditive M f" + +subsubsection \<open>Lambda Systems\<close> + +definition lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set" +where + "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}" + +lemma (in algebra) lambda_system_eq: + "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (x \<inter> l) + f (x - l) = f x}" +proof - + have [simp]: "\<And>l x. l \<in> M \<Longrightarrow> x \<in> M \<Longrightarrow> (\<Omega> - l) \<inter> x = x - l" + by (metis Int_Diff Int_absorb1 Int_commute sets_into_space) + show ?thesis + by (auto simp add: lambda_system_def) (metis Int_commute)+ +qed + +lemma (in algebra) lambda_system_empty: "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f" + by (auto simp add: positive_def lambda_system_eq) + +lemma lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M" + by (simp add: lambda_system_def) + +lemma (in algebra) lambda_system_Compl: + fixes f:: "'a set \<Rightarrow> ennreal" + assumes x: "x \<in> lambda_system \<Omega> M f" + shows "\<Omega> - x \<in> lambda_system \<Omega> M f" +proof - + have "x \<subseteq> \<Omega>" + by (metis sets_into_space lambda_system_sets x) + hence "\<Omega> - (\<Omega> - x) = x" + by (metis double_diff equalityE) + with x show ?thesis + by (force simp add: lambda_system_def ac_simps) +qed + +lemma (in algebra) lambda_system_Int: + fixes f:: "'a set \<Rightarrow> ennreal" + assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" + shows "x \<inter> y \<in> lambda_system \<Omega> M f" +proof - + from xl yl show ?thesis + proof (auto simp add: positive_def lambda_system_eq Int) + fix u + assume x: "x \<in> M" and y: "y \<in> M" and u: "u \<in> M" + and fx: "\<forall>z\<in>M. f (z \<inter> x) + f (z - x) = f z" + and fy: "\<forall>z\<in>M. f (z \<inter> y) + f (z - y) = f z" + have "u - x \<inter> y \<in> M" + by (metis Diff Diff_Int Un u x y) + moreover + have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast + moreover + have "u - x \<inter> y - y = u - y" by blast + ultimately + have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy + by force + have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) + = (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)" + by (simp add: ey ac_simps) + also have "... = (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)" + by (simp add: Int_ac) + also have "... = f (u \<inter> y) + f (u - y)" + using fx [THEN bspec, of "u \<inter> y"] Int y u + by force + also have "... = f u" + by (metis fy u) + finally show "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) = f u" . + qed +qed + +lemma (in algebra) lambda_system_Un: + fixes f:: "'a set \<Rightarrow> ennreal" + assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" + shows "x \<union> y \<in> lambda_system \<Omega> M f" +proof - + have "(\<Omega> - x) \<inter> (\<Omega> - y) \<in> M" + by (metis Diff_Un Un compl_sets lambda_system_sets xl yl) + moreover + have "x \<union> y = \<Omega> - ((\<Omega> - x) \<inter> (\<Omega> - y))" + by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+ + ultimately show ?thesis + by (metis lambda_system_Compl lambda_system_Int xl yl) +qed + +lemma (in algebra) lambda_system_algebra: + "positive M f \<Longrightarrow> algebra \<Omega> (lambda_system \<Omega> M f)" + apply (auto simp add: algebra_iff_Un) + apply (metis lambda_system_sets set_mp sets_into_space) + apply (metis lambda_system_empty) + apply (metis lambda_system_Compl) + apply (metis lambda_system_Un) + done + +lemma (in algebra) lambda_system_strong_additive: + assumes z: "z \<in> M" and disj: "x \<inter> y = {}" + and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" + shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)" +proof - + have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast + moreover + have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast + moreover + have "(z \<inter> (x \<union> y)) \<in> M" + by (metis Int Un lambda_system_sets xl yl z) + ultimately show ?thesis using xl yl + by (simp add: lambda_system_eq) +qed + +lemma (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f" +proof (auto simp add: additive_def) + fix x and y + assume disj: "x \<inter> y = {}" + and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" + hence "x \<in> M" "y \<in> M" by (blast intro: lambda_system_sets)+ + thus "f (x \<union> y) = f x + f y" + using lambda_system_strong_additive [OF top disj xl yl] + by (simp add: Un) +qed + +lemma lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f" + by (simp add: increasing_def lambda_system_def) + +lemma lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f" + by (simp add: positive_def lambda_system_def) + +lemma (in algebra) lambda_system_strong_sum: + fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ennreal" + assumes f: "positive M f" and a: "a \<in> M" + and A: "range A \<subseteq> lambda_system \<Omega> M f" + and disj: "disjoint_family A" + shows "(\<Sum>i = 0..<n. f (a \<inter>A i)) = f (a \<inter> (\<Union>i\<in>{0..<n}. A i))" +proof (induct n) + case 0 show ?case using f by (simp add: positive_def) +next + case (Suc n) + have 2: "A n \<inter> UNION {0..<n} A = {}" using disj + by (force simp add: disjoint_family_on_def neq_iff) + have 3: "A n \<in> lambda_system \<Omega> M f" using A + by blast + interpret l: algebra \<Omega> "lambda_system \<Omega> M f" + using f by (rule lambda_system_algebra) + have 4: "UNION {0..<n} A \<in> lambda_system \<Omega> M f" + using A l.UNION_in_sets by simp + from Suc.hyps show ?case + by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4]) +qed + +lemma (in sigma_algebra) lambda_system_caratheodory: + assumes oms: "outer_measure_space M f" + and A: "range A \<subseteq> lambda_system \<Omega> M f" + and disj: "disjoint_family A" + shows "(\<Union>i. A i) \<in> lambda_system \<Omega> M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)" +proof - + have pos: "positive M f" and inc: "increasing M f" + and csa: "countably_subadditive M f" + by (metis oms outer_measure_space_def)+ + have sa: "subadditive M f" + by (metis countably_subadditive_subadditive csa pos) + have A': "\<And>S. A`S \<subseteq> (lambda_system \<Omega> M f)" using A + by auto + interpret ls: algebra \<Omega> "lambda_system \<Omega> M f" + using pos by (rule lambda_system_algebra) + have A'': "range A \<subseteq> M" + by (metis A image_subset_iff lambda_system_sets) + + have U_in: "(\<Union>i. A i) \<in> M" + by (metis A'' countable_UN) + have U_eq: "f (\<Union>i. A i) = (\<Sum>i. f (A i))" + proof (rule antisym) + show "f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))" + using csa[unfolded countably_subadditive_def] A'' disj U_in by auto + have dis: "\<And>N. disjoint_family_on A {..<N}" by (intro disjoint_family_on_mono[OF _ disj]) auto + show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)" + using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] A'' + by (intro suminf_le_const[OF summableI]) (auto intro!: increasingD[OF inc] countable_UN) + qed + have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a" + if a [iff]: "a \<in> M" for a + proof (rule antisym) + have "range (\<lambda>i. a \<inter> A i) \<subseteq> M" using A'' + by blast + moreover + have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj + by (auto simp add: disjoint_family_on_def) + moreover + have "a \<inter> (\<Union>i. A i) \<in> M" + by (metis Int U_in a) + ultimately + have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))" + using csa[unfolded countably_subadditive_def, rule_format, of "(\<lambda>i. a \<inter> A i)"] + by (simp add: o_def) + hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))" + by (rule add_right_mono) + also have "\<dots> \<le> f a" + proof (intro ennreal_suminf_bound_add) + fix n + have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> M" + by (metis A'' UNION_in_sets) + have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A'' + by (blast intro: increasingD [OF inc] A'' UNION_in_sets) + have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system \<Omega> M f" + using ls.UNION_in_sets by (simp add: A) + hence eq_fa: "f a = f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i))" + by (simp add: lambda_system_eq UNION_in) + have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))" + by (blast intro: increasingD [OF inc] UNION_in U_in) + thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a" + by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric]) + qed + finally show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a" + by simp + next + have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))" + by (blast intro: increasingD [OF inc] U_in) + also have "... \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" + by (blast intro: subadditiveD [OF sa] U_in) + finally show "f a \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" . + qed + thus ?thesis + by (simp add: lambda_system_eq sums_iff U_eq U_in) +qed + +lemma (in sigma_algebra) caratheodory_lemma: + assumes oms: "outer_measure_space M f" + defines "L \<equiv> lambda_system \<Omega> M f" + shows "measure_space \<Omega> L f" +proof - + have pos: "positive M f" + by (metis oms outer_measure_space_def) + have alg: "algebra \<Omega> L" + using lambda_system_algebra [of f, OF pos] + by (simp add: algebra_iff_Un L_def) + then + have "sigma_algebra \<Omega> L" + using lambda_system_caratheodory [OF oms] + by (simp add: sigma_algebra_disjoint_iff L_def) + moreover + have "countably_additive L f" "positive L f" + using pos lambda_system_caratheodory [OF oms] + by (auto simp add: lambda_system_sets L_def countably_additive_def positive_def) + ultimately + show ?thesis + using pos by (simp add: measure_space_def) +qed + +definition outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where + "outer_measure M f X = + (INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))" + +lemma (in ring_of_sets) outer_measure_agrees: + assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \<in> M" + shows "outer_measure M f s = f s" + unfolding outer_measure_def +proof (safe intro!: antisym INF_greatest) + fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" and dA: "disjoint_family A" and sA: "s \<subseteq> (\<Union>x. A x)" + have inc: "increasing M f" + by (metis additive_increasing ca countably_additive_additive posf) + have "f s = f (\<Union>i. A i \<inter> s)" + using sA by (auto simp: Int_absorb1) + also have "\<dots> = (\<Sum>i. f (A i \<inter> s))" + using sA dA A s + by (intro ca[unfolded countably_additive_def, rule_format, symmetric]) + (auto simp: Int_absorb1 disjoint_family_on_def) + also have "... \<le> (\<Sum>i. f (A i))" + using A s by (auto intro!: suminf_le increasingD[OF inc]) + finally show "f s \<le> (\<Sum>i. f (A i))" . +next + have "(\<Sum>i. f (if i = 0 then s else {})) \<le> f s" + using positiveD1[OF posf] by (subst suminf_finite[of "{0}"]) auto + with s show "(INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> s \<subseteq> UNION UNIV A}. \<Sum>i. f (A i)) \<le> f s" + by (intro INF_lower2[of "\<lambda>i. if i = 0 then s else {}"]) + (auto simp: disjoint_family_on_def) +qed + +lemma outer_measure_empty: + "positive M f \<Longrightarrow> {} \<in> M \<Longrightarrow> outer_measure M f {} = 0" + unfolding outer_measure_def + by (intro antisym INF_lower2[of "\<lambda>_. {}"]) (auto simp: disjoint_family_on_def positive_def) + +lemma (in ring_of_sets) positive_outer_measure: + assumes "positive M f" shows "positive (Pow \<Omega>) (outer_measure M f)" + unfolding positive_def by (auto simp: assms outer_measure_empty) + +lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)" + by (force simp: increasing_def outer_measure_def intro!: INF_greatest intro: INF_lower) + +lemma (in ring_of_sets) outer_measure_le: + assumes pos: "positive M f" and inc: "increasing M f" and A: "range A \<subseteq> M" and X: "X \<subseteq> (\<Union>i. A i)" + shows "outer_measure M f X \<le> (\<Sum>i. f (A i))" + unfolding outer_measure_def +proof (safe intro!: INF_lower2[of "disjointed A"] del: subsetI) + show dA: "range (disjointed A) \<subseteq> M" + by (auto intro!: A range_disjointed_sets) + have "\<forall>n. f (disjointed A n) \<le> f (A n)" + by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A) + then show "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))" + by (blast intro!: suminf_le) +qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed) + +lemma (in ring_of_sets) outer_measure_close: + "outer_measure M f X < e \<Longrightarrow> \<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) < e" + unfolding outer_measure_def INF_less_iff by auto + +lemma (in ring_of_sets) countably_subadditive_outer_measure: + assumes posf: "positive M f" and inc: "increasing M f" + shows "countably_subadditive (Pow \<Omega>) (outer_measure M f)" +proof (simp add: countably_subadditive_def, safe) + fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> Pow (\<Omega>)" and sb: "(\<Union>i. A i) \<subseteq> \<Omega>" + let ?O = "outer_measure M f" + show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n))" + proof (rule ennreal_le_epsilon) + fix b and e :: real assume "0 < e" "(\<Sum>n. outer_measure M f (A n)) < top" + then have *: "\<And>n. outer_measure M f (A n) < outer_measure M f (A n) + e * (1/2)^Suc n" + by (auto simp add: less_top dest!: ennreal_suminf_lessD) + obtain B + where B: "\<And>n. range (B n) \<subseteq> M" + and sbB: "\<And>n. A n \<subseteq> (\<Union>i. B n i)" + and Ble: "\<And>n. (\<Sum>i. f (B n i)) \<le> ?O (A n) + e * (1/2)^(Suc n)" + by (metis less_imp_le outer_measure_close[OF *]) + + define C where "C = case_prod B o prod_decode" + from B have B_in_M: "\<And>i j. B i j \<in> M" + by (rule range_subsetD) + then have C: "range C \<subseteq> M" + by (auto simp add: C_def split_def) + have A_C: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)" + using sbB by (auto simp add: C_def subset_eq) (metis prod.case prod_encode_inverse) + + have "?O (\<Union>i. A i) \<le> ?O (\<Union>i. C i)" + using A_C A C by (intro increasing_outer_measure[THEN increasingD]) (auto dest!: sets_into_space) + also have "\<dots> \<le> (\<Sum>i. f (C i))" + using C by (intro outer_measure_le[OF posf inc]) auto + also have "\<dots> = (\<Sum>n. \<Sum>i. f (B n i))" + using B_in_M unfolding C_def comp_def by (intro suminf_ennreal_2dimen) auto + also have "\<dots> \<le> (\<Sum>n. ?O (A n) + e * (1/2) ^ Suc n)" + using B_in_M by (intro suminf_le suminf_nonneg allI Ble) auto + also have "... = (\<Sum>n. ?O (A n)) + (\<Sum>n. ennreal e * ennreal ((1/2) ^ Suc n))" + using \<open>0 < e\<close> by (subst suminf_add[symmetric]) + (auto simp del: ennreal_suminf_cmult simp add: ennreal_mult[symmetric]) + also have "\<dots> = (\<Sum>n. ?O (A n)) + e" + unfolding ennreal_suminf_cmult + by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto + finally show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n)) + e" . + qed +qed + +lemma (in ring_of_sets) outer_measure_space_outer_measure: + "positive M f \<Longrightarrow> increasing M f \<Longrightarrow> outer_measure_space (Pow \<Omega>) (outer_measure M f)" + by (simp add: outer_measure_space_def + positive_outer_measure increasing_outer_measure countably_subadditive_outer_measure) + +lemma (in ring_of_sets) algebra_subset_lambda_system: + assumes posf: "positive M f" and inc: "increasing M f" + and add: "additive M f" + shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (outer_measure M f)" +proof (auto dest: sets_into_space + simp add: algebra.lambda_system_eq [OF algebra_Pow]) + fix x s assume x: "x \<in> M" and s: "s \<subseteq> \<Omega>" + have [simp]: "\<And>x. x \<in> M \<Longrightarrow> s \<inter> (\<Omega> - x) = s - x" using s + by blast + have "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le> outer_measure M f s" + unfolding outer_measure_def[of M f s] + proof (safe intro!: INF_greatest) + fix A :: "nat \<Rightarrow> 'a set" assume A: "disjoint_family A" "range A \<subseteq> M" "s \<subseteq> (\<Union>i. A i)" + have "outer_measure M f (s \<inter> x) \<le> (\<Sum>i. f (A i \<inter> x))" + unfolding outer_measure_def + proof (safe intro!: INF_lower2[of "\<lambda>i. A i \<inter> x"]) + from A(1) show "disjoint_family (\<lambda>i. A i \<inter> x)" + by (rule disjoint_family_on_bisimulation) auto + qed (insert x A, auto) + moreover + have "outer_measure M f (s - x) \<le> (\<Sum>i. f (A i - x))" + unfolding outer_measure_def + proof (safe intro!: INF_lower2[of "\<lambda>i. A i - x"]) + from A(1) show "disjoint_family (\<lambda>i. A i - x)" + by (rule disjoint_family_on_bisimulation) auto + qed (insert x A, auto) + ultimately have "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le> + (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono) + also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))" + using A(2) x posf by (subst suminf_add) (auto simp: positive_def) + also have "\<dots> = (\<Sum>i. f (A i))" + using A x + by (subst add[THEN additiveD, symmetric]) + (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f]) + finally show "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le> (\<Sum>i. f (A i))" . + qed + moreover + have "outer_measure M f s \<le> outer_measure M f (s \<inter> x) + outer_measure M f (s - x)" + proof - + have "outer_measure M f s = outer_measure M f ((s \<inter> x) \<union> (s - x))" + by (metis Un_Diff_Int Un_commute) + also have "... \<le> outer_measure M f (s \<inter> x) + outer_measure M f (s - x)" + apply (rule subadditiveD) + apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow]) + apply (simp add: positive_def outer_measure_empty[OF posf]) + apply (rule countably_subadditive_outer_measure) + using s by (auto intro!: posf inc) + finally show ?thesis . + qed + ultimately + show "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) = outer_measure M f s" + by (rule order_antisym) +qed + +lemma measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>" + by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq) + +subsection \<open>Caratheodory's theorem\<close> + +theorem (in ring_of_sets) caratheodory': + assumes posf: "positive M f" and ca: "countably_additive M f" + shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>" +proof - + have inc: "increasing M f" + by (metis additive_increasing ca countably_additive_additive posf) + let ?O = "outer_measure M f" + define ls where "ls = lambda_system \<Omega> (Pow \<Omega>) ?O" + have mls: "measure_space \<Omega> ls ?O" + using sigma_algebra.caratheodory_lemma + [OF sigma_algebra_Pow outer_measure_space_outer_measure [OF posf inc]] + by (simp add: ls_def) + hence sls: "sigma_algebra \<Omega> ls" + by (simp add: measure_space_def) + have "M \<subseteq> ls" + by (simp add: ls_def) + (metis ca posf inc countably_additive_additive algebra_subset_lambda_system) + hence sgs_sb: "sigma_sets (\<Omega>) (M) \<subseteq> ls" + using sigma_algebra.sigma_sets_subset [OF sls, of "M"] + by simp + have "measure_space \<Omega> (sigma_sets \<Omega> M) ?O" + by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets) + (simp_all add: sgs_sb space_closed) + thus ?thesis using outer_measure_agrees [OF posf ca] + by (intro exI[of _ ?O]) auto +qed + +lemma (in ring_of_sets) caratheodory_empty_continuous: + assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>" + assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0" + shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>" +proof (intro caratheodory' empty_continuous_imp_countably_additive f) + show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto +qed (rule cont) + +subsection \<open>Volumes\<close> + +definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where + "volume M f \<longleftrightarrow> + (f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and> + (\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))" + +lemma volumeI: + assumes "f {} = 0" + assumes "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> f a" + assumes "\<And>C. C \<subseteq> M \<Longrightarrow> disjoint C \<Longrightarrow> finite C \<Longrightarrow> \<Union>C \<in> M \<Longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c)" + shows "volume M f" + using assms by (auto simp: volume_def) + +lemma volume_positive: + "volume M f \<Longrightarrow> a \<in> M \<Longrightarrow> 0 \<le> f a" + by (auto simp: volume_def) + +lemma volume_empty: + "volume M f \<Longrightarrow> f {} = 0" + by (auto simp: volume_def) + +lemma volume_finite_additive: + assumes "volume M f" + assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "UNION I A \<in> M" + shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))" +proof - + have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>(A`I) \<in> M" + using A by (auto simp: disjoint_family_on_disjoint_image) + with \<open>volume M f\<close> have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)" + unfolding volume_def by blast + also have "\<dots> = (\<Sum>i\<in>I. f (A i))" + proof (subst setsum.reindex_nontrivial) + fix i j assume "i \<in> I" "j \<in> I" "i \<noteq> j" "A i = A j" + with \<open>disjoint_family_on A I\<close> have "A i = {}" + by (auto simp: disjoint_family_on_def) + then show "f (A i) = 0" + using volume_empty[OF \<open>volume M f\<close>] by simp + qed (auto intro: \<open>finite I\<close>) + finally show "f (UNION I A) = (\<Sum>i\<in>I. f (A i))" + by simp +qed + +lemma (in ring_of_sets) volume_additiveI: + assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a" + assumes [simp]: "\<mu> {} = 0" + assumes add: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> \<mu> (a \<union> b) = \<mu> a + \<mu> b" + shows "volume M \<mu>" +proof (unfold volume_def, safe) + fix C assume "finite C" "C \<subseteq> M" "disjoint C" + then show "\<mu> (\<Union>C) = setsum \<mu> C" + proof (induct C) + case (insert c C) + from insert(1,2,4,5) have "\<mu> (\<Union>insert c C) = \<mu> c + \<mu> (\<Union>C)" + by (auto intro!: add simp: disjoint_def) + with insert show ?case + by (simp add: disjoint_def) + qed simp +qed fact+ + +lemma (in semiring_of_sets) extend_volume: + assumes "volume M \<mu>" + shows "\<exists>\<mu>'. volume generated_ring \<mu>' \<and> (\<forall>a\<in>M. \<mu>' a = \<mu> a)" +proof - + let ?R = generated_ring + have "\<forall>a\<in>?R. \<exists>m. \<exists>C\<subseteq>M. a = \<Union>C \<and> finite C \<and> disjoint C \<and> m = (\<Sum>c\<in>C. \<mu> c)" + by (auto simp: generated_ring_def) + from bchoice[OF this] guess \<mu>' .. note \<mu>'_spec = this + + { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C" + fix D assume D: "D \<subseteq> M" "finite D" "disjoint D" + assume "\<Union>C = \<Union>D" + have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>d\<in>D. \<Sum>c\<in>C. \<mu> (c \<inter> d))" + proof (intro setsum.cong refl) + fix d assume "d \<in> D" + have Un_eq_d: "(\<Union>c\<in>C. c \<inter> d) = d" + using \<open>d \<in> D\<close> \<open>\<Union>C = \<Union>D\<close> by auto + moreover have "\<mu> (\<Union>c\<in>C. c \<inter> d) = (\<Sum>c\<in>C. \<mu> (c \<inter> d))" + proof (rule volume_finite_additive) + { fix c assume "c \<in> C" then show "c \<inter> d \<in> M" + using C D \<open>d \<in> D\<close> by auto } + show "(\<Union>a\<in>C. a \<inter> d) \<in> M" + unfolding Un_eq_d using \<open>d \<in> D\<close> D by auto + show "disjoint_family_on (\<lambda>a. a \<inter> d) C" + using \<open>disjoint C\<close> by (auto simp: disjoint_family_on_def disjoint_def) + qed fact+ + ultimately show "\<mu> d = (\<Sum>c\<in>C. \<mu> (c \<inter> d))" by simp + qed } + note split_sum = this + + { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C" + fix D assume D: "D \<subseteq> M" "finite D" "disjoint D" + assume "\<Union>C = \<Union>D" + with split_sum[OF C D] split_sum[OF D C] + have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)" + by (simp, subst setsum.commute, simp add: ac_simps) } + note sum_eq = this + + { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C" + then have "\<Union>C \<in> ?R" by (auto simp: generated_ring_def) + with \<mu>'_spec[THEN bspec, of "\<Union>C"] + obtain D where + D: "D \<subseteq> M" "finite D" "disjoint D" "\<Union>C = \<Union>D" and "\<mu>' (\<Union>C) = (\<Sum>d\<in>D. \<mu> d)" + by auto + with sum_eq[OF C D] have "\<mu>' (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" by simp } + note \<mu>' = this + + show ?thesis + proof (intro exI conjI ring_of_sets.volume_additiveI[OF generating_ring] ballI) + fix a assume "a \<in> M" with \<mu>'[of "{a}"] show "\<mu>' a = \<mu> a" + by (simp add: disjoint_def) + next + fix a assume "a \<in> ?R" then guess Ca .. note Ca = this + with \<mu>'[of Ca] \<open>volume M \<mu>\<close>[THEN volume_positive] + show "0 \<le> \<mu>' a" + by (auto intro!: setsum_nonneg) + next + show "\<mu>' {} = 0" using \<mu>'[of "{}"] by auto + next + fix a assume "a \<in> ?R" then guess Ca .. note Ca = this + fix b assume "b \<in> ?R" then guess Cb .. note Cb = this + assume "a \<inter> b = {}" + with Ca Cb have "Ca \<inter> Cb \<subseteq> {{}}" by auto + then have C_Int_cases: "Ca \<inter> Cb = {{}} \<or> Ca \<inter> Cb = {}" by auto + + from \<open>a \<inter> b = {}\<close> have "\<mu>' (\<Union>(Ca \<union> Cb)) = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c)" + using Ca Cb by (intro \<mu>') (auto intro!: disjoint_union) + also have "\<dots> = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c) + (\<Sum>c\<in>Ca \<inter> Cb. \<mu> c)" + using C_Int_cases volume_empty[OF \<open>volume M \<mu>\<close>] by (elim disjE) simp_all + also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)" + using Ca Cb by (simp add: setsum.union_inter) + also have "\<dots> = \<mu>' a + \<mu>' b" + using Ca Cb by (simp add: \<mu>') + finally show "\<mu>' (a \<union> b) = \<mu>' a + \<mu>' b" + using Ca Cb by simp + qed +qed + +subsubsection \<open>Caratheodory on semirings\<close> + +theorem (in semiring_of_sets) caratheodory: + assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>" + shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'" +proof - + have "volume M \<mu>" + proof (rule volumeI) + { fix a assume "a \<in> M" then show "0 \<le> \<mu> a" + using pos unfolding positive_def by auto } + note p = this + + fix C assume sets_C: "C \<subseteq> M" "\<Union>C \<in> M" and "disjoint C" "finite C" + have "\<exists>F'. bij_betw F' {..<card C} C" + by (rule finite_same_card_bij[OF _ \<open>finite C\<close>]) auto + then guess F' .. note F' = this + then have F': "C = F' ` {..< card C}" "inj_on F' {..< card C}" + by (auto simp: bij_betw_def) + { fix i j assume *: "i < card C" "j < card C" "i \<noteq> j" + with F' have "F' i \<in> C" "F' j \<in> C" "F' i \<noteq> F' j" + unfolding inj_on_def by auto + with \<open>disjoint C\<close>[THEN disjointD] + have "F' i \<inter> F' j = {}" + by auto } + note F'_disj = this + define F where "F i = (if i < card C then F' i else {})" for i + then have "disjoint_family F" + using F'_disj by (auto simp: disjoint_family_on_def) + moreover from F' have "(\<Union>i. F i) = \<Union>C" + by (auto simp add: F_def split: if_split_asm) blast + moreover have sets_F: "\<And>i. F i \<in> M" + using F' sets_C by (auto simp: F_def) + moreover note sets_C + ultimately have "\<mu> (\<Union>C) = (\<Sum>i. \<mu> (F i))" + using ca[unfolded countably_additive_def, THEN spec, of F] by auto + also have "\<dots> = (\<Sum>i<card C. \<mu> (F' i))" + proof - + have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) sums (\<Sum>i<card C. \<mu> (F' i))" + by (rule sums_If_finite_set) auto + also have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) = (\<lambda>i. \<mu> (F i))" + using pos by (auto simp: positive_def F_def) + finally show "(\<Sum>i. \<mu> (F i)) = (\<Sum>i<card C. \<mu> (F' i))" + by (simp add: sums_iff) + qed + also have "\<dots> = (\<Sum>c\<in>C. \<mu> c)" + using F'(2) by (subst (2) F') (simp add: setsum.reindex) + finally show "\<mu> (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" . + next + show "\<mu> {} = 0" + using \<open>positive M \<mu>\<close> by (rule positiveD1) + qed + from extend_volume[OF this] obtain \<mu>_r where + V: "volume generated_ring \<mu>_r" "\<And>a. a \<in> M \<Longrightarrow> \<mu> a = \<mu>_r a" + by auto + + interpret G: ring_of_sets \<Omega> generated_ring + by (rule generating_ring) + + have pos: "positive generated_ring \<mu>_r" + using V unfolding positive_def by (auto simp: positive_def intro!: volume_positive volume_empty) + + have "countably_additive generated_ring \<mu>_r" + proof (rule countably_additiveI) + fix A' :: "nat \<Rightarrow> 'a set" assume A': "range A' \<subseteq> generated_ring" "disjoint_family A'" + and Un_A: "(\<Union>i. A' i) \<in> generated_ring" + + from generated_ringE[OF Un_A] guess C' . note C' = this + + { fix c assume "c \<in> C'" + moreover define A where [abs_def]: "A i = A' i \<inter> c" for i + ultimately have A: "range A \<subseteq> generated_ring" "disjoint_family A" + and Un_A: "(\<Union>i. A i) \<in> generated_ring" + using A' C' + by (auto intro!: G.Int G.finite_Union intro: generated_ringI_Basic simp: disjoint_family_on_def) + from A C' \<open>c \<in> C'\<close> have UN_eq: "(\<Union>i. A i) = c" + by (auto simp: A_def) + + have "\<forall>i::nat. \<exists>f::nat \<Rightarrow> 'a set. \<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j)) \<and> disjoint_family f \<and> \<Union>range f = A i \<and> (\<forall>j. f j \<in> M)" + (is "\<forall>i. ?P i") + proof + fix i + from A have Ai: "A i \<in> generated_ring" by auto + from generated_ringE[OF this] guess C . note C = this + + have "\<exists>F'. bij_betw F' {..<card C} C" + by (rule finite_same_card_bij[OF _ \<open>finite C\<close>]) auto + then guess F .. note F = this + define f where [abs_def]: "f i = (if i < card C then F i else {})" for i + then have f: "bij_betw f {..< card C} C" + by (intro bij_betw_cong[THEN iffD1, OF _ F]) auto + with C have "\<forall>j. f j \<in> M" + by (auto simp: Pi_iff f_def dest!: bij_betw_imp_funcset) + moreover + from f C have d_f: "disjoint_family_on f {..<card C}" + by (intro disjoint_image_disjoint_family_on) (auto simp: bij_betw_def) + then have "disjoint_family f" + by (auto simp: disjoint_family_on_def f_def) + moreover + have Ai_eq: "A i = (\<Union>x<card C. f x)" + using f C Ai unfolding bij_betw_def by auto + then have "\<Union>range f = A i" + using f C Ai unfolding bij_betw_def + by (auto simp add: f_def cong del: strong_SUP_cong) + moreover + { have "(\<Sum>j. \<mu>_r (f j)) = (\<Sum>j. if j \<in> {..< card C} then \<mu>_r (f j) else 0)" + using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def) + also have "\<dots> = (\<Sum>j<card C. \<mu>_r (f j))" + by (rule sums_If_finite_set[THEN sums_unique, symmetric]) simp + also have "\<dots> = \<mu>_r (A i)" + using C f[THEN bij_betw_imp_funcset] unfolding Ai_eq + by (intro volume_finite_additive[OF V(1) _ d_f, symmetric]) + (auto simp: Pi_iff Ai_eq intro: generated_ringI_Basic) + finally have "\<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j))" .. } + ultimately show "?P i" + by blast + qed + from choice[OF this] guess f .. note f = this + then have UN_f_eq: "(\<Union>i. case_prod f (prod_decode i)) = (\<Union>i. A i)" + unfolding UN_extend_simps surj_prod_decode by (auto simp: set_eq_iff) + + have d: "disjoint_family (\<lambda>i. case_prod f (prod_decode i))" + unfolding disjoint_family_on_def + proof (intro ballI impI) + fix m n :: nat assume "m \<noteq> n" + then have neq: "prod_decode m \<noteq> prod_decode n" + using inj_prod_decode[of UNIV] by (auto simp: inj_on_def) + show "case_prod f (prod_decode m) \<inter> case_prod f (prod_decode n) = {}" + proof cases + assume "fst (prod_decode m) = fst (prod_decode n)" + then show ?thesis + using neq f by (fastforce simp: disjoint_family_on_def) + next + assume neq: "fst (prod_decode m) \<noteq> fst (prod_decode n)" + have "case_prod f (prod_decode m) \<subseteq> A (fst (prod_decode m))" + "case_prod f (prod_decode n) \<subseteq> A (fst (prod_decode n))" + using f[THEN spec, of "fst (prod_decode m)"] + using f[THEN spec, of "fst (prod_decode n)"] + by (auto simp: set_eq_iff) + with f A neq show ?thesis + by (fastforce simp: disjoint_family_on_def subset_eq set_eq_iff) + qed + qed + from f have "(\<Sum>n. \<mu>_r (A n)) = (\<Sum>n. \<mu>_r (case_prod f (prod_decode n)))" + by (intro suminf_ennreal_2dimen[symmetric] generated_ringI_Basic) + (auto split: prod.split) + also have "\<dots> = (\<Sum>n. \<mu> (case_prod f (prod_decode n)))" + using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split) + also have "\<dots> = \<mu> (\<Union>i. case_prod f (prod_decode i))" + using f \<open>c \<in> C'\<close> C' + by (intro ca[unfolded countably_additive_def, rule_format]) + (auto split: prod.split simp: UN_f_eq d UN_eq) + finally have "(\<Sum>n. \<mu>_r (A' n \<inter> c)) = \<mu> c" + using UN_f_eq UN_eq by (simp add: A_def) } + note eq = this + + have "(\<Sum>n. \<mu>_r (A' n)) = (\<Sum>n. \<Sum>c\<in>C'. \<mu>_r (A' n \<inter> c))" + using C' A' + by (subst volume_finite_additive[symmetric, OF V(1)]) + (auto simp: disjoint_def disjoint_family_on_def + intro!: G.Int G.finite_Union arg_cong[where f="\<lambda>X. suminf (\<lambda>i. \<mu>_r (X i))"] ext + intro: generated_ringI_Basic) + also have "\<dots> = (\<Sum>c\<in>C'. \<Sum>n. \<mu>_r (A' n \<inter> c))" + using C' A' + by (intro suminf_setsum G.Int G.finite_Union) (auto intro: generated_ringI_Basic) + also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)" + using eq V C' by (auto intro!: setsum.cong) + also have "\<dots> = \<mu>_r (\<Union>C')" + using C' Un_A + by (subst volume_finite_additive[symmetric, OF V(1)]) + (auto simp: disjoint_family_on_def disjoint_def + intro: generated_ringI_Basic) + finally show "(\<Sum>n. \<mu>_r (A' n)) = \<mu>_r (\<Union>i. A' i)" + using C' by simp + qed + from G.caratheodory'[OF \<open>positive generated_ring \<mu>_r\<close> \<open>countably_additive generated_ring \<mu>_r\<close>] + guess \<mu>' .. + with V show ?thesis + unfolding sigma_sets_generated_ring_eq + by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic) +qed + +lemma extend_measure_caratheodory: + fixes G :: "'i \<Rightarrow> 'a set" + assumes M: "M = extend_measure \<Omega> I G \<mu>" + assumes "i \<in> I" + assumes "semiring_of_sets \<Omega> (G ` I)" + assumes empty: "\<And>i. i \<in> I \<Longrightarrow> G i = {} \<Longrightarrow> \<mu> i = 0" + assumes inj: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> G i = G j \<Longrightarrow> \<mu> i = \<mu> j" + assumes nonneg: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> \<mu> i" + assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>j. A \<in> UNIV \<rightarrow> I \<Longrightarrow> j \<in> I \<Longrightarrow> disjoint_family (G \<circ> A) \<Longrightarrow> + (\<Union>i. G (A i)) = G j \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> j" + shows "emeasure M (G i) = \<mu> i" +proof - + interpret semiring_of_sets \<Omega> "G ` I" + by fact + have "\<forall>g\<in>G`I. \<exists>i\<in>I. g = G i" + by auto + then obtain sel where sel: "\<And>g. g \<in> G ` I \<Longrightarrow> sel g \<in> I" "\<And>g. g \<in> G ` I \<Longrightarrow> G (sel g) = g" + by metis + + have "\<exists>\<mu>'. (\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'" + proof (rule caratheodory) + show "positive (G ` I) (\<lambda>s. \<mu> (sel s))" + by (auto simp: positive_def intro!: empty sel nonneg) + show "countably_additive (G ` I) (\<lambda>s. \<mu> (sel s))" + proof (rule countably_additiveI) + fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> G ` I" "disjoint_family A" "(\<Union>i. A i) \<in> G ` I" + then show "(\<Sum>i. \<mu> (sel (A i))) = \<mu> (sel (\<Union>i. A i))" + by (intro add) (auto simp: sel image_subset_iff_funcset comp_def Pi_iff intro!: sel) + qed + qed + then obtain \<mu>' where \<mu>': "\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)" "measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'" + by metis + + show ?thesis + proof (rule emeasure_extend_measure[OF M]) + { fix i assume "i \<in> I" then show "\<mu>' (G i) = \<mu> i" + using \<mu>' by (auto intro!: inj sel) } + show "G ` I \<subseteq> Pow \<Omega>" + by fact + then show "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'" + using \<mu>' by (simp_all add: M sets_extend_measure measure_space_def) + qed fact +qed + +lemma extend_measure_caratheodory_pair: + fixes G :: "'i \<Rightarrow> 'j \<Rightarrow> 'a set" + assumes M: "M = extend_measure \<Omega> {(a, b). P a b} (\<lambda>(a, b). G a b) (\<lambda>(a, b). \<mu> a b)" + assumes "P i j" + assumes semiring: "semiring_of_sets \<Omega> {G a b | a b. P a b}" + assumes empty: "\<And>i j. P i j \<Longrightarrow> G i j = {} \<Longrightarrow> \<mu> i j = 0" + assumes inj: "\<And>i j k l. P i j \<Longrightarrow> P k l \<Longrightarrow> G i j = G k l \<Longrightarrow> \<mu> i j = \<mu> k l" + assumes nonneg: "\<And>i j. P i j \<Longrightarrow> 0 \<le> \<mu> i j" + assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>B::nat \<Rightarrow> 'j. \<And>j k. + (\<And>n. P (A n) (B n)) \<Longrightarrow> P j k \<Longrightarrow> disjoint_family (\<lambda>n. G (A n) (B n)) \<Longrightarrow> + (\<Union>i. G (A i) (B i)) = G j k \<Longrightarrow> (\<Sum>n. \<mu> (A n) (B n)) = \<mu> j k" + shows "emeasure M (G i j) = \<mu> i j" +proof - + have "emeasure M ((\<lambda>(a, b). G a b) (i, j)) = (\<lambda>(a, b). \<mu> a b) (i, j)" + proof (rule extend_measure_caratheodory[OF M]) + show "semiring_of_sets \<Omega> ((\<lambda>(a, b). G a b) ` {(a, b). P a b})" + using semiring by (simp add: image_def conj_commute) + next + fix A :: "nat \<Rightarrow> ('i \<times> 'j)" and j assume "A \<in> UNIV \<rightarrow> {(a, b). P a b}" "j \<in> {(a, b). P a b}" + "disjoint_family ((\<lambda>(a, b). G a b) \<circ> A)" + "(\<Union>i. case A i of (a, b) \<Rightarrow> G a b) = (case j of (a, b) \<Rightarrow> G a b)" + then show "(\<Sum>n. case A n of (a, b) \<Rightarrow> \<mu> a b) = (case j of (a, b) \<Rightarrow> \<mu> a b)" + using add[of "\<lambda>i. fst (A i)" "\<lambda>i. snd (A i)" "fst j" "snd j"] + by (simp add: split_beta' comp_def Pi_iff) + qed (auto split: prod.splits intro: assms) + then show ?thesis by simp +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,1426 @@ +section \<open>Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space.\<close> + +theory Cartesian_Euclidean_Space +imports Finite_Cartesian_Product Henstock_Kurzweil_Integration +begin + +lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}" + by (simp add: subspace_def) + +lemma delta_mult_idempotent: + "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" + by simp + +lemma setsum_UNIV_sum: + fixes g :: "'a::finite + 'b::finite \<Rightarrow> _" + shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))" + apply (subst UNIV_Plus_UNIV [symmetric]) + apply (subst setsum.Plus) + apply simp_all + done + +lemma setsum_mult_product: + "setsum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))" + unfolding setsum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric] +proof (rule setsum.cong, simp, rule setsum.reindex_cong) + fix i + show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI) + show "{i * B..<i * B + B} = (\<lambda>j. j + i * B) ` {..<B}" + proof safe + fix j assume "j \<in> {i * B..<i * B + B}" + then show "j \<in> (\<lambda>j. j + i * B) ` {..<B}" + by (auto intro!: image_eqI[of _ _ "j - i * B"]) + qed simp +qed simp + + +subsection\<open>Basic componentwise operations on vectors.\<close> + +instantiation vec :: (times, finite) times +begin + +definition "op * \<equiv> (\<lambda> x y. (\<chi> i. (x$i) * (y$i)))" +instance .. + +end + +instantiation vec :: (one, finite) one +begin + +definition "1 \<equiv> (\<chi> i. 1)" +instance .. + +end + +instantiation vec :: (ord, finite) ord +begin + +definition "x \<le> y \<longleftrightarrow> (\<forall>i. x$i \<le> y$i)" +definition "x < (y::'a^'b) \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x" +instance .. + +end + +text\<open>The ordering on one-dimensional vectors is linear.\<close> + +class cart_one = + assumes UNIV_one: "card (UNIV :: 'a set) = Suc 0" +begin + +subclass finite +proof + from UNIV_one show "finite (UNIV :: 'a set)" + by (auto intro!: card_ge_0_finite) +qed + +end + +instance vec:: (order, finite) order + by standard (auto simp: less_eq_vec_def less_vec_def vec_eq_iff + intro: order.trans order.antisym order.strict_implies_order) + +instance vec :: (linorder, cart_one) linorder +proof + obtain a :: 'b where all: "\<And>P. (\<forall>i. P i) \<longleftrightarrow> P a" + proof - + have "card (UNIV :: 'b set) = Suc 0" by (rule UNIV_one) + then obtain b :: 'b where "UNIV = {b}" by (auto iff: card_Suc_eq) + then have "\<And>P. (\<forall>i\<in>UNIV. P i) \<longleftrightarrow> P b" by auto + then show thesis by (auto intro: that) + qed + fix x y :: "'a^'b::cart_one" + note [simp] = less_eq_vec_def less_vec_def all vec_eq_iff field_simps + show "x \<le> y \<or> y \<le> x" by auto +qed + +text\<open>Constant Vectors\<close> + +definition "vec x = (\<chi> i. x)" + +lemma interval_cbox_cart: "{a::real^'n..b} = cbox a b" + by (auto simp add: less_eq_vec_def mem_box Basis_vec_def inner_axis) + +text\<open>Also the scalar-vector multiplication.\<close> + +definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70) + where "c *s x = (\<chi> i. c * (x$i))" + + +subsection \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space.\<close> + +lemma setsum_cong_aux: + "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A" + by (auto intro: setsum.cong) + +hide_fact (open) setsum_cong_aux + +method_setup vector = \<open> +let + val ss1 = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps [@{thm setsum.distrib} RS sym, + @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, + @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym]) + val ss2 = + simpset_of (@{context} addsimps + [@{thm plus_vec_def}, @{thm times_vec_def}, + @{thm minus_vec_def}, @{thm uminus_vec_def}, + @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def}, + @{thm scaleR_vec_def}, + @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}]) + fun vector_arith_tac ctxt ths = + simp_tac (put_simpset ss1 ctxt) + THEN' (fn i => resolve_tac ctxt @{thms Cartesian_Euclidean_Space.setsum_cong_aux} i + ORELSE resolve_tac ctxt @{thms setsum.neutral} i + ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i) + (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) + THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths) +in + Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths)) +end +\<close> "lift trivial vector statements to real arith statements" + +lemma vec_0[simp]: "vec 0 = 0" by vector +lemma vec_1[simp]: "vec 1 = 1" by vector + +lemma vec_inj[simp]: "vec x = vec y \<longleftrightarrow> x = y" by vector + +lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto + +lemma vec_add: "vec(x + y) = vec x + vec y" by vector +lemma vec_sub: "vec(x - y) = vec x - vec y" by vector +lemma vec_cmul: "vec(c * x) = c *s vec x " by vector +lemma vec_neg: "vec(- x) = - vec x " by vector + +lemma vec_setsum: + assumes "finite S" + shows "vec(setsum f S) = setsum (vec \<circ> f) S" + using assms +proof induct + case empty + then show ?case by simp +next + case insert + then show ?case by (auto simp add: vec_add) +qed + +text\<open>Obvious "component-pushing".\<close> + +lemma vec_component [simp]: "vec x $ i = x" + by vector + +lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i" + by vector + +lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)" + by vector + +lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector + +lemmas vector_component = + vec_component vector_add_component vector_mult_component + vector_smult_component vector_minus_component vector_uminus_component + vector_scaleR_component cond_component + + +subsection \<open>Some frequently useful arithmetic lemmas over vectors.\<close> + +instance vec :: (semigroup_mult, finite) semigroup_mult + by standard (vector mult.assoc) + +instance vec :: (monoid_mult, finite) monoid_mult + by standard vector+ + +instance vec :: (ab_semigroup_mult, finite) ab_semigroup_mult + by standard (vector mult.commute) + +instance vec :: (comm_monoid_mult, finite) comm_monoid_mult + by standard vector + +instance vec :: (semiring, finite) semiring + by standard (vector field_simps)+ + +instance vec :: (semiring_0, finite) semiring_0 + by standard (vector field_simps)+ +instance vec :: (semiring_1, finite) semiring_1 + by standard vector +instance vec :: (comm_semiring, finite) comm_semiring + by standard (vector field_simps)+ + +instance vec :: (comm_semiring_0, finite) comm_semiring_0 .. +instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add .. +instance vec :: (semiring_0_cancel, finite) semiring_0_cancel .. +instance vec :: (comm_semiring_0_cancel, finite) comm_semiring_0_cancel .. +instance vec :: (ring, finite) ring .. +instance vec :: (semiring_1_cancel, finite) semiring_1_cancel .. +instance vec :: (comm_semiring_1, finite) comm_semiring_1 .. + +instance vec :: (ring_1, finite) ring_1 .. + +instance vec :: (real_algebra, finite) real_algebra + by standard (simp_all add: vec_eq_iff) + +instance vec :: (real_algebra_1, finite) real_algebra_1 .. + +lemma of_nat_index: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" +proof (induct n) + case 0 + then show ?case by vector +next + case Suc + then show ?case by vector +qed + +lemma one_index [simp]: "(1 :: 'a :: one ^ 'n) $ i = 1" + by vector + +lemma neg_one_index [simp]: "(- 1 :: 'a :: {one, uminus} ^ 'n) $ i = - 1" + by vector + +instance vec :: (semiring_char_0, finite) semiring_char_0 +proof + fix m n :: nat + show "inj (of_nat :: nat \<Rightarrow> 'a ^ 'b)" + by (auto intro!: injI simp add: vec_eq_iff of_nat_index) +qed + +instance vec :: (numeral, finite) numeral .. +instance vec :: (semiring_numeral, finite) semiring_numeral .. + +lemma numeral_index [simp]: "numeral w $ i = numeral w" + by (induct w) (simp_all only: numeral.simps vector_add_component one_index) + +lemma neg_numeral_index [simp]: "- numeral w $ i = - numeral w" + by (simp only: vector_uminus_component numeral_index) + +instance vec :: (comm_ring_1, finite) comm_ring_1 .. +instance vec :: (ring_char_0, finite) ring_char_0 .. + +lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" + by (vector mult.assoc) +lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" + by (vector field_simps) +lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" + by (vector field_simps) +lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector +lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector +lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" + by (vector field_simps) +lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector +lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector +lemma vector_sneg_minus1: "-x = (-1::'a::ring_1) *s x" by vector +lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector +lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" + by (vector field_simps) + +lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)" + by (simp add: vec_eq_iff) + +lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero) +lemma vector_mul_eq_0[simp]: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0" + by vector +lemma vector_mul_lcancel[simp]: "a *s x = a *s y \<longleftrightarrow> a = (0::real) \<or> x = y" + by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib) +lemma vector_mul_rcancel[simp]: "a *s x = b *s x \<longleftrightarrow> (a::real) = b \<or> x = 0" + by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib) +lemma vector_mul_lcancel_imp: "a \<noteq> (0::real) ==> a *s x = a *s y ==> (x = y)" + by (metis vector_mul_lcancel) +lemma vector_mul_rcancel_imp: "x \<noteq> 0 \<Longrightarrow> (a::real) *s x = b *s x ==> a = b" + by (metis vector_mul_rcancel) + +lemma component_le_norm_cart: "\<bar>x$i\<bar> <= norm x" + apply (simp add: norm_vec_def) + apply (rule member_le_setL2, simp_all) + done + +lemma norm_bound_component_le_cart: "norm x <= e ==> \<bar>x$i\<bar> <= e" + by (metis component_le_norm_cart order_trans) + +lemma norm_bound_component_lt_cart: "norm x < e ==> \<bar>x$i\<bar> < e" + by (metis component_le_norm_cart le_less_trans) + +lemma norm_le_l1_cart: "norm x <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV" + by (simp add: norm_vec_def setL2_le_setsum) + +lemma scalar_mult_eq_scaleR: "c *s x = c *\<^sub>R x" + unfolding scaleR_vec_def vector_scalar_mult_def by simp + +lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y" + unfolding dist_norm scalar_mult_eq_scaleR + unfolding scaleR_right_diff_distrib[symmetric] by simp + +lemma setsum_component [simp]: + fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n" + shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S" +proof (cases "finite S") + case True + then show ?thesis by induct simp_all +next + case False + then show ?thesis by simp +qed + +lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)" + by (simp add: vec_eq_iff) + +lemma setsum_cmul: + fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n" + shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S" + by (simp add: vec_eq_iff setsum_right_distrib) + +lemma setsum_norm_allsubsets_bound_cart: + fixes f:: "'a \<Rightarrow> real ^'n" + assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e" + shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) * e" + using setsum_norm_allsubsets_bound[OF assms] + by simp + +subsection\<open>Closures and interiors of halfspaces\<close> + +lemma interior_halfspace_le [simp]: + assumes "a \<noteq> 0" + shows "interior {x. a \<bullet> x \<le> b} = {x. a \<bullet> x < b}" +proof - + have *: "a \<bullet> x < b" if x: "x \<in> S" and S: "S \<subseteq> {x. a \<bullet> x \<le> b}" and "open S" for S x + proof - + obtain e where "e>0" and e: "cball x e \<subseteq> S" + using \<open>open S\<close> open_contains_cball x by blast + then have "x + (e / norm a) *\<^sub>R a \<in> cball x e" + by (simp add: dist_norm) + then have "x + (e / norm a) *\<^sub>R a \<in> S" + using e by blast + then have "x + (e / norm a) *\<^sub>R a \<in> {x. a \<bullet> x \<le> b}" + using S by blast + moreover have "e * (a \<bullet> a) / norm a > 0" + by (simp add: \<open>0 < e\<close> assms) + ultimately show ?thesis + by (simp add: algebra_simps) + qed + show ?thesis + by (rule interior_unique) (auto simp: open_halfspace_lt *) +qed + +lemma interior_halfspace_ge [simp]: + "a \<noteq> 0 \<Longrightarrow> interior {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x > b}" +using interior_halfspace_le [of "-a" "-b"] by simp + +lemma interior_halfspace_component_le [simp]: + "interior {x. x$k \<le> a} = {x :: (real,'n::finite) vec. x$k < a}" (is "?LE") + and interior_halfspace_component_ge [simp]: + "interior {x. x$k \<ge> a} = {x :: (real,'n::finite) vec. x$k > a}" (is "?GE") +proof - + have "axis k (1::real) \<noteq> 0" + by (simp add: axis_def vec_eq_iff) + moreover have "axis k (1::real) \<bullet> x = x$k" for x + by (simp add: cart_eq_inner_axis inner_commute) + ultimately show ?LE ?GE + using interior_halfspace_le [of "axis k (1::real)" a] + interior_halfspace_ge [of "axis k (1::real)" a] by auto +qed + +lemma closure_halfspace_lt [simp]: + assumes "a \<noteq> 0" + shows "closure {x. a \<bullet> x < b} = {x. a \<bullet> x \<le> b}" +proof - + have [simp]: "-{x. a \<bullet> x < b} = {x. a \<bullet> x \<ge> b}" + by (force simp:) + then show ?thesis + using interior_halfspace_ge [of a b] assms + by (force simp: closure_interior) +qed + +lemma closure_halfspace_gt [simp]: + "a \<noteq> 0 \<Longrightarrow> closure {x. a \<bullet> x > b} = {x. a \<bullet> x \<ge> b}" +using closure_halfspace_lt [of "-a" "-b"] by simp + +lemma closure_halfspace_component_lt [simp]: + "closure {x. x$k < a} = {x :: (real,'n::finite) vec. x$k \<le> a}" (is "?LE") + and closure_halfspace_component_gt [simp]: + "closure {x. x$k > a} = {x :: (real,'n::finite) vec. x$k \<ge> a}" (is "?GE") +proof - + have "axis k (1::real) \<noteq> 0" + by (simp add: axis_def vec_eq_iff) + moreover have "axis k (1::real) \<bullet> x = x$k" for x + by (simp add: cart_eq_inner_axis inner_commute) + ultimately show ?LE ?GE + using closure_halfspace_lt [of "axis k (1::real)" a] + closure_halfspace_gt [of "axis k (1::real)" a] by auto +qed + +lemma interior_hyperplane [simp]: + assumes "a \<noteq> 0" + shows "interior {x. a \<bullet> x = b} = {}" +proof - + have [simp]: "{x. a \<bullet> x = b} = {x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x \<ge> b}" + by (force simp:) + then show ?thesis + by (auto simp: assms) +qed + +lemma frontier_halfspace_le: + assumes "a \<noteq> 0 \<or> b \<noteq> 0" + shows "frontier {x. a \<bullet> x \<le> b} = {x. a \<bullet> x = b}" +proof (cases "a = 0") + case True with assms show ?thesis by simp +next + case False then show ?thesis + by (force simp: frontier_def closed_halfspace_le) +qed + +lemma frontier_halfspace_ge: + assumes "a \<noteq> 0 \<or> b \<noteq> 0" + shows "frontier {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x = b}" +proof (cases "a = 0") + case True with assms show ?thesis by simp +next + case False then show ?thesis + by (force simp: frontier_def closed_halfspace_ge) +qed + +lemma frontier_halfspace_lt: + assumes "a \<noteq> 0 \<or> b \<noteq> 0" + shows "frontier {x. a \<bullet> x < b} = {x. a \<bullet> x = b}" +proof (cases "a = 0") + case True with assms show ?thesis by simp +next + case False then show ?thesis + by (force simp: frontier_def interior_open open_halfspace_lt) +qed + +lemma frontier_halfspace_gt: + assumes "a \<noteq> 0 \<or> b \<noteq> 0" + shows "frontier {x. a \<bullet> x > b} = {x. a \<bullet> x = b}" +proof (cases "a = 0") + case True with assms show ?thesis by simp +next + case False then show ?thesis + by (force simp: frontier_def interior_open open_halfspace_gt) +qed + +lemma interior_standard_hyperplane: + "interior {x :: (real,'n::finite) vec. x$k = a} = {}" +proof - + have "axis k (1::real) \<noteq> 0" + by (simp add: axis_def vec_eq_iff) + moreover have "axis k (1::real) \<bullet> x = x$k" for x + by (simp add: cart_eq_inner_axis inner_commute) + ultimately show ?thesis + using interior_hyperplane [of "axis k (1::real)" a] + by force +qed + +subsection \<open>Matrix operations\<close> + +text\<open>Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"}\<close> + +definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m" + (infixl "**" 70) + where "m ** m' == (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" + +definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm" + (infixl "*v" 70) + where "m *v x \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" + +definition vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n " + (infixl "v*" 70) + where "v v* m == (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" + +definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)" +definition transpose where + "(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))" +definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))" +definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))" +definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}" +definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}" + +lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) +lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)" + by (vector matrix_matrix_mult_def setsum.distrib[symmetric] field_simps) + +lemma matrix_mul_lid: + fixes A :: "'a::semiring_1 ^ 'm ^ 'n" + shows "mat 1 ** A = A" + apply (simp add: matrix_matrix_mult_def mat_def) + apply vector + apply (auto simp only: if_distrib cond_application_beta setsum.delta'[OF finite] + mult_1_left mult_zero_left if_True UNIV_I) + done + + +lemma matrix_mul_rid: + fixes A :: "'a::semiring_1 ^ 'm ^ 'n" + shows "A ** mat 1 = A" + apply (simp add: matrix_matrix_mult_def mat_def) + apply vector + apply (auto simp only: if_distrib cond_application_beta setsum.delta[OF finite] + mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) + done + +lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" + apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult.assoc) + apply (subst setsum.commute) + apply simp + done + +lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" + apply (vector matrix_matrix_mult_def matrix_vector_mult_def + setsum_right_distrib setsum_left_distrib mult.assoc) + apply (subst setsum.commute) + apply simp + done + +lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" + apply (vector matrix_vector_mult_def mat_def) + apply (simp add: if_distrib cond_application_beta setsum.delta' cong del: if_weak_cong) + done + +lemma matrix_transpose_mul: + "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" + by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute) + +lemma matrix_eq: + fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" + shows "A = B \<longleftrightarrow> (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs") + apply auto + apply (subst vec_eq_iff) + apply clarify + apply (clarsimp simp add: matrix_vector_mult_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong) + apply (erule_tac x="axis ia 1" in allE) + apply (erule_tac x="i" in allE) + apply (auto simp add: if_distrib cond_application_beta axis_def + setsum.delta[OF finite] cong del: if_weak_cong) + done + +lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = (A$k) \<bullet> x" + by (simp add: matrix_vector_mult_def inner_vec_def) + +lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)" + apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib ac_simps) + apply (subst setsum.commute) + apply simp + done + +lemma transpose_mat: "transpose (mat n) = mat n" + by (vector transpose_def mat_def) + +lemma transpose_transpose: "transpose(transpose A) = A" + by (vector transpose_def) + +lemma row_transpose: + fixes A:: "'a::semiring_1^_^_" + shows "row i (transpose A) = column i A" + by (simp add: row_def column_def transpose_def vec_eq_iff) + +lemma column_transpose: + fixes A:: "'a::semiring_1^_^_" + shows "column i (transpose A) = row i A" + by (simp add: row_def column_def transpose_def vec_eq_iff) + +lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A" + by (auto simp add: rows_def columns_def row_transpose intro: set_eqI) + +lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" + by (metis transpose_transpose rows_transpose) + +text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close> + +lemma matrix_mult_dot: "A *v x = (\<chi> i. A$i \<bullet> x)" + by (simp add: matrix_vector_mult_def inner_vec_def) + +lemma matrix_mult_vsum: + "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)" + by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute) + +lemma vector_componentwise: + "(x::'a::ring_1^'n) = (\<chi> j. \<Sum>i\<in>UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)" + by (simp add: axis_def if_distrib setsum.If_cases vec_eq_iff) + +lemma basis_expansion: "setsum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)" + by (auto simp add: axis_def vec_eq_iff if_distrib setsum.If_cases cong del: if_weak_cong) + +lemma linear_componentwise: + fixes f:: "real ^'m \<Rightarrow> real ^ _" + assumes lf: "linear f" + shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") +proof - + let ?M = "(UNIV :: 'm set)" + let ?N = "(UNIV :: 'n set)" + have "?rhs = (setsum (\<lambda>i.(x$i) *\<^sub>R f (axis i 1) ) ?M)$j" + unfolding setsum_component by simp + then show ?thesis + unfolding linear_setsum_mul[OF lf, symmetric] + unfolding scalar_mult_eq_scaleR[symmetric] + unfolding basis_expansion + by simp +qed + +text\<open>Inverse matrices (not necessarily square)\<close> + +definition + "invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)" + +definition + "matrix_inv(A:: 'a::semiring_1^'n^'m) = + (SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)" + +text\<open>Correspondence between matrices and linear operators.\<close> + +definition matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n" + where "matrix f = (\<chi> i j. (f(axis j 1))$i)" + +lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::real ^ _))" + by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff + field_simps setsum_right_distrib setsum.distrib) + +lemma matrix_works: + assumes lf: "linear f" + shows "matrix f *v x = f (x::real ^ 'n)" + apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute) + apply clarify + apply (rule linear_componentwise[OF lf, symmetric]) + done + +lemma matrix_vector_mul: "linear f ==> f = (\<lambda>x. matrix f *v (x::real ^ 'n))" + by (simp add: ext matrix_works) + +lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: real ^ 'n)) = A" + by (simp add: matrix_eq matrix_vector_mul_linear matrix_works) + +lemma matrix_compose: + assumes lf: "linear (f::real^'n \<Rightarrow> real^'m)" + and lg: "linear (g::real^'m \<Rightarrow> real^_)" + shows "matrix (g \<circ> f) = matrix g ** matrix f" + using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]] + by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) + +lemma matrix_vector_column: + "(A::'a::comm_semiring_1^'n^_) *v x = setsum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" + by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute) + +lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)" + apply (rule adjoint_unique) + apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def + setsum_left_distrib setsum_right_distrib) + apply (subst setsum.commute) + apply (auto simp add: ac_simps) + done + +lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)" + shows "matrix(adjoint f) = transpose(matrix f)" + apply (subst matrix_vector_mul[OF lf]) + unfolding adjoint_matrix matrix_of_matrix_vector_mul + apply rule + done + + +subsection \<open>lambda skolemization on cartesian products\<close> + +(* FIXME: rename do choice_cart *) + +lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow> + (\<exists>x::'a ^ 'n. \<forall>i. P i (x $ i))" (is "?lhs \<longleftrightarrow> ?rhs") +proof - + let ?S = "(UNIV :: 'n set)" + { assume H: "?rhs" + then have ?lhs by auto } + moreover + { assume H: "?lhs" + then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis + let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n" + { fix i + from f have "P i (f i)" by metis + then have "P i (?x $ i)" by auto + } + hence "\<forall>i. P i (?x$i)" by metis + hence ?rhs by metis } + ultimately show ?thesis by metis +qed + +lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0" + unfolding inner_simps scalar_mult_eq_scaleR by auto + +lemma left_invertible_transpose: + "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)" + by (metis matrix_transpose_mul transpose_mat transpose_transpose) + +lemma right_invertible_transpose: + "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)" + by (metis matrix_transpose_mul transpose_mat transpose_transpose) + +lemma matrix_left_invertible_injective: + "(\<exists>B. (B::real^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x y. A *v x = A *v y \<longrightarrow> x = y)" +proof - + { fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y" + from xy have "B*v (A *v x) = B *v (A*v y)" by simp + hence "x = y" + unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid . } + moreover + { assume A: "\<forall>x y. A *v x = A *v y \<longrightarrow> x = y" + hence i: "inj (op *v A)" unfolding inj_on_def by auto + from linear_injective_left_inverse[OF matrix_vector_mul_linear i] + obtain g where g: "linear g" "g \<circ> op *v A = id" by blast + have "matrix g ** A = mat 1" + unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] + using g(2) by (simp add: fun_eq_iff) + then have "\<exists>B. (B::real ^'m^'n) ** A = mat 1" by blast } + ultimately show ?thesis by blast +qed + +lemma matrix_left_invertible_ker: + "(\<exists>B. (B::real ^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)" + unfolding matrix_left_invertible_injective + using linear_injective_0[OF matrix_vector_mul_linear, of A] + by (simp add: inj_on_def) + +lemma matrix_right_invertible_surjective: + "(\<exists>B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)" +proof - + { fix B :: "real ^'m^'n" + assume AB: "A ** B = mat 1" + { fix x :: "real ^ 'm" + have "A *v (B *v x) = x" + by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB) } + hence "surj (op *v A)" unfolding surj_def by metis } + moreover + { assume sf: "surj (op *v A)" + from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf] + obtain g:: "real ^'m \<Rightarrow> real ^'n" where g: "linear g" "op *v A \<circ> g = id" + by blast + + have "A ** (matrix g) = mat 1" + unfolding matrix_eq matrix_vector_mul_lid + matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] + using g(2) unfolding o_def fun_eq_iff id_def + . + hence "\<exists>B. A ** (B::real^'m^'n) = mat 1" by blast + } + ultimately show ?thesis unfolding surj_def by blast +qed + +lemma matrix_left_invertible_independent_columns: + fixes A :: "real^'n^'m" + shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow> + (\<forall>c. setsum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))" + (is "?lhs \<longleftrightarrow> ?rhs") +proof - + let ?U = "UNIV :: 'n set" + { assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0" + { fix c i + assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0" and i: "i \<in> ?U" + let ?x = "\<chi> i. c i" + have th0:"A *v ?x = 0" + using c + unfolding matrix_mult_vsum vec_eq_iff + by auto + from k[rule_format, OF th0] i + have "c i = 0" by (vector vec_eq_iff)} + hence ?rhs by blast } + moreover + { assume H: ?rhs + { fix x assume x: "A *v x = 0" + let ?c = "\<lambda>i. ((x$i ):: real)" + from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x] + have "x = 0" by vector } + } + ultimately show ?thesis unfolding matrix_left_invertible_ker by blast +qed + +lemma matrix_right_invertible_independent_rows: + fixes A :: "real^'n^'m" + shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow> + (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))" + unfolding left_invertible_transpose[symmetric] + matrix_left_invertible_independent_columns + by (simp add: column_transpose) + +lemma matrix_right_invertible_span_columns: + "(\<exists>(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow> + span (columns A) = UNIV" (is "?lhs = ?rhs") +proof - + let ?U = "UNIV :: 'm set" + have fU: "finite ?U" by simp + have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)" + unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def + apply (subst eq_commute) + apply rule + done + have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> span (columns A))" by blast + { assume h: ?lhs + { fix x:: "real ^'n" + from h[unfolded lhseq, rule_format, of x] obtain y :: "real ^'m" + where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast + have "x \<in> span (columns A)" + unfolding y[symmetric] + apply (rule span_setsum) + unfolding scalar_mult_eq_scaleR + apply (rule span_mul) + apply (rule span_superset) + unfolding columns_def + apply blast + done + } + then have ?rhs unfolding rhseq by blast } + moreover + { assume h:?rhs + let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y" + { fix y + have "?P y" + proof (rule span_induct_alt[of ?P "columns A", folded scalar_mult_eq_scaleR]) + show "\<exists>x::real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0" + by (rule exI[where x=0], simp) + next + fix c y1 y2 + assume y1: "y1 \<in> columns A" and y2: "?P y2" + from y1 obtain i where i: "i \<in> ?U" "y1 = column i A" + unfolding columns_def by blast + from y2 obtain x:: "real ^'m" where + x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast + let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m" + show "?P (c*s y1 + y2)" + proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left cond_application_beta cong del: if_weak_cong) + fix j + have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" + using i(1) by (simp add: field_simps) + have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U" + apply (rule setsum.cong[OF refl]) + using th apply blast + done + also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" + by (simp add: setsum.distrib) + also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" + unfolding setsum.delta[OF fU] + using i(1) by simp + finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" . + qed + next + show "y \<in> span (columns A)" + unfolding h by blast + qed + } + then have ?lhs unfolding lhseq .. + } + ultimately show ?thesis by blast +qed + +lemma matrix_left_invertible_span_rows: + "(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV" + unfolding right_invertible_transpose[symmetric] + unfolding columns_transpose[symmetric] + unfolding matrix_right_invertible_span_columns + .. + +text \<open>The same result in terms of square matrices.\<close> + +lemma matrix_left_right_inverse: + fixes A A' :: "real ^'n^'n" + shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1" +proof - + { fix A A' :: "real ^'n^'n" + assume AA': "A ** A' = mat 1" + have sA: "surj (op *v A)" + unfolding surj_def + apply clarify + apply (rule_tac x="(A' *v y)" in exI) + apply (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid) + done + from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA] + obtain f' :: "real ^'n \<Rightarrow> real ^'n" + where f': "linear f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast + have th: "matrix f' ** A = mat 1" + by (simp add: matrix_eq matrix_works[OF f'(1)] + matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format]) + hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp + hence "matrix f' = A'" + by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid) + hence "matrix f' ** A = A' ** A" by simp + hence "A' ** A = mat 1" by (simp add: th) + } + then show ?thesis by blast +qed + +text \<open>Considering an n-element vector as an n-by-1 or 1-by-n matrix.\<close> + +definition "rowvector v = (\<chi> i j. (v$j))" + +definition "columnvector v = (\<chi> i j. (v$i))" + +lemma transpose_columnvector: "transpose(columnvector v) = rowvector v" + by (simp add: transpose_def rowvector_def columnvector_def vec_eq_iff) + +lemma transpose_rowvector: "transpose(rowvector v) = columnvector v" + by (simp add: transpose_def columnvector_def rowvector_def vec_eq_iff) + +lemma dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v" + by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def) + +lemma dot_matrix_product: + "(x::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1" + by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vec_def) + +lemma dot_matrix_vector_mul: + fixes A B :: "real ^'n ^'n" and x y :: "real ^'n" + shows "(A *v x) \<bullet> (B *v y) = + (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1" + unfolding dot_matrix_product transpose_columnvector[symmetric] + dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc .. + + +lemma infnorm_cart:"infnorm (x::real^'n) = Sup {\<bar>x$i\<bar> |i. i\<in>UNIV}" + by (simp add: infnorm_def inner_axis Basis_vec_def) (metis (lifting) inner_axis real_inner_1_right) + +lemma component_le_infnorm_cart: "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)" + using Basis_le_infnorm[of "axis i 1" x] + by (simp add: Basis_vec_def axis_eq_axis inner_axis) + +lemma continuous_component[continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x $ i)" + unfolding continuous_def by (rule tendsto_vec_nth) + +lemma continuous_on_component[continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x $ i)" + unfolding continuous_on_def by (fast intro: tendsto_vec_nth) + +lemma continuous_on_vec_lambda[continuous_intros]: + "(\<And>i. continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<chi> i. f i x)" + unfolding continuous_on_def by (auto intro: tendsto_vec_lambda) + +lemma closed_positive_orthant: "closed {x::real^'n. \<forall>i. 0 \<le>x$i}" + by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + +lemma bounded_component_cart: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)" + unfolding bounded_def + apply clarify + apply (rule_tac x="x $ i" in exI) + apply (rule_tac x="e" in exI) + apply clarify + apply (rule order_trans [OF dist_vec_nth_le], simp) + done + +lemma compact_lemma_cart: + fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n" + assumes f: "bounded (range f)" + shows "\<exists>l r. subseq r \<and> + (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)" + (is "?th d") +proof - + have "\<forall>d' \<subseteq> d. ?th d'" + by (rule compact_lemma_general[where unproj=vec_lambda]) + (auto intro!: f bounded_component_cart simp: vec_lambda_eta) + then show "?th d" by simp +qed + +instance vec :: (heine_borel, finite) heine_borel +proof + fix f :: "nat \<Rightarrow> 'a ^ 'b" + assume f: "bounded (range f)" + then obtain l r where r: "subseq r" + and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" + using compact_lemma_cart [OF f] by blast + let ?d = "UNIV::'b set" + { fix e::real assume "e>0" + hence "0 < e / (real_of_nat (card ?d))" + using zero_less_card_finite divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto + with l have "eventually (\<lambda>n. \<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" + by simp + moreover + { fix n + assume n: "\<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" + have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))" + unfolding dist_vec_def using zero_le_dist by (rule setL2_le_setsum) + also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))" + by (rule setsum_strict_mono) (simp_all add: n) + finally have "dist (f (r n)) l < e" by simp + } + ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" + by (rule eventually_mono) + } + hence "((f \<circ> r) \<longlongrightarrow> l) sequentially" unfolding o_def tendsto_iff by simp + with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" by auto +qed + +lemma interval_cart: + fixes a :: "real^'n" + shows "box a b = {x::real^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}" + and "cbox a b = {x::real^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}" + by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def mem_box Basis_vec_def inner_axis) + +lemma mem_interval_cart: + fixes a :: "real^'n" + shows "x \<in> box a b \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)" + and "x \<in> cbox a b \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)" + using interval_cart[of a b] by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def) + +lemma interval_eq_empty_cart: + fixes a :: "real^'n" + shows "(box a b = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1) + and "(cbox a b = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2) +proof - + { fix i x assume as:"b$i \<le> a$i" and x:"x\<in>box a b" + hence "a $ i < x $ i \<and> x $ i < b $ i" unfolding mem_interval_cart by auto + hence "a$i < b$i" by auto + hence False using as by auto } + moreover + { assume as:"\<forall>i. \<not> (b$i \<le> a$i)" + let ?x = "(1/2) *\<^sub>R (a + b)" + { fix i + have "a$i < b$i" using as[THEN spec[where x=i]] by auto + hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i" + unfolding vector_smult_component and vector_add_component + by auto } + hence "box a b \<noteq> {}" using mem_interval_cart(1)[of "?x" a b] by auto } + ultimately show ?th1 by blast + + { fix i x assume as:"b$i < a$i" and x:"x\<in>cbox a b" + hence "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" unfolding mem_interval_cart by auto + hence "a$i \<le> b$i" by auto + hence False using as by auto } + moreover + { assume as:"\<forall>i. \<not> (b$i < a$i)" + let ?x = "(1/2) *\<^sub>R (a + b)" + { fix i + have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto + hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i" + unfolding vector_smult_component and vector_add_component + by auto } + hence "cbox a b \<noteq> {}" using mem_interval_cart(2)[of "?x" a b] by auto } + ultimately show ?th2 by blast +qed + +lemma interval_ne_empty_cart: + fixes a :: "real^'n" + shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)" + and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)" + unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le) + (* BH: Why doesn't just "auto" work here? *) + +lemma subset_interval_imp_cart: + fixes a :: "real^'n" + shows "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> cbox c d \<subseteq> cbox a b" + and "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> cbox c d \<subseteq> box a b" + and "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> box c d \<subseteq> cbox a b" + and "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> box c d \<subseteq> box a b" + unfolding subset_eq[unfolded Ball_def] unfolding mem_interval_cart + by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *) + +lemma interval_sing: + fixes a :: "'a::linorder^'n" + shows "{a .. a} = {a} \<and> {a<..<a} = {}" + apply (auto simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff) + done + +lemma subset_interval_cart: + fixes a :: "real^'n" + shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) + and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2) + and "box c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) + and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4) + using subset_box[of c d a b] by (simp_all add: Basis_vec_def inner_axis) + +lemma disjoint_interval_cart: + fixes a::"real^'n" + shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1) + and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2) + and "box a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3) + and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4) + using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis) + +lemma inter_interval_cart: + fixes a :: "real^'n" + shows "cbox a b \<inter> cbox c d = {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}" + unfolding inter_interval + by (auto simp: mem_box less_eq_vec_def) + (auto simp: Basis_vec_def inner_axis) + +lemma closed_interval_left_cart: + fixes b :: "real^'n" + shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}" + by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + +lemma closed_interval_right_cart: + fixes a::"real^'n" + shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}" + by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + +lemma is_interval_cart: + "is_interval (s::(real^'n) set) \<longleftrightarrow> + (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i))) \<longrightarrow> x \<in> s)" + by (simp add: is_interval_def Ball_def Basis_vec_def inner_axis imp_ex) + +lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \<le> a}" + by (simp add: closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + +lemma closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \<ge> a}" + by (simp add: closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + +lemma open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}" + by (simp add: open_Collect_less continuous_on_const continuous_on_id continuous_on_component) + +lemma open_halfspace_component_gt_cart: "open {x::real^'n. x$i > a}" + by (simp add: open_Collect_less continuous_on_const continuous_on_id continuous_on_component) + +lemma Lim_component_le_cart: + fixes f :: "'a \<Rightarrow> real^'n" + assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f x $i \<le> b) net" + shows "l$i \<le> b" + by (rule tendsto_le[OF assms(2) tendsto_const tendsto_vec_nth, OF assms(1, 3)]) + +lemma Lim_component_ge_cart: + fixes f :: "'a \<Rightarrow> real^'n" + assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. b \<le> (f x)$i) net" + shows "b \<le> l$i" + by (rule tendsto_le[OF assms(2) tendsto_vec_nth tendsto_const, OF assms(1, 3)]) + +lemma Lim_component_eq_cart: + fixes f :: "'a \<Rightarrow> real^'n" + assumes net: "(f \<longlongrightarrow> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)$i = b) net" + shows "l$i = b" + using ev[unfolded order_eq_iff eventually_conj_iff] and + Lim_component_ge_cart[OF net, of b i] and + Lim_component_le_cart[OF net, of i b] by auto + +lemma connected_ivt_component_cart: + fixes x :: "real^'n" + shows "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s. z$k = a)" + using connected_ivt_hyperplane[of s x y "axis k 1" a] + by (auto simp add: inner_axis inner_commute) + +lemma subspace_substandard_cart: "subspace {x::real^_. (\<forall>i. P i \<longrightarrow> x$i = 0)}" + unfolding subspace_def by auto + +lemma closed_substandard_cart: + "closed {x::'a::real_normed_vector ^ 'n. \<forall>i. P i \<longrightarrow> x$i = 0}" +proof - + { fix i::'n + have "closed {x::'a ^ 'n. P i \<longrightarrow> x$i = 0}" + by (cases "P i") (simp_all add: closed_Collect_eq continuous_on_const continuous_on_id continuous_on_component) } + thus ?thesis + unfolding Collect_all_eq by (simp add: closed_INT) +qed + +lemma dim_substandard_cart: "dim {x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d" + (is "dim ?A = _") +proof - + let ?a = "\<lambda>x. axis x 1 :: real^'n" + have *: "{x. \<forall>i\<in>Basis. i \<notin> ?a ` d \<longrightarrow> x \<bullet> i = 0} = ?A" + by (auto simp: image_iff Basis_vec_def axis_eq_axis inner_axis) + have "?a ` d \<subseteq> Basis" + by (auto simp: Basis_vec_def) + thus ?thesis + using dim_substandard[of "?a ` d"] card_image[of ?a d] + by (auto simp: axis_eq_axis inj_on_def *) +qed + +lemma affinity_inverses: + assumes m0: "m \<noteq> (0::'a::field)" + shows "(\<lambda>x. m *s x + c) \<circ> (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id" + "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) \<circ> (\<lambda>x. m *s x + c) = id" + using m0 + apply (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff) + apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1 [symmetric]) + done + +lemma vector_affinity_eq: + assumes m0: "(m::'a::field) \<noteq> 0" + shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)" +proof + assume h: "m *s x + c = y" + hence "m *s x = y - c" by (simp add: field_simps) + hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp + then show "x = inverse m *s y + - (inverse m *s c)" + using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) +next + assume h: "x = inverse m *s y + - (inverse m *s c)" + show "m *s x + c = y" unfolding h + using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) +qed + +lemma vector_eq_affinity: + "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)" + using vector_affinity_eq[where m=m and x=x and y=y and c=c] + by metis + +lemma vector_cart: + fixes f :: "real^'n \<Rightarrow> real" + shows "(\<chi> i. f (axis i 1)) = (\<Sum>i\<in>Basis. f i *\<^sub>R i)" + unfolding euclidean_eq_iff[where 'a="real^'n"] + by simp (simp add: Basis_vec_def inner_axis) + +lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)" + by (rule vector_cart) + +subsection "Convex Euclidean Space" + +lemma Cart_1:"(1::real^'n) = \<Sum>Basis" + using const_vector_cart[of 1] by (simp add: one_vec_def) + +declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] +declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] + +lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component + +lemma convex_box_cart: + assumes "\<And>i. convex {x. P i x}" + shows "convex {x. \<forall>i. P i (x$i)}" + using assms unfolding convex_def by auto + +lemma convex_positive_orthant_cart: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}" + by (rule convex_box_cart) (simp add: atLeast_def[symmetric]) + +lemma unit_interval_convex_hull_cart: + "cbox (0::real^'n) 1 = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}" + unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] box_real[symmetric] + by (rule arg_cong[where f="\<lambda>x. convex hull x"]) (simp add: Basis_vec_def inner_axis) + +lemma cube_convex_hull_cart: + assumes "0 < d" + obtains s::"(real^'n) set" + where "finite s" "cbox (x - (\<chi> i. d)) (x + (\<chi> i. d)) = convex hull s" +proof - + from assms obtain s where "finite s" + and "cbox (x - setsum (op *\<^sub>R d) Basis) (x + setsum (op *\<^sub>R d) Basis) = convex hull s" + by (rule cube_convex_hull) + with that[of s] show thesis + by (simp add: const_vector_cart) +qed + + +subsection "Derivative" + +definition "jacobian f net = matrix(frechet_derivative f net)" + +lemma jacobian_works: + "(f::(real^'a) \<Rightarrow> (real^'b)) differentiable net \<longleftrightarrow> + (f has_derivative (\<lambda>h. (jacobian f net) *v h)) net" + apply rule + unfolding jacobian_def + apply (simp only: matrix_works[OF linear_frechet_derivative]) defer + apply (rule differentiableI) + apply assumption + unfolding frechet_derivative_works + apply assumption + done + + +subsection \<open>Component of the differential must be zero if it exists at a local + maximum or minimum for that corresponding component.\<close> + +lemma differential_zero_maxmin_cart: + fixes f::"real^'a \<Rightarrow> real^'b" + assumes "0 < e" "((\<forall>y \<in> ball x e. (f y)$k \<le> (f x)$k) \<or> (\<forall>y\<in>ball x e. (f x)$k \<le> (f y)$k))" + "f differentiable (at x)" + shows "jacobian f (at x) $ k = 0" + using differential_zero_maxmin_component[of "axis k 1" e x f] assms + vector_cart[of "\<lambda>j. frechet_derivative f (at x) j $ k"] + by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def) + +subsection \<open>Lemmas for working on @{typ "real^1"}\<close> + +lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1" + by (metis (full_types) num1_eq_iff) + +lemma ex_1[simp]: "(\<exists>x::1. P x) \<longleftrightarrow> P 1" + by auto (metis (full_types) num1_eq_iff) + +lemma exhaust_2: + fixes x :: 2 + shows "x = 1 \<or> x = 2" +proof (induct x) + case (of_int z) + then have "0 <= z" and "z < 2" by simp_all + then have "z = 0 | z = 1" by arith + then show ?case by auto +qed + +lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2" + by (metis exhaust_2) + +lemma exhaust_3: + fixes x :: 3 + shows "x = 1 \<or> x = 2 \<or> x = 3" +proof (induct x) + case (of_int z) + then have "0 <= z" and "z < 3" by simp_all + then have "z = 0 \<or> z = 1 \<or> z = 2" by arith + then show ?case by auto +qed + +lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3" + by (metis exhaust_3) + +lemma UNIV_1 [simp]: "UNIV = {1::1}" + by (auto simp add: num1_eq_iff) + +lemma UNIV_2: "UNIV = {1::2, 2::2}" + using exhaust_2 by auto + +lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}" + using exhaust_3 by auto + +lemma setsum_1: "setsum f (UNIV::1 set) = f 1" + unfolding UNIV_1 by simp + +lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2" + unfolding UNIV_2 by simp + +lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3" + unfolding UNIV_3 by (simp add: ac_simps) + +instantiation num1 :: cart_one +begin + +instance +proof + show "CARD(1) = Suc 0" by auto +qed + +end + +subsection\<open>The collapse of the general concepts to dimension one.\<close> + +lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))" + by (simp add: vec_eq_iff) + +lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))" + apply auto + apply (erule_tac x= "x$1" in allE) + apply (simp only: vector_one[symmetric]) + done + +lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)" + by (simp add: norm_vec_def) + +lemma norm_real: "norm(x::real ^ 1) = \<bar>x$1\<bar>" + by (simp add: norm_vector_1) + +lemma dist_real: "dist(x::real ^ 1) y = \<bar>(x$1) - (y$1)\<bar>" + by (auto simp add: norm_real dist_norm) + + +subsection\<open>Explicit vector construction from lists.\<close> + +definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)" + +lemma vector_1: "(vector[x]) $1 = x" + unfolding vector_def by simp + +lemma vector_2: + "(vector[x,y]) $1 = x" + "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)" + unfolding vector_def by simp_all + +lemma vector_3: + "(vector [x,y,z] ::('a::zero)^3)$1 = x" + "(vector [x,y,z] ::('a::zero)^3)$2 = y" + "(vector [x,y,z] ::('a::zero)^3)$3 = z" + unfolding vector_def by simp_all + +lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))" + apply auto + apply (erule_tac x="v$1" in allE) + apply (subgoal_tac "vector [v$1] = v") + apply simp + apply (vector vector_def) + apply simp + done + +lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))" + apply auto + apply (erule_tac x="v$1" in allE) + apply (erule_tac x="v$2" in allE) + apply (subgoal_tac "vector [v$1, v$2] = v") + apply simp + apply (vector vector_def) + apply (simp add: forall_2) + done + +lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))" + apply auto + apply (erule_tac x="v$1" in allE) + apply (erule_tac x="v$2" in allE) + apply (erule_tac x="v$3" in allE) + apply (subgoal_tac "vector [v$1, v$2, v$3] = v") + apply simp + apply (vector vector_def) + apply (simp add: forall_3) + done + +lemma bounded_linear_component_cart[intro]: "bounded_linear (\<lambda>x::real^'n. x $ k)" + apply (rule bounded_linearI[where K=1]) + using component_le_norm_cart[of _ k] unfolding real_norm_def by auto + +lemma integral_component_eq_cart[simp]: + fixes f :: "'n::euclidean_space \<Rightarrow> real^'m" + assumes "f integrable_on s" + shows "integral s (\<lambda>x. f x $ k) = integral s f $ k" + using integral_linear[OF assms(1) bounded_linear_component_cart,unfolded o_def] . + +lemma interval_split_cart: + "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}" + "cbox a b \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}" + apply (rule_tac[!] set_eqI) + unfolding Int_iff mem_interval_cart mem_Collect_eq interval_cbox_cart + unfolding vec_lambda_beta + by auto + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,7539 @@ +section \<open>Complex path integrals and Cauchy's integral theorem\<close> + +text\<open>By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\<close> + +theory Cauchy_Integral_Theorem +imports Complex_Transcendental Weierstrass_Theorems Ordered_Euclidean_Space +begin + +subsection\<open>Homeomorphisms of arc images\<close> + +lemma homeomorphism_arc: + fixes g :: "real \<Rightarrow> 'a::t2_space" + assumes "arc g" + obtains h where "homeomorphism {0..1} (path_image g) g h" +using assms by (force simp add: arc_def homeomorphism_compact path_def path_image_def) + +lemma homeomorphic_arc_image_interval: + fixes g :: "real \<Rightarrow> '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 \<Rightarrow> 'a::t2_space" and h :: "real \<Rightarrow> '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 + +subsection \<open>Piecewise differentiable functions\<close> + +definition piecewise_differentiable_on + (infixr "piecewise'_differentiable'_on" 50) + where "f piecewise_differentiable_on i \<equiv> + continuous_on i f \<and> + (\<exists>s. finite s \<and> (\<forall>x \<in> i - s. f differentiable (at x within i)))" + +lemma piecewise_differentiable_on_imp_continuous_on: + "f piecewise_differentiable_on s \<Longrightarrow> continuous_on s f" +by (simp add: piecewise_differentiable_on_def) + +lemma piecewise_differentiable_on_subset: + "f piecewise_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_differentiable_on t" + using continuous_on_subset + unfolding piecewise_differentiable_on_def + apply safe + apply (blast intro: 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} \<Longrightarrow> 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: + "(\<And>x. x \<in> s \<Longrightarrow> f differentiable (at x within s)) + \<Longrightarrow> 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]: "(\<lambda>x. z) piecewise_differentiable_on s" + by (simp add: differentiable_imp_piecewise_differentiable) + +lemma piecewise_differentiable_compose: + "\<lbrakk>f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s); + \<And>x. finite (s \<inter> f-`{x})\<rbrakk> + \<Longrightarrow> (g o 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 \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI) + apply (blast intro: differentiable_chain_within) + done + +lemma piecewise_differentiable_affine: + fixes m::real + assumes "f piecewise_differentiable_on ((\<lambda>x. m *\<^sub>R x + c) ` s)" + shows "(f o (\<lambda>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 \<le> c" "c \<le> b" "f c = g c" + shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_differentiable_on {a..b}" +proof - + obtain s t where st: "finite s" "finite t" + "\<forall>x\<in>{a..c} - s. f differentiable at x within {a..c}" + "\<forall>x\<in>{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} \<union> (s \<union> t))" + by (metis \<open>finite s\<close> \<open>finite t\<close> 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} (\<lambda>x. if x \<le> 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 "\<lambda>x. x\<le>c"] assms + by (force simp: ivl_disj_un_two_touch) + moreover + { fix x + assume x: "x \<in> {a..b} - ({a,b,c} \<union> (s \<union> t))" + have "(\<lambda>x. if x \<le> 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 + apply (rule differentiable_transform_within [where d = "dist x c" and f = f]) + using x le st + apply (simp_all add: dist_real_def) + apply (rule differentiable_at_withinI) + apply (rule differentiable_within_open [where s = "{a<..<c} - s", THEN iffD1], simp_all) + apply (blast intro: open_greaterThanLessThan finite_imp_closed) + apply (force elim!: differentiable_subset)+ + done + next + case ge show ?diff_fg + apply (rule differentiable_transform_within [where d = "dist x c" and f = g]) + using x ge st + apply (simp_all add: dist_real_def) + apply (rule differentiable_at_withinI) + apply (rule differentiable_within_open [where s = "{c<..<b} - t", THEN iffD1], simp_all) + apply (blast intro: open_greaterThanLessThan finite_imp_closed) + apply (force elim!: differentiable_subset)+ + done + qed + } + then have "\<exists>s. finite s \<and> + (\<forall>x\<in>{a..b} - s. (\<lambda>x. if x \<le> 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 \<Longrightarrow> (\<lambda>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 "(\<lambda>x. f x + g x) piecewise_differentiable_on i" +proof - + obtain s t where st: "finite s" "finite t" + "\<forall>x\<in>i - s. f differentiable at x within i" + "\<forall>x\<in>i - t. g differentiable at x within i" + using assms by (auto simp: piecewise_differentiable_on_def) + then have "finite (s \<union> t) \<and> (\<forall>x\<in>i - (s \<union> t). (\<lambda>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: + "\<lbrakk>f piecewise_differentiable_on s; g piecewise_differentiable_on s\<rbrakk> + \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_differentiable_on s" + unfolding diff_conv_add_uminus + by (metis piecewise_differentiable_add piecewise_differentiable_neg) + +lemma continuous_on_joinpaths_D1: + "continuous_on {0..1} (g1 +++ g2) \<Longrightarrow> continuous_on {0..1} g1" + apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (op*(inverse 2))"]) + apply (rule continuous_intros | simp)+ + apply (auto elim!: continuous_on_subset simp: joinpaths_def) + done + +lemma continuous_on_joinpaths_D2: + "\<lbrakk>continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> continuous_on {0..1} g2" + apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (\<lambda>x. inverse 2*x + 1/2)"]) + apply (rule continuous_intros | simp)+ + apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def) + done + +lemma piecewise_differentiable_D1: + "(g1 +++ g2) piecewise_differentiable_on {0..1} \<Longrightarrow> g1 piecewise_differentiable_on {0..1}" + apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D1) + apply (rule_tac x="insert 1 ((op*2)`s)" in exI) + apply simp + apply (intro ballI) + apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" + in differentiable_transform_within) + apply (auto simp: dist_real_def joinpaths_def) + apply (rule differentiable_chain_within derivative_intros | simp)+ + apply (rule differentiable_subset) + apply (force simp:)+ + done + +lemma piecewise_differentiable_D2: + "\<lbrakk>(g1 +++ g2) piecewise_differentiable_on {0..1}; pathfinish g1 = pathstart g2\<rbrakk> + \<Longrightarrow> g2 piecewise_differentiable_on {0..1}" + apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D2) + apply (rule_tac x="insert 0 ((\<lambda>x. 2*x-1)`s)" in exI) + apply simp + apply (intro ballI) + apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)" + in differentiable_transform_within) + apply (auto simp: dist_real_def joinpaths_def abs_if field_simps split: if_split_asm) + apply (rule differentiable_chain_within derivative_intros | simp)+ + apply (rule differentiable_subset) + apply (force simp: divide_simps)+ + done + + +subsubsection\<open>The concept of continuously differentiable\<close> + +text \<open> +John Harrison writes as follows: + +``The usual assumption in complex analysis texts is that a path \<open>\<gamma>\<close> 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... [namely] continuity plus differentiability except on a +finite set ... [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.\<close> + +definition C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool" + (infix "C1'_differentiable'_on" 50) + where + "f C1_differentiable_on s \<longleftrightarrow> + (\<exists>D. (\<forall>x \<in> s. (f has_vector_derivative (D x)) (at x)) \<and> continuous_on s D)" + +lemma C1_differentiable_on_eq: + "f C1_differentiable_on s \<longleftrightarrow> + (\<forall>x \<in> s. f differentiable at x) \<and> continuous_on s (\<lambda>x. vector_derivative f (at x))" + unfolding C1_differentiable_on_def + apply safe + using differentiable_def has_vector_derivative_def apply blast + apply (erule continuous_on_eq) + using vector_derivative_at apply fastforce + using vector_derivative_works apply fastforce + done + +lemma C1_differentiable_on_subset: + "f C1_differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> 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: + "\<lbrakk>f C1_differentiable_on s; g C1_differentiable_on (f ` s); + \<And>x. finite (s \<inter> f-`{x})\<rbrakk> + \<Longrightarrow> (g o f) C1_differentiable_on s" + apply (simp add: C1_differentiable_on_eq, safe) + using differentiable_chain_at apply blast + apply (rule continuous_on_eq [of _ "\<lambda>x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) + apply (rule Limits.continuous_on_scaleR, assumption) + apply (metis (mono_tags, lifting) continuous_on_eq continuous_at_imp_continuous_on continuous_on_compose differentiable_imp_continuous_within o_def) + by (simp add: vector_derivative_chain_at) + +lemma C1_diff_imp_diff: "f C1_differentiable_on s \<Longrightarrow> f differentiable_on s" + by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) + +lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\<lambda>x. x) C1_differentiable_on s" + by (auto simp: C1_differentiable_on_eq continuous_on_const) + +lemma C1_differentiable_on_const [simp, derivative_intros]: "(\<lambda>z. a) C1_differentiable_on s" + by (auto simp: C1_differentiable_on_eq continuous_on_const) + +lemma C1_differentiable_on_add [simp, derivative_intros]: + "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>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 \<Longrightarrow> (\<lambda>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 \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>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 \<Rightarrow> 'a :: real_normed_algebra" + shows "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>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 \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>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 piecewise_C1_differentiable_on + (infixr "piecewise'_C1'_differentiable'_on" 50) + where "f piecewise_C1_differentiable_on i \<equiv> + continuous_on i f \<and> + (\<exists>s. finite s \<and> (f C1_differentiable_on (i - s)))" + +lemma C1_differentiable_imp_piecewise: + "f C1_differentiable_on s \<Longrightarrow> 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 \<Longrightarrow> 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_within) + +lemma piecewise_C1_differentiable_compose: + "\<lbrakk>f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on (f ` s); + \<And>x. finite (s \<inter> f-`{x})\<rbrakk> + \<Longrightarrow> (g o f) piecewise_C1_differentiable_on s" + apply (simp add: piecewise_C1_differentiable_on_def, safe) + apply (blast intro: continuous_on_compose2) + apply (rename_tac A B) + apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI) + apply (rule conjI, blast) + apply (rule C1_differentiable_compose) + apply (blast intro: C1_differentiable_on_subset) + apply (blast intro: C1_differentiable_on_subset) + by (simp add: Diff_Int_distrib2) + +lemma piecewise_C1_differentiable_on_subset: + "f piecewise_C1_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> 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 \<Longrightarrow> 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 ((\<lambda>x. m * x + c) ` s)" + shows "(f o (\<lambda>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 continuous_on_const) +next + case False + show ?thesis + apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) + apply (rule assms derivative_intros | simp add: False vimage_def)+ + using real_vector_affinity_eq [OF False, where c=c, unfolded scaleR_conv_of_real] + apply simp + 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 \<le> c" "c \<le> b" "f c = g c" + shows "(\<lambda>x. if x \<le> 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..<c} - s" + and g_diff: "g differentiable_on {c<..b} - t" + by (simp_all add: C1_differentiable_on_eq differentiable_at_withinI differentiable_on_def) + have "continuous_on {a..c} f" "continuous_on {c..b} g" + using assms piecewise_C1_differentiable_on_def by auto + then have cab: "continuous_on {a..b} (\<lambda>x. if x \<le> 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 "\<lambda>x. x\<le>c"] assms + by (force simp: ivl_disj_un_two_touch) + { fix x + assume x: "x \<in> {a..b} - insert c (s \<union> t)" + have "(\<lambda>x. if x \<le> 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 "(\<forall>x \<in> {a..b} - insert c (s \<union> t). (\<lambda>x. if x \<le> c then f x else g x) differentiable at x)" + by auto + moreover + { assume fcon: "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative f (at x))" + and gcon: "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative g (at x))" + have "open ({a<..<c} - s)" "open ({c<..<b} - t)" + using st by (simp_all add: open_Diff finite_imp_closed) + moreover have "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))" + apply (rule continuous_on_eq [OF fcon]) + apply (simp add:) + apply (rule vector_derivative_at [symmetric]) + apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) + apply (simp_all add: dist_norm vector_derivative_works [symmetric]) + apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff linorder_not_le order_less_irrefl st(1)) + apply auto + done + moreover have "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))" + apply (rule continuous_on_eq [OF gcon]) + apply (simp add:) + apply (rule vector_derivative_at [symmetric]) + apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) + apply (simp_all add: dist_norm vector_derivative_works [symmetric]) + apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff less_irrefl not_le st(2)) + apply auto + done + ultimately have "continuous_on ({a<..<b} - insert c (s \<union> t)) + (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))" + apply (rule continuous_on_subset [OF continuous_on_open_Un], auto) + done + } note * = this + have "continuous_on ({a<..<b} - insert c (s \<union> t)) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> 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 "\<exists>s. finite s \<and> ((\<lambda>x. if x \<le> c then f x else g x) C1_differentiable_on {a..b} - s)" + apply (rule_tac x="{a,b,c} \<union> s \<union> 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 \<Longrightarrow> (\<lambda>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 "(\<lambda>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 \<union> t) \<and> (\<lambda>x. f x + g x) C1_differentiable_on i - (s \<union> 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: + "\<lbrakk>f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on s\<rbrakk> + \<Longrightarrow> (\<lambda>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) + +lemma piecewise_C1_differentiable_D1: + fixes g1 :: "real \<Rightarrow> 'a::real_normed_field" + assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" + shows "g1 piecewise_C1_differentiable_on {0..1}" +proof - + obtain s where "finite s" + and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))" + and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x" + using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + then have g1D: "g1 differentiable at x" if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x + apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" in differentiable_transform_within) + using that + apply (simp_all add: dist_real_def joinpaths_def) + apply (rule differentiable_chain_at derivative_intros | force)+ + done + have [simp]: "vector_derivative (g1 \<circ> op * 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" + if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x + apply (subst vector_derivative_chain_at) + using that + apply (rule derivative_eq_intros g1D | simp)+ + done + have "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))" + using co12 by (rule continuous_on_subset) force + then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 o op*2) (at x))" + apply (rule continuous_on_eq [OF _ vector_derivative_at]) + apply (rule_tac f="g1 o op*2" and d="dist x (1/2)" in has_vector_derivative_transform_within) + apply (simp_all add: dist_norm joinpaths_def vector_derivative_works [symmetric]) + apply (force intro: g1D differentiable_chain_at) + apply auto + done + have "continuous_on ({0..1} - insert 1 (op * 2 ` s)) + ((\<lambda>x. 1/2 * vector_derivative (g1 o op*2) (at x)) o op*(1/2))" + apply (rule continuous_intros)+ + using coDhalf + apply (simp add: scaleR_conv_of_real image_set_diff image_image) + done + then have con_g1: "continuous_on ({0..1} - insert 1 (op * 2 ` s)) (\<lambda>x. vector_derivative g1 (at x))" + by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) + have "continuous_on {0..1} g1" + using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast + with \<open>finite s\<close> show ?thesis + apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + apply (rule_tac x="insert 1 ((op*2)`s)" in exI) + apply (simp add: g1D con_g1) + done +qed + +lemma piecewise_C1_differentiable_D2: + fixes g2 :: "real \<Rightarrow> 'a::real_normed_field" + assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" + shows "g2 piecewise_C1_differentiable_on {0..1}" +proof - + obtain s where "finite s" + and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))" + and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x" + using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + then have g2D: "g2 differentiable at x" if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x + apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)" in differentiable_transform_within) + using that + apply (simp_all add: dist_real_def joinpaths_def) + apply (auto simp: dist_real_def joinpaths_def field_simps) + apply (rule differentiable_chain_at derivative_intros | force)+ + apply (drule_tac x= "(x + 1) / 2" in bspec, force simp: divide_simps) + apply assumption + done + have [simp]: "vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" + if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x + using that by (auto simp: vector_derivative_chain_at divide_simps g2D) + have "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))" + using co12 by (rule continuous_on_subset) force + then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g2 o (\<lambda>x. 2*x-1)) (at x))" + apply (rule continuous_on_eq [OF _ vector_derivative_at]) + apply (rule_tac f="g2 o (\<lambda>x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) + apply (auto simp: dist_real_def field_simps joinpaths_def vector_derivative_works [symmetric] + intro!: g2D differentiable_chain_at) + done + have [simp]: "((\<lambda>x. (x + 1) / 2) ` ({0..1} - insert 0 ((\<lambda>x. 2 * x - 1) ` s))) = ({1/2..1} - insert (1/2) s)" + apply (simp add: image_set_diff inj_on_def image_image) + apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) + done + have "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)) + ((\<lambda>x. 1/2 * vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at x)) o (\<lambda>x. (x+1)/2))" + by (rule continuous_intros | simp add: coDhalf)+ + then have con_g2: "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)) (\<lambda>x. vector_derivative g2 (at x))" + by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) + have "continuous_on {0..1} g2" + using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast + with \<open>finite s\<close> show ?thesis + apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + apply (rule_tac x="insert 0 ((\<lambda>x. 2 * x - 1) ` s)" in exI) + apply (simp add: g2D con_g2) + done +qed + +subsection \<open>Valid paths, and their start and finish\<close> + +lemma Diff_Un_eq: "A - (B \<union> C) = A - B - C" + by blast + +definition valid_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool" + where "valid_path f \<equiv> f piecewise_C1_differentiable_on {0..1::real}" + +definition closed_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool" + where "closed_path g \<equiv> g 0 = g 1" + +subsubsection\<open>In particular, all results for paths apply\<close> + +lemma valid_path_imp_path: "valid_path g \<Longrightarrow> path g" +by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) + +lemma connected_valid_path_image: "valid_path g \<Longrightarrow> connected(path_image g)" + by (metis connected_path_image valid_path_imp_path) + +lemma compact_valid_path_image: "valid_path g \<Longrightarrow> compact(path_image g)" + by (metis compact_path_image valid_path_imp_path) + +lemma bounded_valid_path_image: "valid_path g \<Longrightarrow> bounded(path_image g)" + by (metis bounded_path_image valid_path_imp_path) + +lemma closed_valid_path_image: "valid_path g \<Longrightarrow> closed(path_image g)" + by (metis closed_path_image valid_path_imp_path) + +proposition valid_path_compose: + assumes "valid_path g" + and der: "\<And>x. x \<in> path_image g \<Longrightarrow> \<exists>f'. (f has_field_derivative f') (at x)" + and con: "continuous_on (path_image g) (deriv f)" + shows "valid_path (f o g)" +proof - + obtain s where "finite s" and g_diff: "g C1_differentiable_on {0..1} - s" + using \<open>valid_path g\<close> unfolding valid_path_def piecewise_C1_differentiable_on_def by auto + have "f \<circ> g differentiable at t" when "t\<in>{0..1} - s" for t + proof (rule differentiable_chain_at) + show "g differentiable at t" using \<open>valid_path g\<close> + by (meson C1_differentiable_on_eq \<open>g C1_differentiable_on {0..1} - s\<close> that) + next + have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis + then obtain f' where "(f has_field_derivative f') (at (g t))" + using der by auto + then have " (f has_derivative op * f') (at (g t))" + using has_field_derivative_imp_has_derivative[of f f' "at (g t)"] by auto + then show "f differentiable at (g t)" using differentiableI by auto + qed + moreover have "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (f \<circ> g) (at x))" + proof (rule continuous_on_eq [where f = "\<lambda>x. vector_derivative g (at x) * deriv f (g x)"], + rule continuous_intros) + show "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative g (at x))" + using g_diff C1_differentiable_on_eq by auto + next + have "continuous_on {0..1} (\<lambda>x. deriv f (g x))" + using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] + \<open>valid_path g\<close> piecewise_C1_differentiable_on_def valid_path_def + by blast + then show "continuous_on ({0..1} - s) (\<lambda>x. deriv f (g x))" + using continuous_on_subset by blast + next + show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \<circ> g) (at t)" + when "t \<in> {0..1} - s" for t + proof (rule vector_derivative_chain_at_general[symmetric]) + show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) + next + have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis + then obtain f' where "(f has_field_derivative f') (at (g t))" + using der by auto + then show "\<exists>g'. (f has_field_derivative g') (at (g t))" by auto + qed + qed + ultimately have "f o g C1_differentiable_on {0..1} - s" + using C1_differentiable_on_eq by blast + moreover have "path (f o g)" + proof - + have "isCont f x" when "x\<in>path_image g" for x + proof - + obtain f' where "(f has_field_derivative f') (at x)" + using der[rule_format] \<open>x\<in>path_image g\<close> by auto + thus ?thesis using DERIV_isCont by auto + qed + then have "continuous_on (path_image g) f" using continuous_at_imp_continuous_on by auto + then show ?thesis using path_continuous_image \<open>valid_path g\<close> valid_path_imp_path by auto + qed + ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def + using \<open>finite s\<close> by auto +qed + + +subsection\<open>Contour Integrals along a path\<close> + +text\<open>This definition is for complex numbers only, and does not generalise to line integrals in a vector field\<close> + +text\<open>piecewise differentiable function on [0,1]\<close> + +definition has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool" + (infixr "has'_contour'_integral" 50) + where "(f has_contour_integral i) g \<equiv> + ((\<lambda>x. f(g x) * vector_derivative g (at x within {0..1})) + has_integral i) {0..1}" + +definition contour_integrable_on + (infixr "contour'_integrable'_on" 50) + where "f contour_integrable_on g \<equiv> \<exists>i. (f has_contour_integral i) g" + +definition contour_integral + where "contour_integral g f \<equiv> @i. (f has_contour_integral i) g \<or> ~ f contour_integrable_on g \<and> i=0" + +lemma not_integrable_contour_integral: "~ f contour_integrable_on g \<Longrightarrow> contour_integral g f = 0" + unfolding contour_integrable_on_def contour_integral_def by blast + +lemma contour_integral_unique: "(f has_contour_integral i) g \<Longrightarrow> contour_integral g f = i" + apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) + using has_integral_unique by blast + +corollary has_contour_integral_eqpath: + "\<lbrakk>(f has_contour_integral y) p; f contour_integrable_on \<gamma>; + contour_integral p f = contour_integral \<gamma> f\<rbrakk> + \<Longrightarrow> (f has_contour_integral y) \<gamma>" +using contour_integrable_on_def contour_integral_unique by auto + +lemma has_contour_integral_integral: + "f contour_integrable_on i \<Longrightarrow> (f has_contour_integral (contour_integral i f)) i" + by (metis contour_integral_unique contour_integrable_on_def) + +lemma has_contour_integral_unique: + "(f has_contour_integral i) g \<Longrightarrow> (f has_contour_integral j) g \<Longrightarrow> i = j" + using has_integral_unique + by (auto simp: has_contour_integral_def) + +lemma has_contour_integral_integrable: "(f has_contour_integral i) g \<Longrightarrow> f contour_integrable_on g" + using contour_integrable_on_def by blast + +(* Show that we can forget about the localized derivative.*) + +lemma vector_derivative_within_interior: + "\<lbrakk>x \<in> interior s; NO_MATCH UNIV s\<rbrakk> + \<Longrightarrow> vector_derivative f (at x within s) = vector_derivative f (at x)" + apply (simp add: vector_derivative_def has_vector_derivative_def has_derivative_def netlimit_within_interior) + apply (subst lim_within_interior, auto) + done + +lemma has_integral_localized_vector_derivative: + "((\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \<longleftrightarrow> + ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" +proof - + have "{a..b} - {a,b} = interior {a..b}" + by (simp add: atLeastAtMost_diff_ends) + show ?thesis + apply (rule has_integral_spike_eq [of "{a,b}"]) + apply (auto simp: vector_derivative_within_interior) + done +qed + +lemma integrable_on_localized_vector_derivative: + "(\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \<longleftrightarrow> + (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" + by (simp add: integrable_on_def has_integral_localized_vector_derivative) + +lemma has_contour_integral: + "(f has_contour_integral i) g \<longleftrightarrow> + ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" + by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) + +lemma contour_integrable_on: + "f contour_integrable_on g \<longleftrightarrow> + (\<lambda>t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" + by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) + +subsection\<open>Reversing a path\<close> + +lemma valid_path_imp_reverse: + assumes "valid_path g" + shows "valid_path(reversepath g)" +proof - + obtain s where "finite s" "g C1_differentiable_on ({0..1} - s)" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + then have "finite (op - 1 ` s)" "(reversepath g C1_differentiable_on ({0..1} - op - 1 ` s))" + apply (auto simp: reversepath_def) + apply (rule C1_differentiable_compose [of "\<lambda>x::real. 1-x" _ g, unfolded o_def]) + apply (auto simp: C1_differentiable_on_eq) + apply (rule continuous_intros, force) + apply (force elim!: continuous_on_subset) + apply (simp add: finite_vimageI inj_on_def) + done + then show ?thesis using assms + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) +qed + +lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \<longleftrightarrow> valid_path g" + using valid_path_imp_reverse by force + +lemma has_contour_integral_reversepath: + assumes "valid_path g" "(f has_contour_integral i) g" + shows "(f has_contour_integral (-i)) (reversepath g)" +proof - + { fix s x + assume xs: "g C1_differentiable_on ({0..1} - s)" "x \<notin> op - 1 ` s" "0 \<le> x" "x \<le> 1" + have "vector_derivative (\<lambda>x. g (1 - x)) (at x within {0..1}) = + - vector_derivative g (at (1 - x) within {0..1})" + proof - + obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" + using xs + by (force simp: has_vector_derivative_def C1_differentiable_on_def) + have "(g o (\<lambda>x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" + apply (rule vector_diff_chain_within) + apply (intro vector_diff_chain_within derivative_eq_intros | simp)+ + apply (rule has_vector_derivative_at_within [OF f']) + done + then have mf': "((\<lambda>x. g (1 - x)) has_vector_derivative -f') (at x)" + by (simp add: o_def) + show ?thesis + using xs + by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) + qed + } note * = this + have 01: "{0..1::real} = cbox 0 1" + by simp + show ?thesis using assms + apply (auto simp: has_contour_integral_def) + apply (drule has_integral_affinity01 [where m= "-1" and c=1]) + apply (auto simp: reversepath_def valid_path_def piecewise_C1_differentiable_on_def) + apply (drule has_integral_neg) + apply (rule_tac s = "(\<lambda>x. 1 - x) ` s" in has_integral_spike_finite) + apply (auto simp: *) + done +qed + +lemma contour_integrable_reversepath: + "valid_path g \<Longrightarrow> f contour_integrable_on g \<Longrightarrow> f contour_integrable_on (reversepath g)" + using has_contour_integral_reversepath contour_integrable_on_def by blast + +lemma contour_integrable_reversepath_eq: + "valid_path g \<Longrightarrow> (f contour_integrable_on (reversepath g) \<longleftrightarrow> f contour_integrable_on g)" + using contour_integrable_reversepath valid_path_reversepath by fastforce + +lemma contour_integral_reversepath: + assumes "valid_path g" + shows "contour_integral (reversepath g) f = - (contour_integral g f)" +proof (cases "f contour_integrable_on g") + case True then show ?thesis + by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) +next + case False then have "~ f contour_integrable_on (reversepath g)" + by (simp add: assms contour_integrable_reversepath_eq) + with False show ?thesis by (simp add: not_integrable_contour_integral) +qed + + +subsection\<open>Joining two paths together\<close> + +lemma valid_path_join: + assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" + shows "valid_path(g1 +++ g2)" +proof - + have "g1 1 = g2 0" + using assms by (auto simp: pathfinish_def pathstart_def) + moreover have "(g1 o (\<lambda>x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" + apply (rule piecewise_C1_differentiable_compose) + using assms + apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) + apply (rule continuous_intros | simp)+ + apply (force intro: finite_vimageI [where h = "op*2"] inj_onI) + done + moreover have "(g2 o (\<lambda>x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" + apply (rule piecewise_C1_differentiable_compose) + using assms unfolding valid_path_def piecewise_C1_differentiable_on_def + by (auto intro!: continuous_intros finite_vimageI [where h = "(\<lambda>x. 2*x - 1)"] inj_onI + simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) + ultimately show ?thesis + apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) + apply (rule piecewise_C1_differentiable_cases) + apply (auto simp: o_def) + done +qed + +lemma valid_path_join_D1: + fixes g1 :: "real \<Rightarrow> 'a::real_normed_field" + shows "valid_path (g1 +++ g2) \<Longrightarrow> valid_path g1" + unfolding valid_path_def + by (rule piecewise_C1_differentiable_D1) + +lemma valid_path_join_D2: + fixes g2 :: "real \<Rightarrow> 'a::real_normed_field" + shows "\<lbrakk>valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> valid_path g2" + unfolding valid_path_def + by (rule piecewise_C1_differentiable_D2) + +lemma valid_path_join_eq [simp]: + fixes g2 :: "real \<Rightarrow> 'a::real_normed_field" + shows "pathfinish g1 = pathstart g2 \<Longrightarrow> (valid_path(g1 +++ g2) \<longleftrightarrow> valid_path g1 \<and> valid_path g2)" + using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast + +lemma has_contour_integral_join: + assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" + "valid_path g1" "valid_path g2" + shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" +proof - + obtain s1 s2 + where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x" + and s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x" + using assms + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have 1: "((\<lambda>x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" + and 2: "((\<lambda>x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" + using assms + by (auto simp: has_contour_integral) + have i1: "((\<lambda>x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" + and i2: "((\<lambda>x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" + using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] + has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] + by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) + have g1: "\<lbrakk>0 \<le> z; z*2 < 1; z*2 \<notin> s1\<rbrakk> \<Longrightarrow> + vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = + 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>z - 1/2\<bar>"]]) + apply (simp_all add: dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x" 2 _ g1, simplified o_def]) + apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + using s1 + apply (auto simp: algebra_simps vector_derivative_works) + done + have g2: "\<lbrakk>1 < z*2; z \<le> 1; z*2 - 1 \<notin> s2\<rbrakk> \<Longrightarrow> + vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = + 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2 (2*x - 1))" and d = "\<bar>z - 1/2\<bar>"]]) + apply (simp_all add: dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x - 1" 2 _ g2, simplified o_def]) + apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + using s2 + apply (auto simp: algebra_simps vector_derivative_works) + done + have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" + apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) (op*2 -` s1)"]) + using s1 + apply (force intro: finite_vimageI [where h = "op*2"] inj_onI) + apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) + done + moreover have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" + apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\<lambda>x. 2*x-1) -` s2)"]) + using s2 + apply (force intro: finite_vimageI [where h = "\<lambda>x. 2*x-1"] inj_onI) + apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) + done + ultimately + show ?thesis + apply (simp add: has_contour_integral) + apply (rule has_integral_combine [where c = "1/2"], auto) + done +qed + +lemma contour_integrable_joinI: + assumes "f contour_integrable_on g1" "f contour_integrable_on g2" + "valid_path g1" "valid_path g2" + shows "f contour_integrable_on (g1 +++ g2)" + using assms + by (meson has_contour_integral_join contour_integrable_on_def) + +lemma contour_integrable_joinD1: + assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" + shows "f contour_integrable_on g1" +proof - + obtain s1 + where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have "(\<lambda>x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" + using assms + apply (auto simp: contour_integrable_on) + apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) + apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) + done + then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" + by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) + have g1: "\<lbrakk>0 < z; z < 1; z \<notin> s1\<rbrakk> \<Longrightarrow> + vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = + 2 *\<^sub>R vector_derivative g1 (at z)" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>(z-1)/2\<bar>"]]) + apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\<lambda>x. x*2" 2 _ g1, simplified o_def]) + using s1 + apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + done + show ?thesis + using s1 + apply (auto simp: contour_integrable_on) + apply (rule integrable_spike_finite [of "{0,1} \<union> s1", OF _ _ *]) + apply (auto simp: joinpaths_def scaleR_conv_of_real g1) + done +qed + +lemma contour_integrable_joinD2: + assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" + shows "f contour_integrable_on g2" +proof - + obtain s2 + where s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have "(\<lambda>x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" + using assms + apply (auto simp: contour_integrable_on) + apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) + apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) + apply (simp add: image_affinity_atLeastAtMost_diff) + done + then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) + integrable_on {0..1}" + by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) + have g2: "\<lbrakk>0 < z; z < 1; z \<notin> s2\<rbrakk> \<Longrightarrow> + vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = + 2 *\<^sub>R vector_derivative g2 (at z)" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2(2*x-1))" and d = "\<bar>z/2\<bar>"]]) + apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\<lambda>x. x*2-1" 2 _ g2, simplified o_def]) + using s2 + apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left + vector_derivative_works add_divide_distrib) + done + show ?thesis + using s2 + apply (auto simp: contour_integrable_on) + apply (rule integrable_spike_finite [of "{0,1} \<union> s2", OF _ _ *]) + apply (auto simp: joinpaths_def scaleR_conv_of_real g2) + done +qed + +lemma contour_integrable_join [simp]: + shows + "\<lbrakk>valid_path g1; valid_path g2\<rbrakk> + \<Longrightarrow> f contour_integrable_on (g1 +++ g2) \<longleftrightarrow> f contour_integrable_on g1 \<and> f contour_integrable_on g2" +using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast + +lemma contour_integral_join [simp]: + shows + "\<lbrakk>f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\<rbrakk> + \<Longrightarrow> contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" + by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) + + +subsection\<open>Shifting the starting point of a (closed) path\<close> + +lemma shiftpath_alt_def: "shiftpath a f = (\<lambda>x. if x \<le> 1-a then f (a + x) else f (a + x - 1))" + by (auto simp: shiftpath_def) + +lemma valid_path_shiftpath [intro]: + assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}" + shows "valid_path(shiftpath a g)" + using assms + apply (auto simp: valid_path_def shiftpath_alt_def) + apply (rule piecewise_C1_differentiable_cases) + apply (auto simp: algebra_simps) + apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) + apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) + apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + done + +lemma has_contour_integral_shiftpath: + assumes f: "(f has_contour_integral i) g" "valid_path g" + and a: "a \<in> {0..1}" + shows "(f has_contour_integral i) (shiftpath a g)" +proof - + obtain s + where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have *: "((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" + using assms by (auto simp: has_contour_integral) + then have i: "i = integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)) + + integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x))" + apply (rule has_integral_unique) + apply (subst add.commute) + apply (subst integral_combine) + using assms * integral_unique by auto + { fix x + have "0 \<le> x \<Longrightarrow> x + a < 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a) ` s \<Longrightarrow> + vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" + unfolding shiftpath_def + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x))" and d = "dist(1-a) x"]]) + apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\<lambda>x. x+a" 1 _ g, simplified o_def scaleR_one]) + apply (intro derivative_eq_intros | simp)+ + using g + apply (drule_tac x="x+a" in bspec) + using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) + done + } note vd1 = this + { fix x + have "1 < x + a \<Longrightarrow> x \<le> 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a + 1) ` s \<Longrightarrow> + vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" + unfolding shiftpath_def + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x-1))" and d = "dist (1-a) x"]]) + apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\<lambda>x. x+a-1" 1 _ g, simplified o_def scaleR_one]) + apply (intro derivative_eq_intros | simp)+ + using g + apply (drule_tac x="x+a-1" in bspec) + using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) + done + } note vd2 = this + have va1: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" + using * a by (fastforce intro: integrable_subinterval_real) + have v0a: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" + apply (rule integrable_subinterval_real) + using * a by auto + have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) + has_integral integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x))) {0..1 - a}" + apply (rule has_integral_spike_finite + [where s = "{1-a} \<union> (\<lambda>x. x-a) ` s" and f = "\<lambda>x. f(g(a+x)) * vector_derivative g (at(a+x))"]) + using s apply blast + using a apply (auto simp: algebra_simps vd1) + apply (force simp: shiftpath_def add.commute) + using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] + apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) + done + moreover + have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) + has_integral integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x))) {1 - a..1}" + apply (rule has_integral_spike_finite + [where s = "{1-a} \<union> (\<lambda>x. x-a+1) ` s" and f = "\<lambda>x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) + using s apply blast + using a apply (auto simp: algebra_simps vd2) + apply (force simp: shiftpath_def add.commute) + using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] + apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) + apply (simp add: algebra_simps) + done + ultimately show ?thesis + using a + by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) +qed + +lemma has_contour_integral_shiftpath_D: + assumes "(f has_contour_integral i) (shiftpath a g)" + "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}" + shows "(f has_contour_integral i) g" +proof - + obtain s + where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + { fix x + assume x: "0 < x" "x < 1" "x \<notin> s" + then have gx: "g differentiable at x" + using g by auto + have "vector_derivative g (at x within {0..1}) = + vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" + apply (rule vector_derivative_at_within_ivl + [OF has_vector_derivative_transform_within_open + [where f = "(shiftpath (1 - a) (shiftpath a g))" and s = "{0<..<1}-s"]]) + using s g assms x + apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath + vector_derivative_within_interior vector_derivative_works [symmetric]) + apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) + apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) + done + } note vd = this + have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" + using assms by (auto intro!: has_contour_integral_shiftpath) + show ?thesis + apply (simp add: has_contour_integral_def) + apply (rule has_integral_spike_finite [of "{0,1} \<union> s", OF _ _ fi [unfolded has_contour_integral_def]]) + using s assms vd + apply (auto simp: Path_Connected.shiftpath_shiftpath) + done +qed + +lemma has_contour_integral_shiftpath_eq: + assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}" + shows "(f has_contour_integral i) (shiftpath a g) \<longleftrightarrow> (f has_contour_integral i) g" + using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast + +lemma contour_integrable_on_shiftpath_eq: + assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}" + shows "f contour_integrable_on (shiftpath a g) \<longleftrightarrow> f contour_integrable_on g" +using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto + +lemma contour_integral_shiftpath: + assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}" + shows "contour_integral (shiftpath a g) f = contour_integral g f" + using assms + by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) + + +subsection\<open>More about straight-line paths\<close> + +lemma has_vector_derivative_linepath_within: + "(linepath a b has_vector_derivative (b - a)) (at x within s)" +apply (simp add: linepath_def has_vector_derivative_def algebra_simps) +apply (rule derivative_eq_intros | simp)+ +done + +lemma vector_derivative_linepath_within: + "x \<in> {0..1} \<Longrightarrow> vector_derivative (linepath a b) (at x within {0..1}) = b - a" + apply (rule vector_derivative_within_closed_interval [of 0 "1::real", simplified]) + apply (auto simp: has_vector_derivative_linepath_within) + done + +lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" + by (simp add: has_vector_derivative_linepath_within vector_derivative_at) + +lemma valid_path_linepath [iff]: "valid_path (linepath a b)" + apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) + apply (rule_tac x="{}" in exI) + apply (simp add: differentiable_on_def differentiable_def) + using has_vector_derivative_def has_vector_derivative_linepath_within + apply (fastforce simp add: continuous_on_eq_continuous_within) + done + +lemma has_contour_integral_linepath: + shows "(f has_contour_integral i) (linepath a b) \<longleftrightarrow> + ((\<lambda>x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" + by (simp add: has_contour_integral vector_derivative_linepath_at) + +lemma linepath_in_path: + shows "x \<in> {0..1} \<Longrightarrow> linepath a b x \<in> closed_segment a b" + by (auto simp: segment linepath_def) + +lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" + by (auto simp: segment linepath_def) + +lemma linepath_in_convex_hull: + fixes x::real + assumes a: "a \<in> convex hull s" + and b: "b \<in> convex hull s" + and x: "0\<le>x" "x\<le>1" + shows "linepath a b x \<in> convex hull s" + apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) + using x + apply (auto simp: linepath_image_01 [symmetric]) + done + +lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" + by (simp add: linepath_def) + +lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" + by (simp add: linepath_def) + +lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" + by (simp add: has_contour_integral_linepath) + +lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" + using has_contour_integral_trivial contour_integral_unique by blast + + +subsection\<open>Relation to subpath construction\<close> + +lemma valid_path_subpath: + fixes g :: "real \<Rightarrow> 'a :: real_normed_vector" + assumes "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" + shows "valid_path(subpath u v g)" +proof (cases "v=u") + case True + then show ?thesis + unfolding valid_path_def subpath_def + by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) +next + case False + have "(g o (\<lambda>x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" + apply (rule piecewise_C1_differentiable_compose) + apply (simp add: C1_differentiable_imp_piecewise) + apply (simp add: image_affinity_atLeastAtMost) + using assms False + apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) + apply (subst Int_commute) + apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) + done + then show ?thesis + by (auto simp: o_def valid_path_def subpath_def) +qed + +lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" + by (simp add: has_contour_integral subpath_def) + +lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" + using has_contour_integral_subpath_refl contour_integrable_on_def by blast + +lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" + by (simp add: has_contour_integral_subpath_refl contour_integral_unique) + +lemma has_contour_integral_subpath: + assumes f: "f contour_integrable_on g" and g: "valid_path g" + and uv: "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v" + shows "(f has_contour_integral integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x))) + (subpath u v g)" +proof (cases "v=u") + case True + then show ?thesis + using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) +next + case False + obtain s where s: "\<And>x. x \<in> {0..1} - s \<Longrightarrow> g differentiable at x" and fs: "finite s" + using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast + have *: "((\<lambda>x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) + has_integral (1 / (v - u)) * integral {u..v} (\<lambda>t. f (g t) * vector_derivative g (at t))) + {0..1}" + using f uv + apply (simp add: contour_integrable_on subpath_def has_contour_integral) + apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) + apply (simp_all add: has_integral_integral) + apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) + apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) + apply (simp add: divide_simps False) + done + { fix x + have "x \<in> {0..1} \<Longrightarrow> + x \<notin> (\<lambda>t. (v-u) *\<^sub>R t + u) -` s \<Longrightarrow> + vector_derivative (\<lambda>x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" + apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) + apply (intro derivative_eq_intros | simp)+ + apply (cut_tac s [of "(v - u) * x + u"]) + using uv mult_left_le [of x "v-u"] + apply (auto simp: vector_derivative_works) + done + } note vd = this + show ?thesis + apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) + using fs assms + apply (simp add: False subpath_def has_contour_integral) + apply (rule_tac s = "(\<lambda>t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) + apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) + done +qed + +lemma contour_integrable_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" + shows "f contour_integrable_on (subpath u v g)" + apply (cases u v rule: linorder_class.le_cases) + apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) + apply (subst reversepath_subpath [symmetric]) + apply (rule contour_integrable_reversepath) + using assms apply (blast intro: valid_path_subpath) + apply (simp add: contour_integrable_on_def) + using assms apply (blast intro: has_contour_integral_subpath) + done + +lemma has_integral_integrable_integral: "(f has_integral i) s \<longleftrightarrow> f integrable_on s \<and> integral s f = i" + by blast + +lemma has_integral_contour_integral_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v" + shows "(((\<lambda>x. f(g x) * vector_derivative g (at x))) + has_integral contour_integral (subpath u v g) f) {u..v}" + using assms + apply (auto simp: has_integral_integrable_integral) + apply (rule integrable_on_subcbox [where a=u and b=v and s = "{0..1}", simplified]) + apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) + done + +lemma contour_integral_subcontour_integral: + assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v" + shows "contour_integral (subpath u v g) f = + integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x))" + using assms has_contour_integral_subpath contour_integral_unique by blast + +lemma contour_integral_subpath_combine_less: + assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}" + "u<v" "v<w" + shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = + contour_integral (subpath u w g) f" + using assms apply (auto simp: contour_integral_subcontour_integral) + apply (rule integral_combine, auto) + apply (rule integrable_on_subcbox [where a=u and b=w and s = "{0..1}", simplified]) + apply (auto simp: contour_integrable_on) + done + +lemma contour_integral_subpath_combine: + assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}" + shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = + contour_integral (subpath u w g) f" +proof (cases "u\<noteq>v \<and> v\<noteq>w \<and> u\<noteq>w") + case True + have *: "subpath v u g = reversepath(subpath u v g) \<and> + subpath w u g = reversepath(subpath u w g) \<and> + subpath w v g = reversepath(subpath v w g)" + by (auto simp: reversepath_subpath) + have "u < v \<and> v < w \<or> + u < w \<and> w < v \<or> + v < u \<and> u < w \<or> + v < w \<and> w < u \<or> + w < u \<and> u < v \<or> + w < v \<and> v < u" + using True assms by linarith + with assms show ?thesis + using contour_integral_subpath_combine_less [of f g u v w] + contour_integral_subpath_combine_less [of f g u w v] + contour_integral_subpath_combine_less [of f g v u w] + contour_integral_subpath_combine_less [of f g v w u] + contour_integral_subpath_combine_less [of f g w u v] + contour_integral_subpath_combine_less [of f g w v u] + apply simp + apply (elim disjE) + apply (auto simp: * contour_integral_reversepath contour_integrable_subpath + valid_path_reversepath valid_path_subpath algebra_simps) + done +next + case False + then show ?thesis + apply (auto simp: contour_integral_subpath_refl) + using assms + by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath) +qed + +lemma contour_integral_integral: + "contour_integral g f = integral {0..1} (\<lambda>x. f (g x) * vector_derivative g (at x))" + by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) + + +text\<open>Cauchy's theorem where there's a primitive\<close> + +lemma contour_integral_primitive_lemma: + fixes f :: "complex \<Rightarrow> complex" and g :: "real \<Rightarrow> complex" + assumes "a \<le> b" + and "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)" + and "g piecewise_differentiable_on {a..b}" "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s" + shows "((\<lambda>x. f'(g x) * vector_derivative g (at x within {a..b})) + has_integral (f(g b) - f(g a))) {a..b}" +proof - + obtain k where k: "finite k" "\<forall>x\<in>{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" + using assms by (auto simp: piecewise_differentiable_on_def) + have cfg: "continuous_on {a..b} (\<lambda>x. f (g x))" + apply (rule continuous_on_compose [OF cg, unfolded o_def]) + using assms + apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) + done + { fix x::real + assume a: "a < x" and b: "x < b" and xk: "x \<notin> k" + then have "g differentiable at x within {a..b}" + using k by (simp add: differentiable_at_withinI) + then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" + by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) + then have gdiff: "(g has_derivative (\<lambda>u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" + by (simp add: has_vector_derivative_def scaleR_conv_of_real) + have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" + using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) + then have fdiff: "(f has_derivative op * (f' (g x))) (at (g x) within g ` {a..b})" + by (simp add: has_field_derivative_def) + have "((\<lambda>x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" + using diff_chain_within [OF gdiff fdiff] + by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) + } note * = this + show ?thesis + apply (rule fundamental_theorem_of_calculus_interior_strong) + using k assms cfg * + apply (auto simp: at_within_closed_interval) + done +qed + +lemma contour_integral_primitive: + assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \<subseteq> s" + shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" + using assms + apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) + apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) + done + +corollary Cauchy_theorem_primitive: + assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" + shows "(f' has_contour_integral 0) g" + using assms + by (metis diff_self contour_integral_primitive) + + +text\<open>Existence of path integral for continuous function\<close> +lemma contour_integrable_continuous_linepath: + assumes "continuous_on (closed_segment a b) f" + shows "f contour_integrable_on (linepath a b)" +proof - + have "continuous_on {0..1} ((\<lambda>x. f x * (b - a)) o linepath a b)" + apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) + apply (rule continuous_intros | simp add: assms)+ + done + then show ?thesis + apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) + apply (rule integrable_continuous [of 0 "1::real", simplified]) + apply (rule continuous_on_eq [where f = "\<lambda>x. f(linepath a b x)*(b - a)"]) + apply (auto simp: vector_derivative_linepath_within) + done +qed + +lemma has_field_der_id: "((\<lambda>x. x\<^sup>2 / 2) has_field_derivative x) (at x)" + by (rule has_derivative_imp_has_field_derivative) + (rule derivative_intros | simp)+ + +lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\<lambda>y. y) = (b^2 - a^2)/2" + apply (rule contour_integral_unique) + using contour_integral_primitive [of UNIV "\<lambda>x. x^2/2" "\<lambda>x. x" "linepath a b"] + apply (auto simp: field_simps has_field_der_id) + done + +lemma contour_integrable_on_const [iff]: "(\<lambda>x. c) contour_integrable_on (linepath a b)" + by (simp add: continuous_on_const contour_integrable_continuous_linepath) + +lemma contour_integrable_on_id [iff]: "(\<lambda>x. x) contour_integrable_on (linepath a b)" + by (simp add: continuous_on_id contour_integrable_continuous_linepath) + + +subsection\<open>Arithmetical combining theorems\<close> + +lemma has_contour_integral_neg: + "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. -(f x)) has_contour_integral (-i)) g" + by (simp add: has_integral_neg has_contour_integral_def) + +lemma has_contour_integral_add: + "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk> + \<Longrightarrow> ((\<lambda>x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" + by (simp add: has_integral_add has_contour_integral_def algebra_simps) + +lemma has_contour_integral_diff: + "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk> + \<Longrightarrow> ((\<lambda>x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" + by (simp add: has_integral_sub has_contour_integral_def algebra_simps) + +lemma has_contour_integral_lmul: + "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. c * (f x)) has_contour_integral (c*i)) g" +apply (simp add: has_contour_integral_def) +apply (drule has_integral_mult_right) +apply (simp add: algebra_simps) +done + +lemma has_contour_integral_rmul: + "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. (f x) * c) has_contour_integral (i*c)) g" +apply (drule has_contour_integral_lmul) +apply (simp add: mult.commute) +done + +lemma has_contour_integral_div: + "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. f x/c) has_contour_integral (i/c)) g" + by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) + +lemma has_contour_integral_eq: + "\<lbrakk>(f has_contour_integral y) p; \<And>x. x \<in> path_image p \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> (g has_contour_integral y) p" +apply (simp add: path_image_def has_contour_integral_def) +by (metis (no_types, lifting) image_eqI has_integral_eq) + +lemma has_contour_integral_bound_linepath: + assumes "(f has_contour_integral i) (linepath a b)" + "0 \<le> B" "\<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B" + shows "norm i \<le> B * norm(b - a)" +proof - + { fix x::real + assume x: "0 \<le> x" "x \<le> 1" + have "norm (f (linepath a b x)) * + norm (vector_derivative (linepath a b) (at x within {0..1})) \<le> B * norm (b - a)" + by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) + } note * = this + have "norm i \<le> (B * norm (b - a)) * content (cbox 0 (1::real))" + apply (rule has_integral_bound + [of _ "\<lambda>x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) + using assms * unfolding has_contour_integral_def + apply (auto simp: norm_mult) + done + then show ?thesis + by (auto simp: content_real) +qed + +(*UNUSED +lemma has_contour_integral_bound_linepath_strong: + fixes a :: real and f :: "complex \<Rightarrow> real" + assumes "(f has_contour_integral i) (linepath a b)" + "finite k" + "0 \<le> B" "\<And>x::real. x \<in> closed_segment a b - k \<Longrightarrow> norm(f x) \<le> B" + shows "norm i \<le> B*norm(b - a)" +*) + +lemma has_contour_integral_const_linepath: "((\<lambda>x. c) has_contour_integral c*(b - a))(linepath a b)" + unfolding has_contour_integral_linepath + by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) + +lemma has_contour_integral_0: "((\<lambda>x. 0) has_contour_integral 0) g" + by (simp add: has_contour_integral_def) + +lemma has_contour_integral_is_0: + "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> (f has_contour_integral 0) g" + by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto + +lemma has_contour_integral_setsum: + "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a has_contour_integral i a) p\<rbrakk> + \<Longrightarrow> ((\<lambda>x. setsum (\<lambda>a. f a x) s) has_contour_integral setsum i s) p" + by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) + + +subsection \<open>Operations on path integrals\<close> + +lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\<lambda>x. c) = c*(b - a)" + by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) + +lemma contour_integral_neg: + "f contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. -(f x)) = -(contour_integral g f)" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) + +lemma contour_integral_add: + "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x + f2 x) = + contour_integral g f1 + contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) + +lemma contour_integral_diff: + "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x - f2 x) = + contour_integral g f1 - contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) + +lemma contour_integral_lmul: + shows "f contour_integrable_on g + \<Longrightarrow> contour_integral g (\<lambda>x. c * f x) = c*contour_integral g f" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) + +lemma contour_integral_rmul: + shows "f contour_integrable_on g + \<Longrightarrow> contour_integral g (\<lambda>x. f x * c) = contour_integral g f * c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) + +lemma contour_integral_div: + shows "f contour_integrable_on g + \<Longrightarrow> contour_integral g (\<lambda>x. f x / c) = contour_integral g f / c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) + +lemma contour_integral_eq: + "(\<And>x. x \<in> path_image p \<Longrightarrow> f x = g x) \<Longrightarrow> contour_integral p f = contour_integral p g" + apply (simp add: contour_integral_def) + using has_contour_integral_eq + by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) + +lemma contour_integral_eq_0: + "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> contour_integral g f = 0" + by (simp add: has_contour_integral_is_0 contour_integral_unique) + +lemma contour_integral_bound_linepath: + shows + "\<lbrakk>f contour_integrable_on (linepath a b); + 0 \<le> B; \<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B\<rbrakk> + \<Longrightarrow> norm(contour_integral (linepath a b) f) \<le> B*norm(b - a)" + apply (rule has_contour_integral_bound_linepath [of f]) + apply (auto simp: has_contour_integral_integral) + done + +lemma contour_integral_0 [simp]: "contour_integral g (\<lambda>x. 0) = 0" + by (simp add: contour_integral_unique has_contour_integral_0) + +lemma contour_integral_setsum: + "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk> + \<Longrightarrow> contour_integral p (\<lambda>x. setsum (\<lambda>a. f a x) s) = setsum (\<lambda>a. contour_integral p (f a)) s" + by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral) + +lemma contour_integrable_eq: + "\<lbrakk>f contour_integrable_on p; \<And>x. x \<in> path_image p \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_eq) + + +subsection \<open>Arithmetic theorems for path integrability\<close> + +lemma contour_integrable_neg: + "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. -(f x)) contour_integrable_on g" + using has_contour_integral_neg contour_integrable_on_def by blast + +lemma contour_integrable_add: + "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x + f2 x) contour_integrable_on g" + using has_contour_integral_add contour_integrable_on_def + by fastforce + +lemma contour_integrable_diff: + "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x - f2 x) contour_integrable_on g" + using has_contour_integral_diff contour_integrable_on_def + by fastforce + +lemma contour_integrable_lmul: + "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. c * f x) contour_integrable_on g" + using has_contour_integral_lmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_rmul: + "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x * c) contour_integrable_on g" + using has_contour_integral_rmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_div: + "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x / c) contour_integrable_on g" + using has_contour_integral_div contour_integrable_on_def + by fastforce + +lemma contour_integrable_setsum: + "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk> + \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) s) contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_setsum) + + +subsection\<open>Reversing a path integral\<close> + +lemma has_contour_integral_reverse_linepath: + "(f has_contour_integral i) (linepath a b) + \<Longrightarrow> (f has_contour_integral (-i)) (linepath b a)" + using has_contour_integral_reversepath valid_path_linepath by fastforce + +lemma contour_integral_reverse_linepath: + "continuous_on (closed_segment a b) f + \<Longrightarrow> contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" +apply (rule contour_integral_unique) +apply (rule has_contour_integral_reverse_linepath) +by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) + + +(* Splitting a path integral in a flat way.*) + +lemma has_contour_integral_split: + assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" + and k: "0 \<le> k" "k \<le> 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "(f has_contour_integral (i + j)) (linepath a b)" +proof (cases "k = 0 \<or> k = 1") + case True + then show ?thesis + using assms + apply auto + apply (metis add.left_neutral has_contour_integral_trivial has_contour_integral_unique) + apply (metis add.right_neutral has_contour_integral_trivial has_contour_integral_unique) + done +next + case False + then have k: "0 < k" "k < 1" "complex_of_real k \<noteq> 1" + using assms apply auto + using of_real_eq_iff by fastforce + have c': "c = k *\<^sub>R (b - a) + a" + by (metis diff_add_cancel c) + have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" + by (simp add: algebra_simps c') + { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" + have **: "\<And>x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" + using False + apply (simp add: c' algebra_simps) + apply (simp add: real_vector.scale_left_distrib [symmetric] divide_simps) + done + have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" + using * k + apply - + apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse k" "0", simplified]) + apply (simp_all add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) + apply (drule has_integral_cmul [where c = "inverse k"]) + apply (simp add: has_integral_cmul) + done + } note fi = this + { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" + have **: "\<And>x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" + using k + apply (simp add: c' field_simps) + apply (simp add: scaleR_conv_of_real divide_simps) + apply (simp add: field_simps) + done + have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" + using * k + apply - + apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse(1 - k)" "-(k/(1 - k))", simplified]) + apply (simp_all add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) + apply (drule has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) + apply (simp add: has_integral_cmul) + done + } note fj = this + show ?thesis + using f k + apply (simp add: has_contour_integral_linepath) + apply (simp add: linepath_def) + apply (rule has_integral_combine [OF _ _ fi fj], simp_all) + done +qed + +lemma continuous_on_closed_segment_transform: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \<le> k" "k \<le> 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "continuous_on (closed_segment a c) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + show "continuous_on (closed_segment a c) f" + apply (rule continuous_on_subset [OF f]) + apply (simp add: segment_convex_hull) + apply (rule convex_hull_subset) + using assms + apply (auto simp: hull_inc c' Convex.convexD_alt) + done +qed + +lemma contour_integral_split: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \<le> k" "k \<le> 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" + apply (rule_tac [!] continuous_on_subset [OF f]) + apply (simp_all add: segment_convex_hull) + apply (rule_tac [!] convex_hull_subset) + using assms + apply (auto simp: hull_inc c' Convex.convexD_alt) + done + show ?thesis + apply (rule contour_integral_unique) + apply (rule has_contour_integral_split [OF has_contour_integral_integral has_contour_integral_integral k c]) + apply (rule contour_integrable_continuous_linepath *)+ + done +qed + +lemma contour_integral_split_linepath: + assumes f: "continuous_on (closed_segment a b) f" + and c: "c \<in> closed_segment a b" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" + using c + by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) + +(* The special case of midpoints used in the main quadrisection.*) + +lemma has_contour_integral_midpoint: + assumes "(f has_contour_integral i) (linepath a (midpoint a b))" + "(f has_contour_integral j) (linepath (midpoint a b) b)" + shows "(f has_contour_integral (i + j)) (linepath a b)" + apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) + using assms + apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) + done + +lemma contour_integral_midpoint: + "continuous_on (closed_segment a b) f + \<Longrightarrow> contour_integral (linepath a b) f = + contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" + apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) + apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) + done + + +text\<open>A couple of special case lemmas that are useful below\<close> + +lemma triangle_linear_has_chain_integral: + "((\<lambda>x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" + apply (rule Cauchy_theorem_primitive [of UNIV "\<lambda>x. m/2 * x^2 + d*x"]) + apply (auto intro!: derivative_eq_intros) + done + +lemma has_chain_integral_chain_integral3: + "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) + \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" + apply (subst contour_integral_unique [symmetric], assumption) + apply (drule has_contour_integral_integrable) + apply (simp add: valid_path_join) + done + +lemma has_chain_integral_chain_integral4: + "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) + \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" + apply (subst contour_integral_unique [symmetric], assumption) + apply (drule has_contour_integral_integrable) + apply (simp add: valid_path_join) + done + +subsection\<open>Reversing the order in a double path integral\<close> + +text\<open>The condition is stronger than needed but it's often true in typical situations\<close> + +lemma fst_im_cbox [simp]: "cbox c d \<noteq> {} \<Longrightarrow> (fst ` cbox (a,c) (b,d)) = cbox a b" + by (auto simp: cbox_Pair_eq) + +lemma snd_im_cbox [simp]: "cbox a b \<noteq> {} \<Longrightarrow> (snd ` cbox (a,c) (b,d)) = cbox c d" + by (auto simp: cbox_Pair_eq) + +lemma contour_integral_swap: + assumes fcon: "continuous_on (path_image g \<times> path_image h) (\<lambda>(y1,y2). f y1 y2)" + and vp: "valid_path g" "valid_path h" + and gvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative g (at t))" + and hvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative h (at t))" + shows "contour_integral g (\<lambda>w. contour_integral h (f w)) = + contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))" +proof - + have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + have fgh1: "\<And>x. (\<lambda>t. f (g x) (h t)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g x, h t))" + by (rule ext) simp + have fgh2: "\<And>x. (\<lambda>t. f (g t) (h x)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g t, h x))" + by (rule ext) simp + have fcon_im1: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> continuous_on ((\<lambda>t. (g x, h t)) ` {0..1}) (\<lambda>(x, y). f x y)" + by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) + have fcon_im2: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> continuous_on ((\<lambda>t. (g t, h x)) ` {0..1}) (\<lambda>(x, y). f x y)" + by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) + have vdg: "\<And>y. y \<in> {0..1} \<Longrightarrow> (\<lambda>x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF _ gvcon]) + apply (subst fgh2) + apply (rule fcon_im2 gcon continuous_intros | simp)+ + done + have "(\<lambda>z. vector_derivative g (at (fst z))) = (\<lambda>x. vector_derivative g (at x)) o fst" + by auto + then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\<lambda>x. vector_derivative g (at (fst x)))" + apply (rule ssubst) + apply (rule continuous_intros | simp add: gvcon)+ + done + have "(\<lambda>z. vector_derivative h (at (snd z))) = (\<lambda>x. vector_derivative h (at x)) o snd" + by auto + then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\<lambda>x. vector_derivative h (at (snd x)))" + apply (rule ssubst) + apply (rule continuous_intros | simp add: hvcon)+ + done + have "(\<lambda>x. f (g (fst x)) (h (snd x))) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>w. ((g o fst) w, (h o snd) w))" + by auto + then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\<lambda>x. f (g (fst x)) (h (snd x)))" + apply (rule ssubst) + apply (rule gcon hcon continuous_intros | simp)+ + apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) + done + have "integral {0..1} (\<lambda>x. contour_integral h (f (g x)) * vector_derivative g (at x)) = + integral {0..1} (\<lambda>x. contour_integral h (\<lambda>y. f (g x) y * vector_derivative g (at x)))" + apply (rule integral_cong [OF contour_integral_rmul [symmetric]]) + apply (clarsimp simp: contour_integrable_on) + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF _ hvcon]) + apply (subst fgh1) + apply (rule fcon_im1 hcon continuous_intros | simp)+ + done + also have "... = integral {0..1} + (\<lambda>y. contour_integral g (\<lambda>x. f x (h y) * vector_derivative h (at y)))" + apply (simp only: contour_integral_integral) + apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) + apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ + unfolding integral_mult_left [symmetric] + apply (simp only: mult_ac) + done + also have "... = contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))" + apply (simp add: contour_integral_integral) + apply (rule integral_cong) + unfolding integral_mult_left [symmetric] + apply (simp add: algebra_simps) + done + finally show ?thesis + by (simp add: contour_integral_integral) +qed + + +subsection\<open>The key quadrisection step\<close> + +lemma norm_sum_half: + assumes "norm(a + b) >= e" + shows "norm a >= e/2 \<or> norm b >= e/2" +proof - + have "e \<le> norm (- a - b)" + by (simp add: add.commute assms norm_minus_commute) + thus ?thesis + using norm_triangle_ineq4 order_trans by fastforce +qed + +lemma norm_sum_lemma: + assumes "e \<le> norm (a + b + c + d)" + shows "e / 4 \<le> norm a \<or> e / 4 \<le> norm b \<or> e / 4 \<le> norm c \<or> e / 4 \<le> norm d" +proof - + have "e \<le> norm ((a + b) + (c + d))" using assms + by (simp add: algebra_simps) + then show ?thesis + by (auto dest!: norm_sum_half) +qed + +lemma Cauchy_theorem_quadrisection: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and dist: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K" + and e: "e * K^2 \<le> + norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" + shows "\<exists>a' b' c'. + a' \<in> convex hull {a,b,c} \<and> b' \<in> convex hull {a,b,c} \<and> c' \<in> convex hull {a,b,c} \<and> + dist a' b' \<le> K/2 \<and> dist b' c' \<le> K/2 \<and> dist c' a' \<le> K/2 \<and> + e * (K/2)^2 \<le> norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" +proof - + note divide_le_eq_numeral1 [simp del] + define a' where "a' = midpoint b c" + define b' where "b' = midpoint c a" + define c' where "c' = midpoint a b" + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + have fcont': "continuous_on (closed_segment c' b') f" + "continuous_on (closed_segment a' c') f" + "continuous_on (closed_segment b' a') f" + unfolding a'_def b'_def c'_def + apply (rule continuous_on_subset [OF f], + metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ + done + let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f" + have *: "?pathint a b + ?pathint b c + ?pathint c a = + (?pathint a c' + ?pathint c' b' + ?pathint b' a) + + (?pathint a' c' + ?pathint c' b + ?pathint b a') + + (?pathint a' c + ?pathint c b' + ?pathint b' a') + + (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" + apply (simp add: fcont' contour_integral_reverse_linepath) + apply (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) + done + have [simp]: "\<And>x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" + by (metis left_diff_distrib mult.commute norm_mult_numeral1) + have [simp]: "\<And>x y. cmod (x - y) = cmod (y - x)" + by (simp add: norm_minus_commute) + consider "e * K\<^sup>2 / 4 \<le> cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | + "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | + "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | + "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" + using assms + apply (simp only: *) + apply (blast intro: that dest!: norm_sum_lemma) + done + then show ?thesis + proof cases + case 1 then show ?thesis + apply (rule_tac x=a in exI) + apply (rule exI [where x=c']) + apply (rule exI [where x=b']) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + next + case 2 then show ?thesis + apply (rule_tac x=a' in exI) + apply (rule exI [where x=c']) + apply (rule exI [where x=b]) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + next + case 3 then show ?thesis + apply (rule_tac x=a' in exI) + apply (rule exI [where x=c]) + apply (rule exI [where x=b']) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + next + case 4 then show ?thesis + apply (rule_tac x=a' in exI) + apply (rule exI [where x=b']) + apply (rule exI [where x=c']) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + qed +qed + +subsection\<open>Cauchy's theorem for triangles\<close> + +lemma triangle_points_closer: + fixes a::complex + shows "\<lbrakk>x \<in> convex hull {a,b,c}; y \<in> convex hull {a,b,c}\<rbrakk> + \<Longrightarrow> norm(x - y) \<le> norm(a - b) \<or> + norm(x - y) \<le> norm(b - c) \<or> + norm(x - y) \<le> norm(c - a)" + using simplex_extremal_le [of "{a,b,c}"] + by (auto simp: norm_minus_commute) + +lemma holomorphic_point_small_triangle: + assumes x: "x \<in> s" + and f: "continuous_on s f" + and cd: "f field_differentiable (at x within s)" + and e: "0 < e" + shows "\<exists>k>0. \<forall>a b c. dist a b \<le> k \<and> dist b c \<le> k \<and> dist c a \<le> k \<and> + x \<in> convex hull {a,b,c} \<and> convex hull {a,b,c} \<subseteq> s + \<longrightarrow> norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + + contour_integral(linepath c a) f) + \<le> e*(dist a b + dist b c + dist c a)^2" + (is "\<exists>k>0. \<forall>a b c. _ \<longrightarrow> ?normle a b c") +proof - + have le_of_3: "\<And>a x y z. \<lbrakk>0 \<le> x*y; 0 \<le> x*z; 0 \<le> y*z; a \<le> (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\<rbrakk> + \<Longrightarrow> a \<le> e*(x + y + z)^2" + by (simp add: algebra_simps power2_eq_square) + have disj_le: "\<lbrakk>x \<le> a \<or> x \<le> b \<or> x \<le> c; 0 \<le> a; 0 \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> x \<le> a + b + c" + for x::real and a b c + by linarith + have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" + if "convex hull {a, b, c} \<subseteq> s" for a b c + using segments_subset_convex_hull that + by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ + note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] + { fix f' a b c d + assume d: "0 < d" + and f': "\<And>y. \<lbrakk>cmod (y - x) \<le> d; y \<in> s\<rbrakk> \<Longrightarrow> cmod (f y - f x - f' * (y - x)) \<le> e * cmod (y - x)" + and le: "cmod (a - b) \<le> d" "cmod (b - c) \<le> d" "cmod (c - a) \<le> d" + and xc: "x \<in> convex hull {a, b, c}" + and s: "convex hull {a, b, c} \<subseteq> s" + have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = + contour_integral (linepath a b) (\<lambda>y. f y - f x - f'*(y - x)) + + contour_integral (linepath b c) (\<lambda>y. f y - f x - f'*(y - x)) + + contour_integral (linepath c a) (\<lambda>y. f y - f x - f'*(y - x))" + apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF s]) + apply (simp add: field_simps) + done + { fix y + assume yc: "y \<in> convex hull {a,b,c}" + have "cmod (f y - f x - f' * (y - x)) \<le> e*norm(y - x)" + apply (rule f') + apply (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) + using s yc by blast + also have "... \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" + by (simp add: yc e xc disj_le [OF triangle_points_closer]) + finally have "cmod (f y - f x - f' * (y - x)) \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . + } note cm_le = this + have "?normle a b c" + apply (simp add: dist_norm pa) + apply (rule le_of_3) + using f' xc s e + apply simp_all + apply (intro norm_triangle_le add_mono path_bound) + apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) + apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ + done + } note * = this + show ?thesis + using cd e + apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) + apply (clarify dest!: spec mp) + using * + apply (simp add: dist_norm, blast) + done +qed + + +(* Hence the most basic theorem for a triangle.*) +locale Chain = + fixes x0 At Follows + assumes At0: "At x0 0" + and AtSuc: "\<And>x n. At x n \<Longrightarrow> \<exists>x'. At x' (Suc n) \<and> Follows x' x" +begin + primrec f where + "f 0 = x0" + | "f (Suc n) = (SOME x. At x (Suc n) \<and> Follows x (f n))" + + lemma At: "At (f n) n" + proof (induct n) + case 0 show ?case + by (simp add: At0) + next + case (Suc n) show ?case + by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) + qed + + lemma Follows: "Follows (f(Suc n)) (f n)" + by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) + + declare f.simps(2) [simp del] +end + +lemma Chain3: + assumes At0: "At x0 y0 z0 0" + and AtSuc: "\<And>x y z n. At x y z n \<Longrightarrow> \<exists>x' y' z'. At x' y' z' (Suc n) \<and> Follows x' y' z' x y z" + obtains f g h where + "f 0 = x0" "g 0 = y0" "h 0 = z0" + "\<And>n. At (f n) (g n) (h n) n" + "\<And>n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" +proof - + interpret three: Chain "(x0,y0,z0)" "\<lambda>(x,y,z). At x y z" "\<lambda>(x',y',z'). \<lambda>(x,y,z). Follows x' y' z' x y z" + apply unfold_locales + using At0 AtSuc by auto + show ?thesis + apply (rule that [of "\<lambda>n. fst (three.f n)" "\<lambda>n. fst (snd (three.f n))" "\<lambda>n. snd (snd (three.f n))"]) + apply simp_all + using three.At three.Follows + apply (simp_all add: split_beta') + done +qed + +lemma Cauchy_theorem_triangle: + assumes "f holomorphic_on (convex hull {a,b,c})" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +proof - + have contf: "continuous_on (convex hull {a,b,c}) f" + by (metis assms holomorphic_on_imp_continuous_on) + let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f" + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \<noteq> 0" + define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" + define e where "e = norm y / K^2" + have K1: "K \<ge> 1" by (simp add: K_def max.coboundedI1) + then have K: "K > 0" by linarith + have [iff]: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K" + by (simp_all add: K_def) + have e: "e > 0" + unfolding e_def using ynz K1 by simp + define At where "At x y z n \<longleftrightarrow> + convex hull {x,y,z} \<subseteq> convex hull {a,b,c} \<and> + dist x y \<le> K/2^n \<and> dist y z \<le> K/2^n \<and> dist z x \<le> K/2^n \<and> + norm(?pathint x y + ?pathint y z + ?pathint z x) \<ge> e*(K/2^n)^2" + for x y z n + have At0: "At a b c 0" + using fy + by (simp add: At_def e_def has_chain_integral_chain_integral3) + { fix x y z n + assume At: "At x y z n" + then have contf': "continuous_on (convex hull {x,y,z}) f" + using contf At_def continuous_on_subset by blast + have "\<exists>x' y' z'. At x' y' z' (Suc n) \<and> convex hull {x',y',z'} \<subseteq> convex hull {x,y,z}" + using At + apply (simp add: At_def) + using Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] + apply clarsimp + apply (rule_tac x="a'" in exI) + apply (rule_tac x="b'" in exI) + apply (rule_tac x="c'" in exI) + apply (simp add: algebra_simps) + apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) + done + } note AtSuc = this + obtain fa fb fc + where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" + and cosb: "\<And>n. convex hull {fa n, fb n, fc n} \<subseteq> convex hull {a,b,c}" + and dist: "\<And>n. dist (fa n) (fb n) \<le> K/2^n" + "\<And>n. dist (fb n) (fc n) \<le> K/2^n" + "\<And>n. dist (fc n) (fa n) \<le> K/2^n" + and no: "\<And>n. norm(?pathint (fa n) (fb n) + + ?pathint (fb n) (fc n) + + ?pathint (fc n) (fa n)) \<ge> e * (K/2^n)^2" + and conv_le: "\<And>n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \<subseteq> convex hull {fa n, fb n, fc n}" + apply (rule Chain3 [of At, OF At0 AtSuc]) + apply (auto simp: At_def) + done + have "\<exists>x. \<forall>n. x \<in> convex hull {fa n, fb n, fc n}" + apply (rule bounded_closed_nest) + apply (simp_all add: compact_imp_closed finite_imp_compact_convex_hull finite_imp_bounded_convex_hull) + apply (rule allI) + apply (rule transitive_stepwise_le) + apply (auto simp: conv_le) + done + then obtain x where x: "\<And>n. x \<in> convex hull {fa n, fb n, fc n}" by auto + then have xin: "x \<in> convex hull {a,b,c}" + using assms f0 by blast + then have fx: "f field_differentiable at x within (convex hull {a,b,c})" + using assms holomorphic_on_def by blast + { fix k n + assume k: "0 < k" + and le: + "\<And>x' y' z'. + \<lbrakk>dist x' y' \<le> k; dist y' z' \<le> k; dist z' x' \<le> k; + x \<in> convex hull {x',y',z'}; + convex hull {x',y',z'} \<subseteq> convex hull {a,b,c}\<rbrakk> + \<Longrightarrow> + cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 + \<le> e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" + and Kk: "K / k < 2 ^ n" + have "K / 2 ^ n < k" using Kk k + by (auto simp: field_simps) + then have DD: "dist (fa n) (fb n) \<le> k" "dist (fb n) (fc n) \<le> k" "dist (fc n) (fa n) \<le> k" + using dist [of n] k + by linarith+ + have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 + \<le> (3 * K / 2 ^ n)\<^sup>2" + using dist [of n] e K + by (simp add: abs_le_square_iff [symmetric]) + have less10: "\<And>x y::real. 0 < x \<Longrightarrow> y \<le> 9*x \<Longrightarrow> y < x*10" + by linarith + have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \<le> e * (3 * K / 2 ^ n)\<^sup>2" + using ynz dle e mult_le_cancel_left_pos by blast + also have "... < + cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" + using no [of n] e K + apply (simp add: e_def field_simps) + apply (simp only: zero_less_norm_iff [symmetric]) + done + finally have False + using le [OF DD x cosb] by auto + } then + have ?thesis + using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e + apply clarsimp + apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]]) + apply force+ + done + } + moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" + by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) + segments_subset_convex_hull(3) segments_subset_convex_hull(5)) + ultimately show ?thesis + using has_contour_integral_integral by fastforce +qed + + +subsection\<open>Version needing function holomorphic in interior only\<close> + +lemma Cauchy_theorem_flat_lemma: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and c: "c - a = k *\<^sub>R (b - a)" + and k: "0 \<le> k" + shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof - + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + show ?thesis + proof (cases "k \<le> 1") + case True show ?thesis + by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) + next + case False then show ?thesis + using fabc c + apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) + apply (metis closed_segment_commute fabc(3)) + apply (auto simp: k contour_integral_reverse_linepath) + done + qed +qed + +lemma Cauchy_theorem_flat: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and c: "c - a = k *\<^sub>R (b - a)" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "0 \<le> k") + case True with assms show ?thesis + by (blast intro: Cauchy_theorem_flat_lemma) +next + case False + have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + + contour_integral (linepath c b) f = 0" + apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) + using False c + apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) + done + ultimately show ?thesis + apply (auto simp: contour_integral_reverse_linepath) + using add_eq_0_iff by force +qed + + +lemma Cauchy_theorem_triangle_interior: + assumes contf: "continuous_on (convex hull {a,b,c}) f" + and holf: "f holomorphic_on interior (convex hull {a,b,c})" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +proof - + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using contf continuous_on_subset segments_subset_convex_hull by metis+ + have "bounded (f ` (convex hull {a,b,c}))" + by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) + then obtain B where "0 < B" and Bnf: "\<And>x. x \<in> convex hull {a,b,c} \<Longrightarrow> norm (f x) \<le> B" + by (auto simp: dest!: bounded_pos [THEN iffD1]) + have "bounded (convex hull {a,b,c})" + by (simp add: bounded_convex_hull) + then obtain C where C: "0 < C" and Cno: "\<And>y. y \<in> convex hull {a,b,c} \<Longrightarrow> norm y < C" + using bounded_pos_less by blast + then have diff_2C: "norm(x - y) \<le> 2*C" + if x: "x \<in> convex hull {a, b, c}" and y: "y \<in> convex hull {a, b, c}" for x y + proof - + have "cmod x \<le> C" + using x by (meson Cno not_le not_less_iff_gr_or_eq) + hence "cmod (x - y) \<le> C + C" + using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) + thus "cmod (x - y) \<le> 2 * C" + by (metis mult_2) + qed + have contf': "continuous_on (convex hull {b,a,c}) f" + using contf by (simp add: insert_commute) + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \<noteq> 0" + have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" + by (rule has_chain_integral_chain_integral3 [OF fy]) + have ?thesis + proof (cases "c=a \<or> a=b \<or> b=c") + case True then show ?thesis + using Cauchy_theorem_flat [OF contf, of 0] + using has_chain_integral_chain_integral3 [OF fy] ynz + by (force simp: fabc contour_integral_reverse_linepath) + next + case False + then have car3: "card {a, b, c} = Suc (DIM(complex))" + by auto + { assume "interior(convex hull {a,b,c}) = {}" + then have "collinear{a,b,c}" + using interior_convex_hull_eq_empty [OF car3] + by (simp add: collinear_3_eq_affine_dependent) + then have "False" + using False + apply (clarsimp simp add: collinear_3 collinear_lemma) + apply (drule Cauchy_theorem_flat [OF contf']) + using pi_eq_y ynz + apply (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) + done + } + then obtain d where d: "d \<in> interior (convex hull {a, b, c})" + by blast + { fix d1 + assume d1_pos: "0 < d1" + and d1: "\<And>x x'. \<lbrakk>x\<in>convex hull {a, b, c}; x'\<in>convex hull {a, b, c}; cmod (x' - x) < d1\<rbrakk> + \<Longrightarrow> cmod (f x' - f x) < cmod y / (24 * C)" + define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" + define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x + let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f" + have e: "0 < e" "e \<le> 1" "e \<le> d1 / (4 * C)" "e \<le> cmod y / 24 / C / B" + using d1_pos \<open>C>0\<close> \<open>B>0\<close> ynz by (simp_all add: e_def) + then have eCB: "24 * e * C * B \<le> cmod y" + using \<open>C>0\<close> \<open>B>0\<close> by (simp add: field_simps) + have e_le_d1: "e * (4 * C) \<le> d1" + using e \<open>C>0\<close> by (simp add: field_simps) + have "shrink a \<in> interior(convex hull {a,b,c})" + "shrink b \<in> interior(convex hull {a,b,c})" + "shrink c \<in> interior(convex hull {a,b,c})" + using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) + then have fhp0: "(f has_contour_integral 0) + (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" + by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal convex_interior) + then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" + by (simp add: has_chain_integral_chain_integral3) + have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" + "f contour_integrable_on linepath (shrink b) (shrink c)" + "f contour_integrable_on linepath (shrink c) (shrink a)" + using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) + have cmod_shr: "\<And>x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" + using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) + have sh_eq: "\<And>a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" + by (simp add: algebra_simps) + have "cmod y / (24 * C) \<le> cmod y / cmod (b - a) / 12" + using False \<open>C>0\<close> diff_2C [of b a] ynz + by (auto simp: divide_simps hull_inc) + have less_C: "\<lbrakk>u \<in> convex hull {a, b, c}; 0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> x * cmod u < C" for x u + apply (cases "x=0", simp add: \<open>0<C\<close>) + using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast + { fix u v + assume uv: "u \<in> convex hull {a, b, c}" "v \<in> convex hull {a, b, c}" "u\<noteq>v" + and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" + have shr_uv: "shrink u \<in> interior(convex hull {a,b,c})" + "shrink v \<in> interior(convex hull {a,b,c})" + using d e uv + by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) + have cmod_fuv: "\<And>x. 0\<le>x \<Longrightarrow> x\<le>1 \<Longrightarrow> cmod (f (linepath (shrink u) (shrink v) x)) \<le> B" + using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) + have By_uv: "B * (12 * (e * cmod (u - v))) \<le> cmod y" + apply (rule order_trans [OF _ eCB]) + using e \<open>B>0\<close> diff_2C [of u v] uv + by (auto simp: field_simps) + { fix x::real assume x: "0\<le>x" "x\<le>1" + have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" + apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) + using uv x d interior_subset + apply (auto simp: hull_inc intro!: less_C) + done + have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" + by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) + have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" + using \<open>e>0\<close> + apply (simp add: ll norm_mult scaleR_diff_right) + apply (rule less_le_trans [OF _ e_le_d1]) + using cmod_less_4C + apply (force intro: norm_triangle_lt) + done + have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" + using x uv shr_uv cmod_less_dt + by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) + also have "... \<le> cmod y / cmod (v - u) / 12" + using False uv \<open>C>0\<close> diff_2C [of v u] ynz + by (auto simp: divide_simps hull_inc) + finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \<le> cmod y / cmod (v - u) / 12" + by simp + then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \<le> cmod y" + using uv False by (auto simp: field_simps) + have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) + \<le> cmod y / 6" + apply (rule order_trans [of _ "B*((norm y / 24 / C / B)*2*C) + (2*C)*(norm y /24 / C)"]) + apply (rule add_mono [OF mult_mono]) + using By_uv e \<open>0 < B\<close> \<open>0 < C\<close> x ynz + apply (simp_all add: cmod_fuv cmod_shr cmod_12_le hull_inc) + apply (simp add: field_simps) + done + } note cmod_diff_le = this + have f_uv: "continuous_on (closed_segment u v) f" + by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) + have **: "\<And>f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" + by (simp add: algebra_simps) + have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \<le> norm y / 6" + apply (rule order_trans) + apply (rule has_integral_bound + [of "B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)" + "\<lambda>x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" + _ 0 1 ]) + using ynz \<open>0 < B\<close> \<open>0 < C\<close> + apply (simp_all del: le_divide_eq_numeral1) + apply (simp add: has_integral_sub has_contour_integral_linepath [symmetric] has_contour_integral_integral + fpi_uv f_uv contour_integrable_continuous_linepath, clarify) + apply (simp only: **) + apply (simp add: norm_triangle_le norm_mult cmod_diff_le del: le_divide_eq_numeral1) + done + } note * = this + have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \<le> norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + moreover + have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \<le> norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + moreover + have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \<le> norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + ultimately + have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + + (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) + \<le> norm y / 6 + norm y / 6 + norm y / 6" + by (metis norm_triangle_le add_mono) + also have "... = norm y / 2" + by simp + finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - + (?pathint a b + ?pathint b c + ?pathint c a)) + \<le> norm y / 2" + by (simp add: algebra_simps) + then + have "norm(?pathint a b + ?pathint b c + ?pathint c a) \<le> norm y / 2" + by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) + then have "False" + using pi_eq_y ynz by auto + } + moreover have "uniformly_continuous_on (convex hull {a,b,c}) f" + by (simp add: contf compact_convex_hull compact_uniformly_continuous) + ultimately have "False" + unfolding uniformly_continuous_on_def + by (force simp: ynz \<open>0 < C\<close> dist_norm) + then show ?thesis .. + qed + } + moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" + using fabc contour_integrable_continuous_linepath by auto + ultimately show ?thesis + using has_contour_integral_integral by fastforce +qed + + +subsection\<open>Version allowing finite number of exceptional points\<close> + +lemma Cauchy_theorem_triangle_cofinite: + assumes "continuous_on (convex hull {a,b,c}) f" + and "finite s" + and "(\<And>x. x \<in> interior(convex hull {a,b,c}) - s \<Longrightarrow> f field_differentiable (at x))" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +using assms +proof (induction "card s" arbitrary: a b c s rule: less_induct) + case (less s a b c) + show ?case + proof (cases "s={}") + case True with less show ?thesis + by (fastforce simp: holomorphic_on_def field_differentiable_at_within + Cauchy_theorem_triangle_interior) + next + case False + then obtain d s' where d: "s = insert d s'" "d \<notin> s'" + by (meson Set.set_insert all_not_in_conv) + then show ?thesis + proof (cases "d \<in> convex hull {a,b,c}") + case False + show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" + apply (rule less.hyps [of "s'"]) + using False d \<open>finite s\<close> interior_subset + apply (auto intro!: less.prems) + done + next + case True + have *: "convex hull {a, b, d} \<subseteq> convex hull {a, b, c}" + by (meson True hull_subset insert_subset convex_hull_subset) + have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)" + apply (rule less.hyps [of "s'"]) + using True d \<open>finite s\<close> not_in_interior_convex_hull_3 + apply (auto intro!: less.prems continuous_on_subset [OF _ *]) + apply (metis * insert_absorb insert_subset interior_mono) + done + have *: "convex hull {b, c, d} \<subseteq> convex hull {a, b, c}" + by (meson True hull_subset insert_subset convex_hull_subset) + have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)" + apply (rule less.hyps [of "s'"]) + using True d \<open>finite s\<close> not_in_interior_convex_hull_3 + apply (auto intro!: less.prems continuous_on_subset [OF _ *]) + apply (metis * insert_absorb insert_subset interior_mono) + done + have *: "convex hull {c, a, d} \<subseteq> convex hull {a, b, c}" + by (meson True hull_subset insert_subset convex_hull_subset) + have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)" + apply (rule less.hyps [of "s'"]) + using True d \<open>finite s\<close> not_in_interior_convex_hull_3 + apply (auto intro!: less.prems continuous_on_subset [OF _ *]) + apply (metis * insert_absorb insert_subset interior_mono) + done + have "f contour_integrable_on linepath a b" + using less.prems + by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3)) + moreover have "f contour_integrable_on linepath b c" + using less.prems + by (metis continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(3)) + moreover have "f contour_integrable_on linepath c a" + using less.prems + by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3)) + ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" + by auto + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \<noteq> 0" + have cont_ad: "continuous_on (closed_segment a d) f" + by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3)) + have cont_bd: "continuous_on (closed_segment b d) f" + by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1)) + have cont_cd: "continuous_on (closed_segment c d) f" + by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2)) + have "contour_integral (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))" + "contour_integral (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))" + "contour_integral (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)" + using has_chain_integral_chain_integral3 [OF abd] + has_chain_integral_chain_integral3 [OF bcd] + has_chain_integral_chain_integral3 [OF cad] + by (simp_all add: algebra_simps add_eq_0_iff) + then have ?thesis + using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce + } + then show ?thesis + using fpi contour_integrable_on_def by blast + qed + qed +qed + + +subsection\<open>Cauchy's theorem for an open starlike set\<close> + +lemma starlike_convex_subset: + assumes s: "a \<in> s" "closed_segment b c \<subseteq> s" and subs: "\<And>x. x \<in> s \<Longrightarrow> closed_segment a x \<subseteq> s" + shows "convex hull {a,b,c} \<subseteq> s" + using s + apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull) + apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE) + done + +lemma triangle_contour_integrals_starlike_primitive: + assumes contf: "continuous_on s f" + and s: "a \<in> s" "open s" + and x: "x \<in> s" + and subs: "\<And>y. y \<in> s \<Longrightarrow> closed_segment a y \<subseteq> s" + and zer: "\<And>b c. closed_segment b c \<subseteq> s + \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)" +proof - + let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f" + { fix e y + assume e: "0 < e" and bxe: "ball x e \<subseteq> s" and close: "cmod (y - x) < e" + have y: "y \<in> s" + using bxe close by (force simp: dist_norm norm_minus_commute) + have cont_ayf: "continuous_on (closed_segment a y) f" + using contf continuous_on_subset subs y by blast + have xys: "closed_segment x y \<subseteq> s" + apply (rule order_trans [OF _ bxe]) + using close + by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound) + have "?pathint a y - ?pathint a x = ?pathint x y" + using zer [OF xys] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force + } note [simp] = this + { fix e::real + assume e: "0 < e" + have cont_atx: "continuous (at x) f" + using x s contf continuous_on_eq_continuous_at by blast + then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. cmod (y - x) < d1 \<Longrightarrow> cmod (f y - f x) < e/2" + unfolding continuous_at Lim_at dist_norm using e + by (drule_tac x="e/2" in spec) force + obtain d2 where d2: "d2>0" "ball x d2 \<subseteq> s" using \<open>open s\<close> x + by (auto simp: open_contains_ball) + have dpos: "min d1 d2 > 0" using d1 d2 by simp + { fix y + assume yx: "y \<noteq> x" and close: "cmod (y - x) < min d1 d2" + have y: "y \<in> s" + using d2 close by (force simp: dist_norm norm_minus_commute) + have fxy: "f contour_integrable_on linepath x y" + apply (rule contour_integrable_continuous_linepath) + apply (rule continuous_on_subset [OF contf]) + using close d2 + apply (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1)) + done + then obtain i where i: "(f has_contour_integral i) (linepath x y)" + by (auto simp: contour_integrable_on_def) + then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" + by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) + then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)" + apply (rule has_contour_integral_bound_linepath [where B = "e/2"]) + using e apply simp + apply (rule d1_less [THEN less_imp_le]) + using close segment_bound + apply force + done + also have "... < e * cmod (y - x)" + by (simp add: e yx) + finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using i yx by (simp add: contour_integral_unique divide_less_eq) + } + then have "\<exists>d>0. \<forall>y. y \<noteq> x \<and> cmod (y-x) < d \<longrightarrow> cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using dpos by blast + } + then have *: "(\<lambda>y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \<midarrow>x\<rightarrow> 0" + by (simp add: Lim_at dist_norm inverse_eq_divide) + show ?thesis + apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right) + apply (rule Lim_transform [OF * Lim_eventually]) + apply (simp add: inverse_eq_divide [symmetric] eventually_at) + using \<open>open s\<close> x + apply (force simp: dist_norm open_contains_ball) + done +qed + +(** Existence of a primitive.*) + +lemma holomorphic_starlike_primitive: + fixes f :: "complex \<Rightarrow> complex" + assumes contf: "continuous_on s f" + and s: "starlike s" and os: "open s" + and k: "finite k" + and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x" + shows "\<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x)" +proof - + obtain a where a: "a\<in>s" and a_cs: "\<And>x. x\<in>s \<Longrightarrow> closed_segment a x \<subseteq> s" + using s by (auto simp: starlike_def) + { fix x b c + assume "x \<in> s" "closed_segment b c \<subseteq> s" + then have abcs: "convex hull {a, b, c} \<subseteq> s" + by (simp add: a a_cs starlike_convex_subset) + then have *: "continuous_on (convex hull {a, b, c}) f" + by (simp add: continuous_on_subset [OF contf]) + have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" + apply (rule Cauchy_theorem_triangle_cofinite [OF _ k]) + using abcs apply (simp add: continuous_on_subset [OF contf]) + using * abcs interior_subset apply (auto intro: fcd) + done + } note 0 = this + show ?thesis + apply (intro exI ballI) + apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption) + apply (metis a_cs) + apply (metis has_chain_integral_chain_integral3 0) + done +qed + +lemma Cauchy_theorem_starlike: + "\<lbrakk>open s; starlike s; finite k; continuous_on s f; + \<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x; + valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk> + \<Longrightarrow> (f has_contour_integral 0) g" + by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open) + +lemma Cauchy_theorem_starlike_simple: + "\<lbrakk>open s; starlike s; f holomorphic_on s; valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk> + \<Longrightarrow> (f has_contour_integral 0) g" +apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI]) +apply (simp_all add: holomorphic_on_imp_continuous_on) +apply (metis at_within_open holomorphic_on_def) +done + + +subsection\<open>Cauchy's theorem for a convex set\<close> + +text\<open>For a convex set we can avoid assuming openness and boundary analyticity\<close> + +lemma triangle_contour_integrals_convex_primitive: + assumes contf: "continuous_on s f" + and s: "a \<in> s" "convex s" + and x: "x \<in> s" + and zer: "\<And>b c. \<lbrakk>b \<in> s; c \<in> s\<rbrakk> + \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within s)" +proof - + let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f" + { fix y + assume y: "y \<in> s" + have cont_ayf: "continuous_on (closed_segment a y) f" + using s y by (meson contf continuous_on_subset convex_contains_segment) + have xys: "closed_segment x y \<subseteq> s" (*?*) + using convex_contains_segment s x y by auto + have "?pathint a y - ?pathint a x = ?pathint x y" + using zer [OF x y] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force + } note [simp] = this + { fix e::real + assume e: "0 < e" + have cont_atx: "continuous (at x within s) f" + using x s contf by (simp add: continuous_on_eq_continuous_within) + then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. \<lbrakk>y \<in> s; cmod (y - x) < d1\<rbrakk> \<Longrightarrow> cmod (f y - f x) < e/2" + unfolding continuous_within Lim_within dist_norm using e + by (drule_tac x="e/2" in spec) force + { fix y + assume yx: "y \<noteq> x" and close: "cmod (y - x) < d1" and y: "y \<in> s" + have fxy: "f contour_integrable_on linepath x y" + using convex_contains_segment s x y + by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) + then obtain i where i: "(f has_contour_integral i) (linepath x y)" + by (auto simp: contour_integrable_on_def) + then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" + by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) + then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)" + apply (rule has_contour_integral_bound_linepath [where B = "e/2"]) + using e apply simp + apply (rule d1_less [THEN less_imp_le]) + using convex_contains_segment s(2) x y apply blast + using close segment_bound(1) apply fastforce + done + also have "... < e * cmod (y - x)" + by (simp add: e yx) + finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using i yx by (simp add: contour_integral_unique divide_less_eq) + } + then have "\<exists>d>0. \<forall>y\<in>s. y \<noteq> x \<and> cmod (y-x) < d \<longrightarrow> cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using d1 by blast + } + then have *: "((\<lambda>y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \<longlongrightarrow> 0) (at x within s)" + by (simp add: Lim_within dist_norm inverse_eq_divide) + show ?thesis + apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) + apply (rule Lim_transform [OF * Lim_eventually]) + using linordered_field_no_ub + apply (force simp: inverse_eq_divide [symmetric] eventually_at) + done +qed + +lemma contour_integral_convex_primitive: + "\<lbrakk>convex s; continuous_on s f; + \<And>a b c. \<lbrakk>a \<in> s; b \<in> s; c \<in> s\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)\<rbrakk> + \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)" + apply (cases "s={}") + apply (simp_all add: ex_in_conv [symmetric]) + apply (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3) + done + +lemma holomorphic_convex_primitive: + fixes f :: "complex \<Rightarrow> complex" + shows + "\<lbrakk>convex s; finite k; continuous_on s f; + \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x\<rbrakk> + \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)" +apply (rule contour_integral_convex_primitive [OF _ _ Cauchy_theorem_triangle_cofinite]) +prefer 3 +apply (erule continuous_on_subset) +apply (simp add: subset_hull continuous_on_subset, assumption+) +by (metis Diff_iff convex_contains_segment insert_absorb insert_subset interior_mono segment_convex_hull subset_hull) + +lemma Cauchy_theorem_convex: + "\<lbrakk>continuous_on s f; convex s; finite k; + \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x; + valid_path g; path_image g \<subseteq> s; + pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g" + by (metis holomorphic_convex_primitive Cauchy_theorem_primitive) + +lemma Cauchy_theorem_convex_simple: + "\<lbrakk>f holomorphic_on s; convex s; + valid_path g; path_image g \<subseteq> s; + pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g" + apply (rule Cauchy_theorem_convex) + apply (simp_all add: holomorphic_on_imp_continuous_on) + apply (rule finite.emptyI) + using at_within_interior holomorphic_on_def interior_subset by fastforce + + +text\<open>In particular for a disc\<close> +lemma Cauchy_theorem_disc: + "\<lbrakk>finite k; continuous_on (cball a e) f; + \<And>x. x \<in> ball a e - k \<Longrightarrow> f field_differentiable at x; + valid_path g; path_image g \<subseteq> cball a e; + pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g" + apply (rule Cauchy_theorem_convex) + apply (auto simp: convex_cball interior_cball) + done + +lemma Cauchy_theorem_disc_simple: + "\<lbrakk>f holomorphic_on (ball a e); valid_path g; path_image g \<subseteq> ball a e; + pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g" +by (simp add: Cauchy_theorem_convex_simple) + + +subsection\<open>Generalize integrability to local primitives\<close> + +lemma contour_integral_local_primitive_lemma: + fixes f :: "complex\<Rightarrow>complex" + shows + "\<lbrakk>g piecewise_differentiable_on {a..b}; + \<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s); + \<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s\<rbrakk> + \<Longrightarrow> (\<lambda>x. f' (g x) * vector_derivative g (at x within {a..b})) + integrable_on {a..b}" + apply (cases "cbox a b = {}", force) + apply (simp add: integrable_on_def) + apply (rule exI) + apply (rule contour_integral_primitive_lemma, assumption+) + using atLeastAtMost_iff by blast + +lemma contour_integral_local_primitive_any: + fixes f :: "complex \<Rightarrow> complex" + assumes gpd: "g piecewise_differentiable_on {a..b}" + and dh: "\<And>x. x \<in> s + \<Longrightarrow> \<exists>d h. 0 < d \<and> + (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))" + and gs: "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s" + shows "(\<lambda>x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}" +proof - + { fix x + assume x: "a \<le> x" "x \<le> b" + obtain d h where d: "0 < d" + and h: "(\<And>y. norm(y - g x) < d \<Longrightarrow> (h has_field_derivative f y) (at y within s))" + using x gs dh by (metis atLeastAtMost_iff) + have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast + then obtain e where e: "e>0" and lessd: "\<And>x'. x' \<in> {a..b} \<Longrightarrow> \<bar>x' - x\<bar> < e \<Longrightarrow> cmod (g x' - g x) < d" + using x d + apply (auto simp: dist_norm continuous_on_iff) + apply (drule_tac x=x in bspec) + using x apply simp + apply (drule_tac x=d in spec, auto) + done + have "\<exists>d>0. \<forall>u v. u \<le> x \<and> x \<le> v \<and> {u..v} \<subseteq> ball x d \<and> (u \<le> v \<longrightarrow> a \<le> u \<and> v \<le> b) \<longrightarrow> + (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}" + apply (rule_tac x=e in exI) + using e + apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify) + apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma) + apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset) + apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force) + done + } then + show ?thesis + by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified]) +qed + +lemma contour_integral_local_primitive: + fixes f :: "complex \<Rightarrow> complex" + assumes g: "valid_path g" "path_image g \<subseteq> s" + and dh: "\<And>x. x \<in> s + \<Longrightarrow> \<exists>d h. 0 < d \<and> + (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))" + shows "f contour_integrable_on g" + using g + apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def + has_integral_localized_vector_derivative integrable_on_def [symmetric]) + using contour_integral_local_primitive_any [OF _ dh] + by (meson image_subset_iff piecewise_C1_imp_differentiable) + + +text\<open>In particular if a function is holomorphic\<close> + +lemma contour_integrable_holomorphic: + assumes contf: "continuous_on s f" + and os: "open s" + and k: "finite k" + and g: "valid_path g" "path_image g \<subseteq> s" + and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x" + shows "f contour_integrable_on g" +proof - + { fix z + assume z: "z \<in> s" + obtain d where d: "d>0" "ball z d \<subseteq> s" using \<open>open s\<close> z + by (auto simp: open_contains_ball) + then have contfb: "continuous_on (ball z d) f" + using contf continuous_on_subset by blast + obtain h where "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within ball z d)" + using holomorphic_convex_primitive [OF convex_ball k contfb fcd] d + interior_subset by force + then have "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within s)" + by (metis Topology_Euclidean_Space.open_ball at_within_open d(2) os subsetCE) + then have "\<exists>h. (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))" + by (force simp: dist_norm norm_minus_commute) + then have "\<exists>d h. 0 < d \<and> (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))" + using d by blast + } + then show ?thesis + by (rule contour_integral_local_primitive [OF g]) +qed + +lemma contour_integrable_holomorphic_simple: + assumes fh: "f holomorphic_on s" + and os: "open s" + and g: "valid_path g" "path_image g \<subseteq> s" + shows "f contour_integrable_on g" + apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g]) + apply (simp add: fh holomorphic_on_imp_continuous_on) + using fh by (simp add: field_differentiable_def holomorphic_on_open os) + +lemma continuous_on_inversediff: + fixes z:: "'a::real_normed_field" shows "z \<notin> s \<Longrightarrow> continuous_on s (\<lambda>w. 1 / (w - z))" + by (rule continuous_intros | force)+ + +corollary contour_integrable_inversediff: + "\<lbrakk>valid_path g; z \<notin> path_image g\<rbrakk> \<Longrightarrow> (\<lambda>w. 1 / (w-z)) contour_integrable_on g" +apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"]) +apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros) +done + +text\<open>Key fact that path integral is the same for a "nearby" path. This is the + main lemma for the homotopy form of Cauchy's theorem and is also useful + if we want "without loss of generality" to assume some nice properties of a + path (e.g. smoothness). It can also be used to define the integrals of + analytic functions over arbitrary continuous paths. This is just done for + winding numbers now. +\<close> + +text\<open>A technical definition to avoid duplication of similar proofs, + for paths joined at the ends versus looping paths\<close> +definition linked_paths :: "bool \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool" + where "linked_paths atends g h == + (if atends then pathstart h = pathstart g \<and> pathfinish h = pathfinish g + else pathfinish g = pathstart g \<and> pathfinish h = pathstart h)" + +text\<open>This formulation covers two cases: @{term g} and @{term h} share their + start and end points; @{term g} and @{term h} both loop upon themselves.\<close> +lemma contour_integral_nearby: + assumes os: "open s" and p: "path p" "path_image p \<subseteq> s" + shows + "\<exists>d. 0 < d \<and> + (\<forall>g h. valid_path g \<and> valid_path h \<and> + (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and> + linked_paths atends g h + \<longrightarrow> path_image g \<subseteq> s \<and> path_image h \<subseteq> s \<and> + (\<forall>f. f holomorphic_on s \<longrightarrow> contour_integral h f = contour_integral g f))" +proof - + have "\<forall>z. \<exists>e. z \<in> path_image p \<longrightarrow> 0 < e \<and> ball z e \<subseteq> s" + using open_contains_ball os p(2) by blast + then obtain ee where ee: "\<And>z. z \<in> path_image p \<Longrightarrow> 0 < ee z \<and> ball z (ee z) \<subseteq> s" + by metis + define cover where "cover = (\<lambda>z. ball z (ee z/3)) ` (path_image p)" + have "compact (path_image p)" + by (metis p(1) compact_path_image) + moreover have "path_image p \<subseteq> (\<Union>c\<in>path_image p. ball c (ee c / 3))" + using ee by auto + ultimately have "\<exists>D \<subseteq> cover. finite D \<and> path_image p \<subseteq> \<Union>D" + by (simp add: compact_eq_heine_borel cover_def) + then obtain D where D: "D \<subseteq> cover" "finite D" "path_image p \<subseteq> \<Union>D" + by blast + then obtain k where k: "k \<subseteq> {0..1}" "finite k" and D_eq: "D = ((\<lambda>z. ball z (ee z / 3)) \<circ> p) ` k" + apply (simp add: cover_def path_image_def image_comp) + apply (blast dest!: finite_subset_image [OF \<open>finite D\<close>]) + done + then have kne: "k \<noteq> {}" + using D by auto + have pi: "\<And>i. i \<in> k \<Longrightarrow> p i \<in> path_image p" + using k by (auto simp: path_image_def) + then have eepi: "\<And>i. i \<in> k \<Longrightarrow> 0 < ee((p i))" + by (metis ee) + define e where "e = Min((ee o p) ` k)" + have fin_eep: "finite ((ee o p) ` k)" + using k by blast + have enz: "0 < e" + using ee k by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi) + have "uniformly_continuous_on {0..1} p" + using p by (simp add: path_def compact_uniformly_continuous) + then obtain d::real where d: "d>0" + and de: "\<And>x x'. \<bar>x' - x\<bar> < d \<Longrightarrow> x\<in>{0..1} \<Longrightarrow> x'\<in>{0..1} \<Longrightarrow> cmod (p x' - p x) < e/3" + unfolding uniformly_continuous_on_def dist_norm real_norm_def + by (metis divide_pos_pos enz zero_less_numeral) + then obtain N::nat where N: "N>0" "inverse N < d" + using real_arch_inverse [of d] by auto + { fix g h + assume g: "valid_path g" and gp: "\<forall>t\<in>{0..1}. cmod (g t - p t) < e / 3" + and h: "valid_path h" and hp: "\<forall>t\<in>{0..1}. cmod (h t - p t) < e / 3" + and joins: "linked_paths atends g h" + { fix t::real + assume t: "0 \<le> t" "t \<le> 1" + then obtain u where u: "u \<in> k" and ptu: "p t \<in> ball(p u) (ee(p u) / 3)" + using \<open>path_image p \<subseteq> \<Union>D\<close> D_eq by (force simp: path_image_def) + then have ele: "e \<le> ee (p u)" using fin_eep + by (simp add: e_def) + have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3" + using gp hp t by auto + with ele have "cmod (g t - p t) < ee (p u) / 3" + "cmod (h t - p t) < ee (p u) / 3" + by linarith+ + then have "g t \<in> ball(p u) (ee(p u))" "h t \<in> ball(p u) (ee(p u))" + using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"] + norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u + by (force simp: dist_norm ball_def norm_minus_commute)+ + then have "g t \<in> s" "h t \<in> s" using ee u k + by (auto simp: path_image_def ball_def) + } + then have ghs: "path_image g \<subseteq> s" "path_image h \<subseteq> s" + by (auto simp: path_image_def) + moreover + { fix f + assume fhols: "f holomorphic_on s" + then have fpa: "f contour_integrable_on g" "f contour_integrable_on h" + using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple + by blast+ + have contf: "continuous_on s f" + by (simp add: fhols holomorphic_on_imp_continuous_on) + { fix z + assume z: "z \<in> path_image p" + have "f holomorphic_on ball z (ee z)" + using fhols ee z holomorphic_on_subset by blast + then have "\<exists>ff. (\<forall>w \<in> ball z (ee z). (ff has_field_derivative f w) (at w))" + using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified] + by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball) + } + then obtain ff where ff: + "\<And>z w. \<lbrakk>z \<in> path_image p; w \<in> ball z (ee z)\<rbrakk> \<Longrightarrow> (ff z has_field_derivative f w) (at w)" + by metis + { fix n + assume n: "n \<le> N" + then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f = + contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f" + proof (induct n) + case 0 show ?case by simp + next + case (Suc n) + obtain t where t: "t \<in> k" and "p (n/N) \<in> ball(p t) (ee(p t) / 3)" + using \<open>path_image p \<subseteq> \<Union>D\<close> [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems + by (force simp: path_image_def) + then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3" + by (simp add: dist_norm) + have e3le: "e/3 \<le> ee (p t) / 3" using fin_eep t + by (simp add: e_def) + { fix x + assume x: "n/N \<le> x" "x \<le> (1 + n)/N" + then have nN01: "0 \<le> n/N" "(1 + n)/N \<le> 1" + using Suc.prems by auto + then have x01: "0 \<le> x" "x \<le> 1" + using x by linarith+ + have "cmod (p t - p x) < ee (p t) / 3 + e/3" + apply (rule norm_diff_triangle_less [OF ptu de]) + using x N x01 Suc.prems + apply (auto simp: field_simps) + done + then have ptx: "cmod (p t - p x) < 2*ee (p t)/3" + using e3le eepi [OF t] by simp + have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 " + apply (rule norm_diff_triangle_less [OF ptx]) + using gp x01 by (simp add: norm_minus_commute) + also have "... \<le> ee (p t)" + using e3le eepi [OF t] by simp + finally have gg: "cmod (p t - g x) < ee (p t)" . + have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 " + apply (rule norm_diff_triangle_less [OF ptx]) + using hp x01 by (simp add: norm_minus_commute) + also have "... \<le> ee (p t)" + using e3le eepi [OF t] by simp + finally have "cmod (p t - g x) < ee (p t)" + "cmod (p t - h x) < ee (p t)" + using gg by auto + } note ptgh_ee = this + have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \<subseteq> ball (p t) (ee (p t))" + using ptgh_ee [of "n/N"] Suc.prems + by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) + then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \<subseteq> s" + using \<open>N>0\<close> Suc.prems + apply (simp add: path_image_join field_simps closed_segment_commute) + apply (erule order_trans) + apply (simp add: ee pi t) + done + have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) + \<subseteq> ball (p t) (ee (p t))" + using ptgh_ee [of "(1+n)/N"] Suc.prems + by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) + then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \<subseteq> s" + using \<open>N>0\<close> Suc.prems ee pi t + by (auto simp: Path_Connected.path_image_join field_simps) + have pi_subset_ball: + "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++ + subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N))) + \<subseteq> ball (p t) (ee (p t))" + apply (intro subset_path_image_join pi_hgn pi_ghn') + using \<open>N>0\<close> Suc.prems + apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) + done + have pi0: "(f has_contour_integral 0) + (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ + subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))" + apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"]) + apply (metis ff open_ball at_within_open pi t) + apply (intro valid_path_join) + using Suc.prems pi_subset_ball apply (simp_all add: valid_path_subpath g h) + done + have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g" + using Suc.prems by (simp add: contour_integrable_subpath g fpa) + have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))" + using gh_n's + by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) + have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))" + using gh_ns + by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) + have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f + + contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f + + contour_integral (subpath ((Suc n) / N) (n/N) h) f + + contour_integral (linepath (h (n/N)) (g (n/N))) f = 0" + using contour_integral_unique [OF pi0] Suc.prems + by (simp add: g h fpa valid_path_subpath contour_integrable_subpath + fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) + have *: "\<And>hn he hn' gn gd gn' hgn ghn gh0 ghn'. + \<lbrakk>hn - gn = ghn - gh0; + gd + ghn' + he + hgn = (0::complex); + hn - he = hn'; gn + gd = gn'; hgn = -ghn\<rbrakk> \<Longrightarrow> hn' - gn' = ghn' - gh0" + by (auto simp: algebra_simps) + have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = + contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f" + unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"] + using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath) + also have "... = contour_integral (subpath 0 ((Suc n) / N) h) f" + using Suc.prems by (simp add: contour_integral_subpath_combine h fpa) + finally have pi0_eq: + "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = + contour_integral (subpath 0 ((Suc n) / N) h) f" . + show ?case + apply (rule * [OF Suc.hyps eq0 pi0_eq]) + using Suc.prems + apply (simp_all add: g h fpa contour_integral_subpath_combine + contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath + continuous_on_subset [OF contf gh_ns]) + done + qed + } note ind = this + have "contour_integral h f = contour_integral g f" + using ind [OF order_refl] N joins + by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm) + } + ultimately + have "path_image g \<subseteq> s \<and> path_image h \<subseteq> s \<and> (\<forall>f. f holomorphic_on s \<longrightarrow> contour_integral h f = contour_integral g f)" + by metis + } note * = this + show ?thesis + apply (rule_tac x="e/3" in exI) + apply (rule conjI) + using enz apply simp + apply (clarsimp simp only: ball_conj_distrib) + apply (rule *; assumption) + done +qed + + +lemma + assumes "open s" "path p" "path_image p \<subseteq> s" + shows contour_integral_nearby_ends: + "\<exists>d. 0 < d \<and> + (\<forall>g h. valid_path g \<and> valid_path h \<and> + (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and> + pathstart h = pathstart g \<and> pathfinish h = pathfinish g + \<longrightarrow> path_image g \<subseteq> s \<and> + path_image h \<subseteq> s \<and> + (\<forall>f. f holomorphic_on s + \<longrightarrow> contour_integral h f = contour_integral g f))" + and contour_integral_nearby_loops: + "\<exists>d. 0 < d \<and> + (\<forall>g h. valid_path g \<and> valid_path h \<and> + (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and> + pathfinish g = pathstart g \<and> pathfinish h = pathstart h + \<longrightarrow> path_image g \<subseteq> s \<and> + path_image h \<subseteq> s \<and> + (\<forall>f. f holomorphic_on s + \<longrightarrow> contour_integral h f = contour_integral g f))" + using contour_integral_nearby [OF assms, where atends=True] + using contour_integral_nearby [OF assms, where atends=False] + unfolding linked_paths_def by simp_all + +corollary differentiable_polynomial_function: + fixes p :: "real \<Rightarrow> 'a::euclidean_space" + shows "polynomial_function p \<Longrightarrow> p differentiable_on s" +by (meson has_vector_derivative_polynomial_function differentiable_at_imp_differentiable_on differentiable_def has_vector_derivative_def) + +lemma C1_differentiable_polynomial_function: + fixes p :: "real \<Rightarrow> 'a::euclidean_space" + shows "polynomial_function p \<Longrightarrow> p C1_differentiable_on s" + by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) + +lemma valid_path_polynomial_function: + fixes p :: "real \<Rightarrow> 'a::euclidean_space" + shows "polynomial_function p \<Longrightarrow> valid_path p" +by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) + +lemma valid_path_subpath_trivial [simp]: + fixes g :: "real \<Rightarrow> 'a::euclidean_space" + shows "z \<noteq> g x \<Longrightarrow> valid_path (subpath x x g)" + by (simp add: subpath_def valid_path_polynomial_function) + +lemma contour_integral_bound_exists: +assumes s: "open s" + and g: "valid_path g" + and pag: "path_image g \<subseteq> s" + shows "\<exists>L. 0 < L \<and> + (\<forall>f B. f holomorphic_on s \<and> (\<forall>z \<in> s. norm(f z) \<le> B) + \<longrightarrow> norm(contour_integral g f) \<le> L*B)" +proof - +have "path g" using g + by (simp add: valid_path_imp_path) +then obtain d::real and p + where d: "0 < d" + and p: "polynomial_function p" "path_image p \<subseteq> s" + and pi: "\<And>f. f holomorphic_on s \<Longrightarrow> contour_integral g f = contour_integral p f" + using contour_integral_nearby_ends [OF s \<open>path g\<close> pag] + apply clarify + apply (drule_tac x=g in spec) + apply (simp only: assms) + apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function) + done +then obtain p' where p': "polynomial_function p'" + "\<And>x. (p has_vector_derivative (p' x)) (at x)" + using has_vector_derivative_polynomial_function by force +then have "bounded(p' ` {0..1})" + using continuous_on_polymonial_function + by (force simp: intro!: compact_imp_bounded compact_continuous_image) +then obtain L where L: "L>0" and nop': "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (p' x) \<le> L" + by (force simp: bounded_pos) +{ fix f B + assume f: "f holomorphic_on s" + and B: "\<And>z. z\<in>s \<Longrightarrow> cmod (f z) \<le> B" + then have "f contour_integrable_on p \<and> valid_path p" + using p s + by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on) + moreover have "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (vector_derivative p (at x)) * cmod (f (p x)) \<le> L * B" + apply (rule mult_mono) + apply (subst Derivative.vector_derivative_at; force intro: p' nop') + using L B p + apply (auto simp: path_image_def image_subset_iff) + done + ultimately have "cmod (contour_integral g f) \<le> L * B" + apply (simp add: pi [OF f]) + apply (simp add: contour_integral_integral) + apply (rule order_trans [OF integral_norm_bound_integral]) + apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult) + done +} then +show ?thesis + by (force simp: L contour_integral_integral) +qed + +subsection\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close> + +text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close> + +lemma continuous_disconnected_range_constant: + assumes s: "connected s" + and conf: "continuous_on s f" + and fim: "f ` s \<subseteq> t" + and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}" + shows "\<exists>a. \<forall>x \<in> s. f x = a" +proof (cases "s = {}") + case True then show ?thesis by force +next + case False + { fix x assume "x \<in> s" + then have "f ` s \<subseteq> {f x}" + by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct) + } + with False show ?thesis + by blast +qed + +lemma discrete_subset_disconnected: + fixes s :: "'a::topological_space set" + fixes t :: "'b::real_normed_vector set" + assumes conf: "continuous_on s f" + and no: "\<And>x. x \<in> s \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" + shows "f ` s \<subseteq> {y. connected_component_set (f ` s) y = {y}}" +proof - + { fix x assume x: "x \<in> s" + then obtain e where "e>0" and ele: "\<And>y. \<lbrakk>y \<in> s; f y \<noteq> f x\<rbrakk> \<Longrightarrow> e \<le> norm (f y - f x)" + using conf no [OF x] by auto + then have e2: "0 \<le> e / 2" + by simp + have "f y = f x" if "y \<in> s" and ccs: "f y \<in> connected_component_set (f ` s) (f x)" for y + apply (rule ccontr) + using connected_closed [of "connected_component_set (f ` s) (f x)"] \<open>e>0\<close> + apply (simp add: del: ex_simps) + apply (drule spec [where x="cball (f x) (e / 2)"]) + apply (drule spec [where x="- ball(f x) e"]) + apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in) + apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less) + using centre_in_cball connected_component_refl_eq e2 x apply blast + using ccs + apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> s\<close>]) + done + moreover have "connected_component_set (f ` s) (f x) \<subseteq> f ` s" + by (auto simp: connected_component_in) + ultimately have "connected_component_set (f ` s) (f x) = {f x}" + by (auto simp: x) + } + with assms show ?thesis + by blast +qed + +lemma finite_implies_discrete: + fixes s :: "'a::topological_space set" + assumes "finite (f ` s)" + shows "(\<forall>x \<in> s. \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x))" +proof - + have "\<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" if "x \<in> s" for x + proof (cases "f ` s - {f x} = {}") + case True + with zero_less_numeral show ?thesis + by (fastforce simp add: Set.image_subset_iff cong: conj_cong) + next + case False + then obtain z where z: "z \<in> s" "f z \<noteq> f x" + by blast + have finn: "finite {norm (z - f x) |z. z \<in> f ` s - {f x}}" + using assms by simp + then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` s - {f x}}" + apply (rule finite_imp_less_Inf) + using z apply force+ + done + show ?thesis + by (force intro!: * cInf_le_finite [OF finn]) + qed + with assms show ?thesis + by blast +qed + +text\<open>This proof requires the existence of two separate values of the range type.\<close> +lemma finite_range_constant_imp_connected: + assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. + \<lbrakk>continuous_on s f; finite(f ` s)\<rbrakk> \<Longrightarrow> \<exists>a. \<forall>x \<in> s. f x = a" + shows "connected s" +proof - + { fix t u + assume clt: "closedin (subtopology euclidean s) t" + and clu: "closedin (subtopology euclidean s) u" + and tue: "t \<inter> u = {}" and tus: "t \<union> u = s" + have conif: "continuous_on s (\<lambda>x. if x \<in> t then 0 else 1)" + apply (subst tus [symmetric]) + apply (rule continuous_on_cases_local) + using clt clu tue + apply (auto simp: tus continuous_on_const) + done + have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` s)" + by (rule finite_subset [of _ "{0,1}"]) auto + have "t = {} \<or> u = {}" + using assms [OF conif fi] tus [symmetric] + by (auto simp: Ball_def) (metis IntI empty_iff one_neq_zero tue) + } + then show ?thesis + by (simp add: connected_closedin_eq) +qed + +lemma continuous_disconnected_range_constant_eq: + "(connected s \<longleftrightarrow> + (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. + \<forall>t. continuous_on s f \<and> f ` s \<subseteq> t \<and> (\<forall>y \<in> t. connected_component_set t y = {y}) + \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis1) + and continuous_discrete_range_constant_eq: + "(connected s \<longleftrightarrow> + (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. + continuous_on s f \<and> + (\<forall>x \<in> s. \<exists>e. 0 < e \<and> (\<forall>y. y \<in> s \<and> (f y \<noteq> f x) \<longrightarrow> e \<le> norm(f y - f x))) + \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis2) + and continuous_finite_range_constant_eq: + "(connected s \<longleftrightarrow> + (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. + continuous_on s f \<and> finite (f ` s) + \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis3) +proof - + have *: "\<And>s t u v. \<lbrakk>s \<Longrightarrow> t; t \<Longrightarrow> u; u \<Longrightarrow> v; v \<Longrightarrow> s\<rbrakk> + \<Longrightarrow> (s \<longleftrightarrow> t) \<and> (s \<longleftrightarrow> u) \<and> (s \<longleftrightarrow> v)" + by blast + have "?thesis1 \<and> ?thesis2 \<and> ?thesis3" + apply (rule *) + using continuous_disconnected_range_constant apply metis + apply clarify + apply (frule discrete_subset_disconnected; blast) + apply (blast dest: finite_implies_discrete) + apply (blast intro!: finite_range_constant_imp_connected) + done + then show ?thesis1 ?thesis2 ?thesis3 + by blast+ +qed + +lemma continuous_discrete_range_constant: + fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1" + assumes s: "connected s" + and "continuous_on s f" + and "\<And>x. x \<in> s \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" + shows "\<exists>a. \<forall>x \<in> s. f x = a" + using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms + by blast + +lemma continuous_finite_range_constant: + fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1" + assumes "connected s" + and "continuous_on s f" + and "finite (f ` s)" + shows "\<exists>a. \<forall>x \<in> s. f x = a" + using assms continuous_finite_range_constant_eq + by blast + + +text\<open>We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\<close> + +subsection\<open>Winding Numbers\<close> + +definition winding_number:: "[real \<Rightarrow> complex, complex] \<Rightarrow> complex" where + "winding_number \<gamma> z \<equiv> + @n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and> + pathstart p = pathstart \<gamma> \<and> + pathfinish p = pathfinish \<gamma> \<and> + (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and> + contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n" + +lemma winding_number: + assumes "path \<gamma>" "z \<notin> path_image \<gamma>" "0 < e" + shows "\<exists>p. valid_path p \<and> z \<notin> path_image p \<and> + pathstart p = pathstart \<gamma> \<and> + pathfinish p = pathfinish \<gamma> \<and> + (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and> + contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * winding_number \<gamma> z" +proof - + have "path_image \<gamma> \<subseteq> UNIV - {z}" + using assms by blast + then obtain d + where d: "d>0" + and pi_eq: "\<And>h1 h2. valid_path h1 \<and> valid_path h2 \<and> + (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d) \<and> + pathstart h2 = pathstart h1 \<and> pathfinish h2 = pathfinish h1 \<longrightarrow> + path_image h1 \<subseteq> UNIV - {z} \<and> path_image h2 \<subseteq> UNIV - {z} \<and> + (\<forall>f. f holomorphic_on UNIV - {z} \<longrightarrow> contour_integral h2 f = contour_integral h1 f)" + using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete) + then obtain h where h: "polynomial_function h \<and> pathstart h = pathstart \<gamma> \<and> pathfinish h = pathfinish \<gamma> \<and> + (\<forall>t \<in> {0..1}. norm(h t - \<gamma> t) < d/2)" + using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "d/2"] d by auto + define nn where "nn = 1/(2* pi*\<i>) * contour_integral h (\<lambda>w. 1/(w - z))" + have "\<exists>n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and> + pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> + (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and> + contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n" + (is "\<exists>n. \<forall>e > 0. ?PP e n") + proof (rule_tac x=nn in exI, clarify) + fix e::real + assume e: "e>0" + obtain p where p: "polynomial_function p \<and> + pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> (\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < min e (d / 2))" + using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "min e (d/2)"] d \<open>0<e\<close> by auto + have "(\<lambda>w. 1 / (w - z)) holomorphic_on UNIV - {z}" + by (auto simp: intro!: holomorphic_intros) + then show "?PP e nn" + apply (rule_tac x=p in exI) + using pi_eq [of h p] h p d + apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def) + done + qed + then show ?thesis + unfolding winding_number_def + apply (rule someI2_ex) + apply (blast intro: \<open>0<e\<close>) + done +qed + +lemma winding_number_unique: + assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>" + and pi: + "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and> + pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> + (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and> + contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n" + shows "winding_number \<gamma> z = n" +proof - + have "path_image \<gamma> \<subseteq> UNIV - {z}" + using assms by blast + then obtain e + where e: "e>0" + and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2; + (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e); + pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow> + contour_integral h2 f = contour_integral h1 f" + using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete) + obtain p where p: + "valid_path p \<and> z \<notin> path_image p \<and> + pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> + (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and> + contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n" + using pi [OF e] by blast + obtain q where q: + "valid_path q \<and> z \<notin> path_image q \<and> + pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and> + (\<forall>t\<in>{0..1}. cmod (\<gamma> t - q t) < e) \<and> contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z" + using winding_number [OF \<gamma> e] by blast + have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))" + using p by auto + also have "... = contour_integral q (\<lambda>w. 1 / (w - z))" + apply (rule pi_eq) + using p q + by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros) + also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" + using q by auto + finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" . + then show ?thesis + by simp +qed + +lemma winding_number_unique_loop: + assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>" + and loop: "pathfinish \<gamma> = pathstart \<gamma>" + and pi: + "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and> + pathfinish p = pathstart p \<and> + (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and> + contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n" + shows "winding_number \<gamma> z = n" +proof - + have "path_image \<gamma> \<subseteq> UNIV - {z}" + using assms by blast + then obtain e + where e: "e>0" + and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2; + (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e); + pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow> + contour_integral h2 f = contour_integral h1 f" + using contour_integral_nearby_loops [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete) + obtain p where p: + "valid_path p \<and> z \<notin> path_image p \<and> + pathfinish p = pathstart p \<and> + (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and> + contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n" + using pi [OF e] by blast + obtain q where q: + "valid_path q \<and> z \<notin> path_image q \<and> + pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and> + (\<forall>t\<in>{0..1}. cmod (\<gamma> t - q t) < e) \<and> contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z" + using winding_number [OF \<gamma> e] by blast + have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))" + using p by auto + also have "... = contour_integral q (\<lambda>w. 1 / (w - z))" + apply (rule pi_eq) + using p q loop + by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros) + also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" + using q by auto + finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" . + then show ?thesis + by simp +qed + +lemma winding_number_valid_path: + assumes "valid_path \<gamma>" "z \<notin> path_image \<gamma>" + shows "winding_number \<gamma> z = 1/(2*pi*\<i>) * contour_integral \<gamma> (\<lambda>w. 1/(w - z))" + using assms by (auto simp: valid_path_imp_path intro!: winding_number_unique) + +lemma has_contour_integral_winding_number: + assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>" + shows "((\<lambda>w. 1/(w - z)) has_contour_integral (2*pi*\<i>*winding_number \<gamma> z)) \<gamma>" +by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) + +lemma winding_number_trivial [simp]: "z \<noteq> a \<Longrightarrow> winding_number(linepath a a) z = 0" + by (simp add: winding_number_valid_path) + +lemma winding_number_subpath_trivial [simp]: "z \<noteq> g x \<Longrightarrow> winding_number (subpath x x g) z = 0" + by (simp add: path_image_subpath winding_number_valid_path) + +lemma winding_number_join: + assumes g1: "path g1" "z \<notin> path_image g1" + and g2: "path g2" "z \<notin> path_image g2" + and "pathfinish g1 = pathstart g2" + shows "winding_number(g1 +++ g2) z = winding_number g1 z + winding_number g2 z" + apply (rule winding_number_unique) + using assms apply (simp_all add: not_in_path_image_join) + apply (frule winding_number [OF g2]) + apply (frule winding_number [OF g1], clarify) + apply (rename_tac p2 p1) + apply (rule_tac x="p1+++p2" in exI) + apply (simp add: not_in_path_image_join contour_integrable_inversediff algebra_simps) + apply (auto simp: joinpaths_def) + done + +lemma winding_number_reversepath: + assumes "path \<gamma>" "z \<notin> path_image \<gamma>" + shows "winding_number(reversepath \<gamma>) z = - (winding_number \<gamma> z)" + apply (rule winding_number_unique) + using assms + apply simp_all + apply (frule winding_number [OF assms], clarify) + apply (rule_tac x="reversepath p" in exI) + apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) + apply (auto simp: reversepath_def) + done + +lemma winding_number_shiftpath: + assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>" + and "pathfinish \<gamma> = pathstart \<gamma>" "a \<in> {0..1}" + shows "winding_number(shiftpath a \<gamma>) z = winding_number \<gamma> z" + apply (rule winding_number_unique_loop) + using assms + apply (simp_all add: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath) + apply (frule winding_number [OF \<gamma>], clarify) + apply (rule_tac x="shiftpath a p" in exI) + apply (simp add: contour_integral_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath valid_path_shiftpath) + apply (auto simp: shiftpath_def) + done + +lemma winding_number_split_linepath: + assumes "c \<in> closed_segment a b" "z \<notin> closed_segment a b" + shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" +proof - + have "z \<notin> closed_segment a c" "z \<notin> closed_segment c b" + using assms apply (meson convex_contains_segment convex_segment ends_in_segment(1) subsetCE) + using assms by (meson convex_contains_segment convex_segment ends_in_segment(2) subsetCE) + then show ?thesis + using assms + by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) +qed + +lemma winding_number_cong: + "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p t = q t) \<Longrightarrow> winding_number p z = winding_number q z" + by (simp add: winding_number_def pathstart_def pathfinish_def) + +lemma winding_number_offset: "winding_number p z = winding_number (\<lambda>w. p w - z) 0" + apply (simp add: winding_number_def contour_integral_integral path_image_def valid_path_def pathstart_def pathfinish_def) + apply (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) + apply (rename_tac g) + apply (rule_tac x="\<lambda>t. g t - z" in exI) + apply (force simp: vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) + apply (rename_tac g) + apply (rule_tac x="\<lambda>t. g t + z" in exI) + apply (simp add: piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) + apply (force simp: algebra_simps) + done + +(* A combined theorem deducing several things piecewise.*) +lemma winding_number_join_pos_combined: + "\<lbrakk>valid_path \<gamma>1; z \<notin> path_image \<gamma>1; 0 < Re(winding_number \<gamma>1 z); + valid_path \<gamma>2; z \<notin> path_image \<gamma>2; 0 < Re(winding_number \<gamma>2 z); pathfinish \<gamma>1 = pathstart \<gamma>2\<rbrakk> + \<Longrightarrow> valid_path(\<gamma>1 +++ \<gamma>2) \<and> z \<notin> path_image(\<gamma>1 +++ \<gamma>2) \<and> 0 < Re(winding_number(\<gamma>1 +++ \<gamma>2) z)" + by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) + + +(* Useful sufficient conditions for the winding number to be positive etc.*) + +lemma Re_winding_number: + "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>\<rbrakk> + \<Longrightarrow> Re(winding_number \<gamma> z) = Im(contour_integral \<gamma> (\<lambda>w. 1/(w - z))) / (2*pi)" +by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) + +lemma winding_number_pos_le: + assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>" + and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> 0 \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))" + shows "0 \<le> Re(winding_number \<gamma> z)" +proof - + have *: "0 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))" if x: "0 < x" "x < 1" for x + using ge by (simp add: Complex.Im_divide algebra_simps x) + show ?thesis + apply (simp add: Re_winding_number [OF \<gamma>] field_simps) + apply (rule has_integral_component_nonneg + [of \<i> "\<lambda>x. if x \<in> {0<..<1} + then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else 0", simplified]) + prefer 3 apply (force simp: *) + apply (simp add: Basis_complex_def) + apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)"]) + apply simp + apply (simp only: box_real) + apply (subst has_contour_integral [symmetric]) + using \<gamma> + apply (simp add: contour_integrable_inversediff has_contour_integral_integral) + done +qed + +lemma winding_number_pos_lt_lemma: + assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>" + and e: "0 < e" + and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))" + shows "0 < Re(winding_number \<gamma> z)" +proof - + have "e \<le> Im (contour_integral \<gamma> (\<lambda>w. 1 / (w - z)))" + apply (rule has_integral_component_le + [of \<i> "\<lambda>x. \<i>*e" "\<i>*e" "{0..1}" + "\<lambda>x. if x \<in> {0<..<1} then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else \<i>*e" + "contour_integral \<gamma> (\<lambda>w. 1/(w - z))", simplified]) + using e + apply (simp_all add: Basis_complex_def) + using has_integral_const_real [of _ 0 1] apply force + apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)", simplified box_real]) + apply simp + apply (subst has_contour_integral [symmetric]) + using \<gamma> + apply (simp_all add: contour_integrable_inversediff has_contour_integral_integral ge) + done + with e show ?thesis + by (simp add: Re_winding_number [OF \<gamma>] field_simps) +qed + +lemma winding_number_pos_lt: + assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>" + and e: "0 < e" + and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))" + shows "0 < Re (winding_number \<gamma> z)" +proof - + have bm: "bounded ((\<lambda>w. w - z) ` (path_image \<gamma>))" + using bounded_translation [of _ "-z"] \<gamma> by (simp add: bounded_valid_path_image) + then obtain B where B: "B > 0" and Bno: "\<And>x. x \<in> (\<lambda>w. w - z) ` (path_image \<gamma>) \<Longrightarrow> norm x \<le> B" + using bounded_pos [THEN iffD1, OF bm] by blast + { fix x::real assume x: "0 < x" "x < 1" + then have B2: "cmod (\<gamma> x - z)^2 \<le> B^2" using Bno [of "\<gamma> x - z"] + by (simp add: path_image_def power2_eq_square mult_mono') + with x have "\<gamma> x \<noteq> z" using \<gamma> + using path_image_def by fastforce + then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) * cnj (\<gamma> x - z)) / (cmod (\<gamma> x - z))\<^sup>2" + using B ge [OF x] B2 e + apply (rule_tac y="e / (cmod (\<gamma> x - z))\<^sup>2" in order_trans) + apply (auto simp: divide_left_mono divide_right_mono) + done + then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))" + by (simp add: Im_divide_Reals complex_div_cnj [of _ "\<gamma> x - z" for x] del: complex_cnj_diff times_complex.sel) + } note * = this + show ?thesis + using e B by (simp add: * winding_number_pos_lt_lemma [OF \<gamma>, of "e/B^2"]) +qed + +subsection\<open>The winding number is an integer\<close> + +text\<open>Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, + Also on page 134 of Serge Lang's book with the name title, etc.\<close> + +lemma exp_fg: + fixes z::complex + assumes g: "(g has_vector_derivative g') (at x within s)" + and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" + and z: "g x \<noteq> z" + shows "((\<lambda>x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" +proof - + have *: "(exp o (\<lambda>x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" + using assms unfolding has_vector_derivative_def scaleR_conv_of_real + by (auto intro!: derivative_eq_intros) + show ?thesis + apply (rule has_vector_derivative_eq_rhs) + apply (rule bounded_bilinear.has_vector_derivative [OF bounded_bilinear_mult]) + using z + apply (auto simp: intro!: derivative_eq_intros * [unfolded o_def] g) + done +qed + +lemma winding_number_exp_integral: + fixes z::complex + assumes \<gamma>: "\<gamma> piecewise_C1_differentiable_on {a..b}" + and ab: "a \<le> b" + and z: "z \<notin> \<gamma> ` {a..b}" + shows "(\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)) integrable_on {a..b}" + (is "?thesis1") + "exp (- (integral {a..b} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))) * (\<gamma> b - z) = \<gamma> a - z" + (is "?thesis2") +proof - + let ?D\<gamma> = "\<lambda>x. vector_derivative \<gamma> (at x)" + have [simp]: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<gamma> x \<noteq> z" + using z by force + have cong: "continuous_on {a..b} \<gamma>" + using \<gamma> by (simp add: piecewise_C1_differentiable_on_def) + obtain k where fink: "finite k" and g_C1_diff: "\<gamma> C1_differentiable_on ({a..b} - k)" + using \<gamma> by (force simp: piecewise_C1_differentiable_on_def) + have o: "open ({a<..<b} - k)" + using \<open>finite k\<close> by (simp add: finite_imp_closed open_Diff) + moreover have "{a<..<b} - k \<subseteq> {a..b} - k" + by force + ultimately have g_diff_at: "\<And>x. \<lbrakk>x \<notin> k; x \<in> {a<..<b}\<rbrakk> \<Longrightarrow> \<gamma> differentiable at x" + by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open) + { fix w + assume "w \<noteq> z" + have "continuous_on (ball w (cmod (w - z))) (\<lambda>w. 1 / (w - z))" + by (auto simp: dist_norm intro!: continuous_intros) + moreover have "\<And>x. cmod (w - x) < cmod (w - z) \<Longrightarrow> \<exists>f'. ((\<lambda>w. 1 / (w - z)) has_field_derivative f') (at x)" + by (auto simp: intro!: derivative_eq_intros) + ultimately have "\<exists>h. \<forall>y. norm(y - w) < norm(w - z) \<longrightarrow> (h has_field_derivative 1/(y - z)) (at y)" + using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\<lambda>w. 1/(w - z)"] + by (simp add: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) + } + then obtain h where h: "\<And>w y. w \<noteq> z \<Longrightarrow> norm(y - w) < norm(w - z) \<Longrightarrow> (h w has_field_derivative 1/(y - z)) (at y)" + by meson + have exy: "\<exists>y. ((\<lambda>x. inverse (\<gamma> x - z) * ?D\<gamma> x) has_integral y) {a..b}" + unfolding integrable_on_def [symmetric] + apply (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \<gamma>], of "-{z}"]) + apply (rename_tac w) + apply (rule_tac x="norm(w - z)" in exI) + apply (simp_all add: inverse_eq_divide) + apply (metis has_field_derivative_at_within h) + done + have vg_int: "(\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)) integrable_on {a..b}" + unfolding box_real [symmetric] divide_inverse_commute + by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) + with ab show ?thesis1 + by (simp add: divide_inverse_commute integral_def integrable_on_def) + { fix t + assume t: "t \<in> {a..b}" + have cball: "continuous_on (ball (\<gamma> t) (dist (\<gamma> t) z)) (\<lambda>x. inverse (x - z))" + using z by (auto intro!: continuous_intros simp: dist_norm) + have icd: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow> (\<lambda>w. inverse (w - z)) field_differentiable at x" + unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) + obtain h where h: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow> + (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\<gamma> t - y) < cmod (\<gamma> t - z)})" + using holomorphic_convex_primitive [where f = "\<lambda>w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] + by simp (auto simp: ball_def dist_norm that) + { fix x D + assume x: "x \<notin> k" "a < x" "x < b" + then have "x \<in> interior ({a..b} - k)" + using open_subset_interior [OF o] by fastforce + then have con: "isCont (\<lambda>x. ?D\<gamma> x) x" + using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) + then have con_vd: "continuous (at x within {a..b}) (\<lambda>x. ?D\<gamma> x)" + by (rule continuous_at_imp_continuous_within) + have gdx: "\<gamma> differentiable at x" + using x by (simp add: g_diff_at) + have "((\<lambda>c. exp (- integral {a..c} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))) * (\<gamma> c - z)) has_derivative (\<lambda>h. 0)) + (at x within {a..b})" + using x gdx t + apply (clarsimp simp add: differentiable_iff_scaleR) + apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_within) + apply (simp_all add: has_vector_derivative_def [symmetric]) + apply (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at]) + apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ + done + } note * = this + have "exp (- (integral {a..t} (\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)))) * (\<gamma> t - z) =\<gamma> a - z" + apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \<union> k" a b]) + using t + apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous [OF vg_int] simp add: ab)+ + done + } + with ab show ?thesis2 + by (simp add: divide_inverse_commute integral_def) +qed + +corollary winding_number_exp_2pi: + "\<lbrakk>path p; z \<notin> path_image p\<rbrakk> + \<Longrightarrow> pathfinish p - z = exp (2 * pi * \<i> * winding_number p z) * (pathstart p - z)" +using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def + by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) + + +subsection\<open>The version with complex integers and equality\<close> + +lemma integer_winding_number_eq: + assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>" + shows "winding_number \<gamma> z \<in> \<int> \<longleftrightarrow> pathfinish \<gamma> = pathstart \<gamma>" +proof - + have *: "\<And>i::complex. \<And>g0 g1. \<lbrakk>i \<noteq> 0; g0 \<noteq> z; (g1 - z) / i = g0 - z\<rbrakk> \<Longrightarrow> (i = 1 \<longleftrightarrow> g1 = g0)" + by (simp add: field_simps) algebra + obtain p where p: "valid_path p" "z \<notin> path_image p" + "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>" + "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z" + using winding_number [OF assms, of 1] by auto + have [simp]: "(winding_number \<gamma> z \<in> \<int>) = (exp (contour_integral p (\<lambda>w. 1 / (w - z))) = 1)" + using p by (simp add: exp_eq_1 complex_is_Int_iff) + have "winding_number p z \<in> \<int> \<longleftrightarrow> pathfinish p = pathstart p" + using p z + apply (simp add: winding_number_valid_path valid_path_def path_image_def pathstart_def pathfinish_def) + using winding_number_exp_integral(2) [of p 0 1 z] + apply (simp add: field_simps contour_integral_integral exp_minus) + apply (rule *) + apply (auto simp: path_image_def field_simps) + done + then show ?thesis using p + by (auto simp: winding_number_valid_path) +qed + +theorem integer_winding_number: + "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> path_image \<gamma>\<rbrakk> \<Longrightarrow> winding_number \<gamma> z \<in> \<int>" +by (metis integer_winding_number_eq) + + +text\<open>If the winding number's magnitude is at least one, then the path must contain points in every direction.*) + We can thus bound the winding number of a path that doesn't intersect a given ray. \<close> + +lemma winding_number_pos_meets: + fixes z::complex + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and 1: "Re (winding_number \<gamma> z) \<ge> 1" + and w: "w \<noteq> z" + shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>" +proof - + have [simp]: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<gamma> x \<noteq> z" + using z by (auto simp: path_image_def) + have [simp]: "z \<notin> \<gamma> ` {0..1}" + using path_image_def z by auto + have gpd: "\<gamma> piecewise_C1_differentiable_on {0..1}" + using \<gamma> valid_path_def by blast + define r where "r = (w - z) / (\<gamma> 0 - z)" + have [simp]: "r \<noteq> 0" + using w z by (auto simp: r_def) + have "Arg r \<le> 2*pi" + by (simp add: Arg less_eq_real_def) + also have "... \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))" + using 1 + apply (simp add: winding_number_valid_path [OF \<gamma> z] contour_integral_integral) + apply (simp add: Complex.Re_divide field_simps power2_eq_square) + done + finally have "Arg r \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))" . + then have "\<exists>t. t \<in> {0..1} \<and> Im(integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r" + apply (simp add:) + apply (rule IVT') + apply (simp_all add: Arg_ge_0) + apply (intro continuous_intros indefinite_integral_continuous winding_number_exp_integral [OF gpd]; simp) + done + then obtain t where t: "t \<in> {0..1}" + and eqArg: "Im (integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r" + by blast + define i where "i = integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))" + have iArg: "Arg r = Im i" + using eqArg by (simp add: i_def) + have gpdt: "\<gamma> piecewise_C1_differentiable_on {0..t}" + by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) + have "exp (- i) * (\<gamma> t - z) = \<gamma> 0 - z" + unfolding i_def + apply (rule winding_number_exp_integral [OF gpdt]) + using t z unfolding path_image_def + apply force+ + done + then have *: "\<gamma> t - z = exp i * (\<gamma> 0 - z)" + by (simp add: exp_minus field_simps) + then have "(w - z) = r * (\<gamma> 0 - z)" + by (simp add: r_def) + then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \<gamma> t" + apply (simp add:) + apply (subst Complex_Transcendental.Arg_eq [of r]) + apply (simp add: iArg) + using * + apply (simp add: exp_eq_polar field_simps) + done + with t show ?thesis + by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) +qed + +lemma winding_number_big_meets: + fixes z::complex + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "\<bar>Re (winding_number \<gamma> z)\<bar> \<ge> 1" + and w: "w \<noteq> z" + shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>" +proof - + { assume "Re (winding_number \<gamma> z) \<le> - 1" + then have "Re (winding_number (reversepath \<gamma>) z) \<ge> 1" + by (simp add: \<gamma> valid_path_imp_path winding_number_reversepath z) + moreover have "valid_path (reversepath \<gamma>)" + using \<gamma> valid_path_imp_reverse by auto + moreover have "z \<notin> path_image (reversepath \<gamma>)" + by (simp add: z) + ultimately have "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image (reversepath \<gamma>)" + using winding_number_pos_meets w by blast + then have ?thesis + by simp + } + then show ?thesis + using assms + by (simp add: abs_if winding_number_pos_meets split: if_split_asm) +qed + +lemma winding_number_less_1: + fixes z::complex + shows + "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>; w \<noteq> z; + \<And>a::real. 0 < a \<Longrightarrow> z + a*(w - z) \<notin> path_image \<gamma>\<rbrakk> + \<Longrightarrow> \<bar>Re(winding_number \<gamma> z)\<bar> < 1" + by (auto simp: not_less dest: winding_number_big_meets) + +text\<open>One way of proving that WN=1 for a loop.\<close> +lemma winding_number_eq_1: + fixes z::complex + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + and 0: "0 < Re(winding_number \<gamma> z)" and 2: "Re(winding_number \<gamma> z) < 2" + shows "winding_number \<gamma> z = 1" +proof - + have "winding_number \<gamma> z \<in> Ints" + by (simp add: \<gamma> integer_winding_number loop valid_path_imp_path z) + then show ?thesis + using 0 2 by (auto simp: Ints_def) +qed + + +subsection\<open>Continuity of winding number and invariance on connected sets.\<close> + +lemma continuous_at_winding_number: + fixes z::complex + assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>" + shows "continuous (at z) (winding_number \<gamma>)" +proof - + obtain e where "e>0" and cbg: "cball z e \<subseteq> - path_image \<gamma>" + using open_contains_cball [of "- path_image \<gamma>"] z + by (force simp: closed_def [symmetric] closed_path_image [OF \<gamma>]) + then have ppag: "path_image \<gamma> \<subseteq> - cball z (e/2)" + by (force simp: cball_def dist_norm) + have oc: "open (- cball z (e / 2))" + by (simp add: closed_def [symmetric]) + obtain d where "d>0" and pi_eq: + "\<And>h1 h2. \<lbrakk>valid_path h1; valid_path h2; + (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d); + pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\<rbrakk> + \<Longrightarrow> + path_image h1 \<subseteq> - cball z (e / 2) \<and> + path_image h2 \<subseteq> - cball z (e / 2) \<and> + (\<forall>f. f holomorphic_on - cball z (e / 2) \<longrightarrow> contour_integral h2 f = contour_integral h1 f)" + using contour_integral_nearby_ends [OF oc \<gamma> ppag] by metis + obtain p where p: "valid_path p" "z \<notin> path_image p" + "pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma>" + and pg: "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (\<gamma> t - p t) < min d e / 2" + and pi: "contour_integral p (\<lambda>x. 1 / (x - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z" + using winding_number [OF \<gamma> z, of "min d e / 2"] \<open>d>0\<close> \<open>e>0\<close> by auto + { fix w + assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" + then have wnotp: "w \<notin> path_image p" + using cbg \<open>d>0\<close> \<open>e>0\<close> + apply (simp add: path_image_def cball_def dist_norm, clarify) + apply (frule pg) + apply (drule_tac c="\<gamma> x" in subsetD) + apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) + done + have wnotg: "w \<notin> path_image \<gamma>" + using cbg e2 \<open>e>0\<close> by (force simp: dist_norm norm_minus_commute) + { fix k::real + assume k: "k>0" + then obtain q where q: "valid_path q" "w \<notin> path_image q" + "pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma>" + and qg: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (\<gamma> t - q t) < min k (min d e) / 2" + and qi: "contour_integral q (\<lambda>u. 1 / (u - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w" + using winding_number [OF \<gamma> wnotg, of "min k (min d e) / 2"] \<open>d>0\<close> \<open>e>0\<close> k + by (force simp: min_divide_distrib_right) + have "contour_integral p (\<lambda>u. 1 / (u - w)) = contour_integral q (\<lambda>u. 1 / (u - w))" + apply (rule pi_eq [OF \<open>valid_path q\<close> \<open>valid_path p\<close>, THEN conjunct2, THEN conjunct2, rule_format]) + apply (frule pg) + apply (frule qg) + using p q \<open>d>0\<close> e2 + apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) + done + then have "contour_integral p (\<lambda>x. 1 / (x - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w" + by (simp add: pi qi) + } note pip = this + have "path p" + using p by (simp add: valid_path_imp_path) + then have "winding_number p w = winding_number \<gamma> w" + apply (rule winding_number_unique [OF _ wnotp]) + apply (rule_tac x=p in exI) + apply (simp add: p wnotp min_divide_distrib_right pip) + done + } note wnwn = this + obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \<subseteq> - path_image p" + using p open_contains_cball [of "- path_image p"] + by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) + obtain L + where "L>0" + and L: "\<And>f B. \<lbrakk>f holomorphic_on - cball z (3 / 4 * pe); + \<forall>z \<in> - cball z (3 / 4 * pe). cmod (f z) \<le> B\<rbrakk> \<Longrightarrow> + cmod (contour_integral p f) \<le> L * B" + using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \<open>valid_path p\<close> by force + { fix e::real and w::complex + assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" + then have [simp]: "w \<notin> path_image p" + using cbp p(2) \<open>0 < pe\<close> + by (force simp: dist_norm norm_minus_commute path_image_def cball_def) + have [simp]: "contour_integral p (\<lambda>x. 1/(x - w)) - contour_integral p (\<lambda>x. 1/(x - z)) = + contour_integral p (\<lambda>x. 1/(x - w) - 1/(x - z))" + by (simp add: p contour_integrable_inversediff contour_integral_diff) + { fix x + assume pe: "3/4 * pe < cmod (z - x)" + have "cmod (w - x) < pe/4 + cmod (z - x)" + by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) + then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp + have "cmod (z - x) \<le> cmod (z - w) + cmod (w - x)" + using norm_diff_triangle_le by blast + also have "... < pe/4 + cmod (w - x)" + using w by (simp add: norm_minus_commute) + finally have "pe/2 < cmod (w - x)" + using pe by auto + then have "(pe/2)^2 < cmod (w - x) ^ 2" + apply (rule power_strict_mono) + using \<open>pe>0\<close> by auto + then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" + by (simp add: power_divide) + have "8 * L * cmod (w - z) < e * pe\<^sup>2" + using w \<open>L>0\<close> by (simp add: field_simps) + also have "... < e * 4 * cmod (w - x) * cmod (w - x)" + using pe2 \<open>e>0\<close> by (simp add: power2_eq_square) + also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" + using wx + apply (rule mult_strict_left_mono) + using pe2 e not_less_iff_gr_or_eq by fastforce + finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" + by simp + also have "... \<le> e * cmod (w - x) * cmod (z - x)" + using e by simp + finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . + have "L * cmod (1 / (x - w) - 1 / (x - z)) \<le> e" + apply (cases "x=z \<or> x=w") + using pe \<open>pe>0\<close> w \<open>L>0\<close> + apply (force simp: norm_minus_commute) + using wx w(2) \<open>L>0\<close> pe pe2 Lwz + apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) + done + } note L_cmod_le = this + have *: "cmod (contour_integral p (\<lambda>x. 1 / (x - w) - 1 / (x - z))) \<le> L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" + apply (rule L) + using \<open>pe>0\<close> w + apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) + using \<open>pe>0\<close> w \<open>L>0\<close> + apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) + done + have "cmod (contour_integral p (\<lambda>x. 1 / (x - w)) - contour_integral p (\<lambda>x. 1 / (x - z))) < 2*e" + apply (simp add:) + apply (rule le_less_trans [OF *]) + using \<open>L>0\<close> e + apply (force simp: field_simps) + done + then have "cmod (winding_number p w - winding_number p z) < e" + using pi_ge_two e + by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) + } note cmod_wn_diff = this + then have "isCont (winding_number p) z" + apply (simp add: continuous_at_eps_delta, clarify) + apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) + using \<open>pe>0\<close> \<open>L>0\<close> + apply (simp add: dist_norm cmod_wn_diff) + done + then show ?thesis + apply (rule continuous_transform_within [where d = "min d e / 2"]) + apply (auto simp: \<open>d>0\<close> \<open>e>0\<close> dist_norm wnwn) + done +qed + +corollary continuous_on_winding_number: + "path \<gamma> \<Longrightarrow> continuous_on (- path_image \<gamma>) (\<lambda>w. winding_number \<gamma> w)" + by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) + + +subsection\<open>The winding number is constant on a connected region\<close> + +lemma winding_number_constant: + assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and cs: "connected s" and sg: "s \<inter> path_image \<gamma> = {}" + shows "\<exists>k. \<forall>z \<in> s. winding_number \<gamma> z = k" +proof - + { fix y z + assume ne: "winding_number \<gamma> y \<noteq> winding_number \<gamma> z" + assume "y \<in> s" "z \<in> s" + then have "winding_number \<gamma> y \<in> \<int>" "winding_number \<gamma> z \<in> \<int>" + using integer_winding_number [OF \<gamma> loop] sg \<open>y \<in> s\<close> by auto + with ne have "1 \<le> cmod (winding_number \<gamma> y - winding_number \<gamma> z)" + by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff) + } note * = this + show ?thesis + apply (rule continuous_discrete_range_constant [OF cs]) + using continuous_on_winding_number [OF \<gamma>] sg + apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset) + apply (rule_tac x=1 in exI) + apply (auto simp: *) + done +qed + +lemma winding_number_eq: + "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; w \<in> s; z \<in> s; connected s; s \<inter> path_image \<gamma> = {}\<rbrakk> + \<Longrightarrow> winding_number \<gamma> w = winding_number \<gamma> z" +using winding_number_constant by fastforce + +lemma open_winding_number_levelsets: + assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + shows "open {z. z \<notin> path_image \<gamma> \<and> winding_number \<gamma> z = k}" +proof - + have op: "open (- path_image \<gamma>)" + by (simp add: closed_path_image \<gamma> open_Compl) + { fix z assume z: "z \<notin> path_image \<gamma>" and k: "k = winding_number \<gamma> z" + obtain e where e: "e>0" "ball z e \<subseteq> - path_image \<gamma>" + using open_contains_ball [of "- path_image \<gamma>"] op z + by blast + have "\<exists>e>0. \<forall>y. dist y z < e \<longrightarrow> y \<notin> path_image \<gamma> \<and> winding_number \<gamma> y = winding_number \<gamma> z" + apply (rule_tac x=e in exI) + using e apply (simp add: dist_norm ball_def norm_minus_commute) + apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where s = "ball z e"]) + done + } then + show ?thesis + by (auto simp: open_dist) +qed + +subsection\<open>Winding number is zero "outside" a curve, in various senses\<close> + +lemma winding_number_zero_in_outside: + assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and z: "z \<in> outside (path_image \<gamma>)" + shows "winding_number \<gamma> z = 0" +proof - + obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto + obtain w::complex where w: "w \<notin> ball 0 (B + 1)" + by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) + have "- ball 0 (B + 1) \<subseteq> outside (path_image \<gamma>)" + apply (rule outside_subset_convex) + using B subset_ball by auto + then have wout: "w \<in> outside (path_image \<gamma>)" + using w by blast + moreover obtain k where "\<And>z. z \<in> outside (path_image \<gamma>) \<Longrightarrow> winding_number \<gamma> z = k" + using winding_number_constant [OF \<gamma> loop, of "outside(path_image \<gamma>)"] connected_outside + by (metis DIM_complex bounded_path_image dual_order.refl \<gamma> outside_no_overlap) + ultimately have "winding_number \<gamma> z = winding_number \<gamma> w" + using z by blast + also have "... = 0" + proof - + have wnot: "w \<notin> path_image \<gamma>" using wout by (simp add: outside_def) + { fix e::real assume "0<e" + obtain p where p: "polynomial_function p" "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>" + and pg1: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < 1)" + and pge: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < e)" + using path_approx_polynomial_function [OF \<gamma>, of "min 1 e"] \<open>e>0\<close> by force + have pip: "path_image p \<subseteq> ball 0 (B + 1)" + using B + apply (clarsimp simp add: path_image_def dist_norm ball_def) + apply (frule (1) pg1) + apply (fastforce dest: norm_add_less) + done + then have "w \<notin> path_image p" using w by blast + then have "\<exists>p. valid_path p \<and> w \<notin> path_image p \<and> + pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> + (\<forall>t\<in>{0..1}. cmod (\<gamma> t - p t) < e) \<and> contour_integral p (\<lambda>wa. 1 / (wa - w)) = 0" + apply (rule_tac x=p in exI) + apply (simp add: p valid_path_polynomial_function) + apply (intro conjI) + using pge apply (simp add: norm_minus_commute) + apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) + apply (rule holomorphic_intros | simp add: dist_norm)+ + using mem_ball_0 w apply blast + using p apply (simp_all add: valid_path_polynomial_function loop pip) + done + } + then show ?thesis + by (auto intro: winding_number_unique [OF \<gamma>] simp add: wnot) + qed + finally show ?thesis . +qed + +lemma winding_number_zero_outside: + "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> s; path_image \<gamma> \<subseteq> s\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0" + by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) + +lemma winding_number_zero_at_infinity: + assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + shows "\<exists>B. \<forall>z. B \<le> norm z \<longrightarrow> winding_number \<gamma> z = 0" +proof - + obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto + then show ?thesis + apply (rule_tac x="B+1" in exI, clarify) + apply (rule winding_number_zero_outside [OF \<gamma> convex_cball [of 0 B] loop]) + apply (meson less_add_one mem_cball_0 not_le order_trans) + using ball_subset_cball by blast +qed + +lemma winding_number_zero_point: + "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; open s; path_image \<gamma> \<subseteq> s\<rbrakk> + \<Longrightarrow> \<exists>z. z \<in> s \<and> winding_number \<gamma> z = 0" + using outside_compact_in_open [of "path_image \<gamma>" s] path_image_nonempty winding_number_zero_in_outside + by (fastforce simp add: compact_path_image) + + +text\<open>If a path winds round a set, it winds rounds its inside.\<close> +lemma winding_number_around_inside: + assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + and cls: "closed s" and cos: "connected s" and s_disj: "s \<inter> path_image \<gamma> = {}" + and z: "z \<in> s" and wn_nz: "winding_number \<gamma> z \<noteq> 0" and w: "w \<in> s \<union> inside s" + shows "winding_number \<gamma> w = winding_number \<gamma> z" +proof - + have ssb: "s \<subseteq> inside(path_image \<gamma>)" + proof + fix x :: complex + assume "x \<in> s" + hence "x \<notin> path_image \<gamma>" + by (meson disjoint_iff_not_equal s_disj) + thus "x \<in> inside (path_image \<gamma>)" + using \<open>x \<in> s\<close> by (metis (no_types) ComplI UnE cos \<gamma> loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) +qed + show ?thesis + apply (rule winding_number_eq [OF \<gamma> loop w]) + using z apply blast + apply (simp add: cls connected_with_inside cos) + apply (simp add: Int_Un_distrib2 s_disj, safe) + by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \<gamma>"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \<gamma> inside_no_overlap) + qed + + +text\<open>Bounding a WN by 1/2 for a path and point in opposite halfspaces.\<close> +lemma winding_number_subpath_continuous: + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" + shows "continuous_on {0..1} (\<lambda>x. winding_number(subpath 0 x \<gamma>) z)" +proof - + have *: "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) = + winding_number (subpath 0 x \<gamma>) z" + if x: "0 \<le> x" "x \<le> 1" for x + proof - + have "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) = + 1 / (2*pi*\<i>) * contour_integral (subpath 0 x \<gamma>) (\<lambda>w. 1/(w - z))" + using assms x + apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) + done + also have "... = winding_number (subpath 0 x \<gamma>) z" + apply (subst winding_number_valid_path) + using assms x + apply (simp_all add: path_image_subpath valid_path_subpath) + by (force simp: path_image_def) + finally show ?thesis . + qed + show ?thesis + apply (rule continuous_on_eq + [where f = "\<lambda>x. 1 / (2*pi*\<i>) * + integral {0..x} (\<lambda>t. 1/(\<gamma> t - z) * vector_derivative \<gamma> (at t))"]) + apply (rule continuous_intros)+ + apply (rule indefinite_integral_continuous) + apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) + using assms + apply (simp add: *) + done +qed + +lemma winding_number_ivt_pos: + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> Re(winding_number \<gamma> z)" + shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w" + apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right]) + apply (simp add:) + apply (rule winding_number_subpath_continuous [OF \<gamma> z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_neg: + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "Re(winding_number \<gamma> z) \<le> w" "w \<le> 0" + shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w" + apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right]) + apply (simp add:) + apply (rule winding_number_subpath_continuous [OF \<gamma> z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_abs: + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> \<bar>Re(winding_number \<gamma> z)\<bar>" + shows "\<exists>t \<in> {0..1}. \<bar>Re (winding_number (subpath 0 t \<gamma>) z)\<bar> = w" + using assms winding_number_ivt_pos [of \<gamma> z w] winding_number_ivt_neg [of \<gamma> z "-w"] + by force + +lemma winding_number_lt_half_lemma: + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and az: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}" + shows "Re(winding_number \<gamma> z) < 1/2" +proof - + { assume "Re(winding_number \<gamma> z) \<ge> 1/2" + then obtain t::real where t: "0 \<le> t" "t \<le> 1" and sub12: "Re (winding_number (subpath 0 t \<gamma>) z) = 1/2" + using winding_number_ivt_pos [OF \<gamma> z, of "1/2"] by auto + have gt: "\<gamma> t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \<gamma>) z)))) * (\<gamma> 0 - z))" + using winding_number_exp_2pi [of "subpath 0 t \<gamma>" z] + apply (simp add: t \<gamma> valid_path_imp_path) + using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) + have "b < a \<bullet> \<gamma> 0" + proof - + have "\<gamma> 0 \<in> {c. b < a \<bullet> c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) + thus ?thesis + by blast + qed + moreover have "b < a \<bullet> \<gamma> t" + proof - + have "\<gamma> t \<in> {c. b < a \<bullet> c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) + thus ?thesis + by blast + qed + ultimately have "0 < a \<bullet> (\<gamma> 0 - z)" "0 < a \<bullet> (\<gamma> t - z)" using az + by (simp add: inner_diff_right)+ + then have False + by (simp add: gt inner_mult_right mult_less_0_iff) + } + then show ?thesis by force +qed + +lemma winding_number_lt_half: + assumes "valid_path \<gamma>" "a \<bullet> z \<le> b" "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}" + shows "\<bar>Re (winding_number \<gamma> z)\<bar> < 1/2" +proof - + have "z \<notin> path_image \<gamma>" using assms by auto + with assms show ?thesis + apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) + apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \<gamma> z a b] + winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) + done +qed + +lemma winding_number_le_half: + assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" + and anz: "a \<noteq> 0" and azb: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w \<ge> b}" + shows "\<bar>Re (winding_number \<gamma> z)\<bar> \<le> 1/2" +proof - + { assume wnz_12: "\<bar>Re (winding_number \<gamma> z)\<bar> > 1/2" + have "isCont (winding_number \<gamma>) z" + by (metis continuous_at_winding_number valid_path_imp_path \<gamma> z) + then obtain d where "d>0" and d: "\<And>x'. dist x' z < d \<Longrightarrow> dist (winding_number \<gamma> x') (winding_number \<gamma> z) < \<bar>Re(winding_number \<gamma> z)\<bar> - 1/2" + using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast + define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" + have *: "a \<bullet> z' \<le> b - d / 3 * cmod a" + unfolding z'_def inner_mult_right' divide_inverse + apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz) + apply (metis \<open>0 < d\<close> add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) + done + have "cmod (winding_number \<gamma> z' - winding_number \<gamma> z) < \<bar>Re (winding_number \<gamma> z)\<bar> - 1/2" + using d [of z'] anz \<open>d>0\<close> by (simp add: dist_norm z'_def) + then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - cmod (winding_number \<gamma> z' - winding_number \<gamma> z)" + by simp + then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - \<bar>Re (winding_number \<gamma> z') - Re (winding_number \<gamma> z)\<bar>" + using abs_Re_le_cmod [of "winding_number \<gamma> z' - winding_number \<gamma> z"] by simp + then have wnz_12': "\<bar>Re (winding_number \<gamma> z')\<bar> > 1/2" + by linarith + moreover have "\<bar>Re (winding_number \<gamma> z')\<bar> < 1/2" + apply (rule winding_number_lt_half [OF \<gamma> *]) + using azb \<open>d>0\<close> pag + apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD) + done + ultimately have False + by simp + } + then show ?thesis by force +qed + +lemma winding_number_lt_half_linepath: "z \<notin> closed_segment a b \<Longrightarrow> \<bar>Re (winding_number (linepath a b) z)\<bar> < 1/2" + using separating_hyperplane_closed_point [of "closed_segment a b" z] + apply auto + apply (simp add: closed_segment_def) + apply (drule less_imp_le) + apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) + apply (auto simp: segment) + done + + +text\<open> Positivity of WN for a linepath.\<close> +lemma winding_number_linepath_pos_lt: + assumes "0 < Im ((b - a) * cnj (b - z))" + shows "0 < Re(winding_number(linepath a b) z)" +proof - + have z: "z \<notin> path_image (linepath a b)" + using assms + by (simp add: closed_segment_def) (force simp: algebra_simps) + show ?thesis + apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) + apply (simp add: linepath_def algebra_simps) + done +qed + + +subsection\<open>Cauchy's integral formula, again for a convex enclosing set.\<close> + +lemma Cauchy_integral_formula_weak: + assumes s: "convex s" and "finite k" and conf: "continuous_on s f" + and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)" + and z: "z \<in> interior s - k" and vpg: "valid_path \<gamma>" + and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>" +proof - + obtain f' where f': "(f has_field_derivative f') (at z)" + using fcd [OF z] by (auto simp: field_differentiable_def) + have pas: "path_image \<gamma> \<subseteq> s" and znotin: "z \<notin> path_image \<gamma>" using pasz by blast+ + have c: "continuous (at x within s) (\<lambda>w. if w = z then f' else (f w - f z) / (w - z))" if "x \<in> s" for x + proof (cases "x = z") + case True then show ?thesis + apply (simp add: continuous_within) + apply (rule Lim_transform_away_within [of _ "z+1" _ "\<lambda>w::complex. (f w - f z)/(w - z)"]) + using has_field_derivative_at_within DERIV_within_iff f' + apply (fastforce simp add:)+ + done + next + case False + then have dxz: "dist x z > 0" by auto + have cf: "continuous (at x within s) f" + using conf continuous_on_eq_continuous_within that by blast + have "continuous (at x within s) (\<lambda>w. (f w - f z) / (w - z))" + by (rule cf continuous_intros | simp add: False)+ + then show ?thesis + apply (rule continuous_transform_within [OF _ dxz that, of "\<lambda>w::complex. (f w - f z)/(w - z)"]) + apply (force simp: dist_commute) + done + qed + have fink': "finite (insert z k)" using \<open>finite k\<close> by blast + have *: "((\<lambda>w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>" + apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) + using c apply (force simp: continuous_on_eq_continuous_within) + apply (rename_tac w) + apply (rule_tac d="dist w z" and f = "\<lambda>w. (f w - f z)/(w - z)" in field_differentiable_transform_within) + apply (simp_all add: dist_pos_lt dist_commute) + apply (metis less_irrefl) + apply (rule derivative_intros fcd | simp)+ + done + show ?thesis + apply (rule has_contour_integral_eq) + using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] + apply (auto simp: mult_ac divide_simps) + done +qed + +theorem Cauchy_integral_formula_convex_simple: + "\<lbrakk>convex s; f holomorphic_on s; z \<in> interior s; valid_path \<gamma>; path_image \<gamma> \<subseteq> s - {z}; + pathfinish \<gamma> = pathstart \<gamma>\<rbrakk> + \<Longrightarrow> ((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>" + apply (rule Cauchy_integral_formula_weak [where k = "{}"]) + using holomorphic_on_imp_continuous_on + by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) + + +subsection\<open>Homotopy forms of Cauchy's theorem\<close> + +proposition Cauchy_theorem_homotopic: + assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" +proof - + have pathsf: "linked_paths atends g h" + using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) + obtain k :: "real \<times> real \<Rightarrow> complex" + where contk: "continuous_on ({0..1} \<times> {0..1}) k" + and ks: "k ` ({0..1} \<times> {0..1}) \<subseteq> s" + and k [simp]: "\<forall>x. k (0, x) = g x" "\<forall>x. k (1, x) = h x" + and ksf: "\<forall>t\<in>{0..1}. linked_paths atends g (\<lambda>x. k (t, x))" + using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) + have ucontk: "uniformly_continuous_on ({0..1} \<times> {0..1}) k" + by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) + { fix t::real assume t: "t \<in> {0..1}" + have pak: "path (k o (\<lambda>u. (t, u)))" + unfolding path_def + apply (rule continuous_intros continuous_on_subset [OF contk])+ + using t by force + have pik: "path_image (k \<circ> Pair t) \<subseteq> s" + using ks t by (auto simp: path_image_def) + obtain e where "e>0" and e: + "\<And>g h. \<lbrakk>valid_path g; valid_path h; + \<forall>u\<in>{0..1}. cmod (g u - (k \<circ> Pair t) u) < e \<and> cmod (h u - (k \<circ> Pair t) u) < e; + linked_paths atends g h\<rbrakk> + \<Longrightarrow> contour_integral h f = contour_integral g f" + using contour_integral_nearby [OF \<open>open s\<close> pak pik, of atends] f by metis + obtain d where "d>0" and d: + "\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> {0..1}; x' \<in> {0..1} \<times> {0..1}; norm (x'-x) < d\<rbrakk> \<Longrightarrow> norm (k x' - k x) < e/4" + by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \<open>e>0\<close>) + { fix t1 t2 + assume t1: "0 \<le> t1" "t1 \<le> 1" and t2: "0 \<le> t2" "t2 \<le> 1" and ltd: "\<bar>t1 - t\<bar> < d" "\<bar>t2 - t\<bar> < d" + have no2: "\<And>g1 k1 kt. \<lbrakk>norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\<rbrakk> \<Longrightarrow> norm(g1 - kt) < e" + using \<open>e > 0\<close> + apply (rule_tac y = k1 in norm_triangle_half_l) + apply (auto simp: norm_minus_commute intro: order_less_trans) + done + have "\<exists>d>0. \<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and> + (\<forall>u\<in>{0..1}. cmod (g1 u - k (t1, u)) < d \<and> cmod (g2 u - k (t2, u)) < d) \<and> + linked_paths atends g1 g2 \<longrightarrow> + contour_integral g2 f = contour_integral g1 f" + apply (rule_tac x="e/4" in exI) + using t t1 t2 ltd \<open>e > 0\<close> + apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) + done + } + then have "\<exists>e. 0 < e \<and> + (\<forall>t1 t2. t1 \<in> {0..1} \<and> t2 \<in> {0..1} \<and> \<bar>t1 - t\<bar> < e \<and> \<bar>t2 - t\<bar> < e + \<longrightarrow> (\<exists>d. 0 < d \<and> + (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and> + (\<forall>u \<in> {0..1}. + norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and> + linked_paths atends g1 g2 + \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))" + by (rule_tac x=d in exI) (simp add: \<open>d > 0\<close>) + } + then obtain ee where ee: + "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0 \<and> + (\<forall>t1 t2. t1 \<in> {0..1} \<longrightarrow> t2 \<in> {0..1} \<longrightarrow> \<bar>t1 - t\<bar> < ee t \<longrightarrow> \<bar>t2 - t\<bar> < ee t + \<longrightarrow> (\<exists>d. 0 < d \<and> + (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and> + (\<forall>u \<in> {0..1}. + norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and> + linked_paths atends g1 g2 + \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))" + by metis + note ee_rule = ee [THEN conjunct2, rule_format] + define C where "C = (\<lambda>t. ball t (ee t / 3)) ` {0..1}" + have "\<forall>t \<in> C. open t" by (simp add: C_def) + moreover have "{0..1} \<subseteq> \<Union>C" + using ee [THEN conjunct1] by (auto simp: C_def dist_norm) + ultimately obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'" + by (rule compactE [OF compact_interval]) + define kk where "kk = {t \<in> {0..1}. ball t (ee t / 3) \<in> C'}" + have kk01: "kk \<subseteq> {0..1}" by (auto simp: kk_def) + define e where "e = Min (ee ` kk)" + have C'_eq: "C' = (\<lambda>t. ball t (ee t / 3)) ` kk" + using C' by (auto simp: kk_def C_def) + have ee_pos[simp]: "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0" + by (simp add: kk_def ee) + moreover have "finite kk" + using \<open>finite C'\<close> kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) + moreover have "kk \<noteq> {}" using \<open>{0..1} \<subseteq> \<Union>C'\<close> C'_eq by force + ultimately have "e > 0" + using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) + then obtain N::nat where "N > 0" and N: "1/N < e/3" + by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) + have e_le_ee: "\<And>i. i \<in> kk \<Longrightarrow> e \<le> ee i" + using \<open>finite kk\<close> by (simp add: e_def Min_le_iff [of "ee ` kk"]) + have plus: "\<exists>t \<in> kk. x \<in> ball t (ee t / 3)" if "x \<in> {0..1}" for x + using C' subsetD [OF C'01 that] unfolding C'_eq by blast + have [OF order_refl]: + "\<exists>d. 0 < d \<and> (\<forall>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (n/N, u)) < d) \<and> linked_paths atends g j + \<longrightarrow> contour_integral j f = contour_integral g f)" + if "n \<le> N" for n + using that + proof (induct n) + case 0 show ?case using ee_rule [of 0 0 0] + apply clarsimp + apply (rule_tac x=d in exI, safe) + by (metis diff_self vpg norm_zero) + next + case (Suc n) + then have N01: "n/N \<in> {0..1}" "(Suc n)/N \<in> {0..1}" by auto + then obtain t where t: "t \<in> kk" "n/N \<in> ball t (ee t / 3)" + using plus [of "n/N"] by blast + then have nN_less: "\<bar>n/N - t\<bar> < ee t" + by (simp add: dist_norm del: less_divide_eq_numeral1) + have n'N_less: "\<bar>real (Suc n) / real N - t\<bar> < ee t" + using t N \<open>N > 0\<close> e_le_ee [of t] + by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) + have t01: "t \<in> {0..1}" using \<open>kk \<subseteq> {0..1}\<close> \<open>t \<in> kk\<close> by blast + obtain d1 where "d1 > 0" and d1: + "\<And>g1 g2. \<lbrakk>valid_path g1; valid_path g2; + \<forall>u\<in>{0..1}. cmod (g1 u - k (n/N, u)) < d1 \<and> cmod (g2 u - k ((Suc n) / N, u)) < d1; + linked_paths atends g1 g2\<rbrakk> + \<Longrightarrow> contour_integral g2 f = contour_integral g1 f" + using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce + have "n \<le> N" using Suc.prems by auto + with Suc.hyps + obtain d2 where "d2 > 0" + and d2: "\<And>j. \<lbrakk>valid_path j; \<forall>u\<in>{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\<rbrakk> + \<Longrightarrow> contour_integral j f = contour_integral g f" + by auto + have "continuous_on {0..1} (k o (\<lambda>u. (n/N, u)))" + apply (rule continuous_intros continuous_on_subset [OF contk])+ + using N01 by auto + then have pkn: "path (\<lambda>u. k (n/N, u))" + by (simp add: path_def) + have min12: "min d1 d2 > 0" by (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>) + obtain p where "polynomial_function p" + and psf: "pathstart p = pathstart (\<lambda>u. k (n/N, u))" + "pathfinish p = pathfinish (\<lambda>u. k (n/N, u))" + and pk_le: "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (p t - k (n/N, t)) < min d1 d2" + using path_approx_polynomial_function [OF pkn min12] by blast + then have vpp: "valid_path p" using valid_path_polynomial_function by blast + have lpa: "linked_paths atends g p" + by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) + show ?case + apply (rule_tac x="min d1 d2" in exI) + apply (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>, clarify) + apply (rule_tac s="contour_integral p f" in trans) + using pk_le N01(1) ksf pathfinish_def pathstart_def + apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf) + using pk_le N01 apply (force intro!: vpp d2 lpa simp add: linked_paths_def psf ksf) + done + qed + then obtain d where "0 < d" + "\<And>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (1,u)) < d) \<and> + linked_paths atends g j + \<Longrightarrow> contour_integral j f = contour_integral g f" + using \<open>N>0\<close> by auto + then have "linked_paths atends g h \<Longrightarrow> contour_integral h f = contour_integral g f" + using \<open>N>0\<close> vph by fastforce + then show ?thesis + by (simp add: pathsf) +qed + +proposition Cauchy_theorem_homotopic_paths: + assumes hom: "homotopic_paths s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" + using Cauchy_theorem_homotopic [of True s g h] assms by simp + +proposition Cauchy_theorem_homotopic_loops: + assumes hom: "homotopic_loops s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" + using Cauchy_theorem_homotopic [of False s g h] assms by simp + +lemma has_contour_integral_newpath: + "\<lbrakk>(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\<rbrakk> + \<Longrightarrow> (f has_contour_integral y) g" + using has_contour_integral_integral contour_integral_unique by auto + +lemma Cauchy_theorem_null_homotopic: + "\<lbrakk>f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g" + apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) + using contour_integrable_holomorphic_simple + apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) + by (simp add: Cauchy_theorem_homotopic_loops) + + + +subsection\<open>More winding number properties\<close> + +text\<open>including the fact that it's +-1 inside a simple closed curve.\<close> + +lemma winding_number_homotopic_paths: + assumes "homotopic_paths (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto + moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h" + using homotopic_paths_imp_subset [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\<And>p. \<lbrakk>path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk> + \<Longrightarrow> homotopic_paths (-{z}) g p" + and e: "\<And>q. \<lbrakk>path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk> + \<Longrightarrow> homotopic_paths (-{z}) h q" + using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \<notin> path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z" + using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast + obtain q where q: + "valid_path q" "z \<notin> path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z" + using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast + have gp: "homotopic_paths (- {z}) g p" + by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) + have hq: "homotopic_paths (- {z}) h q" + by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) + have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))" + apply (rule Cauchy_theorem_homotopic_paths [of "-{z}"]) + apply (blast intro: homotopic_paths_trans homotopic_paths_sym gp hq assms) + apply (auto intro!: holomorphic_intros simp: p q) + done + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_homotopic_loops: + assumes "homotopic_loops (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto + moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h" + using homotopic_loops_imp_subset [OF assms] by auto + moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" + using homotopic_loops_imp_loop [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\<And>p. \<lbrakk>path p; pathfinish p = pathstart p; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk> + \<Longrightarrow> homotopic_loops (-{z}) g p" + and e: "\<And>q. \<lbrakk>path q; pathfinish q = pathstart q; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk> + \<Longrightarrow> homotopic_loops (-{z}) h q" + using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \<notin> path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z" + using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast + obtain q where q: + "valid_path q" "z \<notin> path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z" + using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast + have gp: "homotopic_loops (- {z}) g p" + by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) + have hq: "homotopic_loops (- {z}) h q" + by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) + have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))" + apply (rule Cauchy_theorem_homotopic_loops [of "-{z}"]) + apply (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) + apply (auto intro!: holomorphic_intros simp: p q) + done + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_paths_linear_eq: + "\<lbrakk>path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; + \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk> + \<Longrightarrow> winding_number h z = winding_number g z" + by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths elim: ) + +lemma winding_number_loops_linear_eq: + "\<lbrakk>path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; + \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk> + \<Longrightarrow> winding_number h z = winding_number g z" + by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops elim: ) + +lemma winding_number_nearby_paths_eq: + "\<lbrakk>path g; path h; + pathstart h = pathstart g; pathfinish h = pathfinish g; + \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk> + \<Longrightarrow> winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) + +lemma winding_number_nearby_loops_eq: + "\<lbrakk>path g; path h; + pathfinish g = pathstart g; + pathfinish h = pathstart h; + \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk> + \<Longrightarrow> winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) + + +proposition winding_number_subpath_combine: + "\<lbrakk>path g; z \<notin> path_image g; + u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk> + \<Longrightarrow> winding_number (subpath u v g) z + winding_number (subpath v w g) z = + winding_number (subpath u w g) z" +apply (rule trans [OF winding_number_join [THEN sym] + winding_number_homotopic_paths [OF homotopic_join_subpaths]]) +apply (auto dest: path_image_subpath_subset) +done + + +subsection\<open>Partial circle path\<close> + +definition part_circlepath :: "[complex, real, real, real, real] \<Rightarrow> complex" + where "part_circlepath z r s t \<equiv> \<lambda>x. z + of_real r * exp (\<i> * of_real (linepath s t x))" + +lemma pathstart_part_circlepath [simp]: + "pathstart(part_circlepath z r s t) = z + r*exp(\<i> * s)" +by (metis part_circlepath_def pathstart_def pathstart_linepath) + +lemma pathfinish_part_circlepath [simp]: + "pathfinish(part_circlepath z r s t) = z + r*exp(\<i>*t)" +by (metis part_circlepath_def pathfinish_def pathfinish_linepath) + +proposition has_vector_derivative_part_circlepath [derivative_intros]: + "((part_circlepath z r s t) has_vector_derivative + (\<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x))) + (at x within X)" + apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) + apply (rule has_vector_derivative_real_complex) + apply (rule derivative_eq_intros | simp)+ + done + +corollary vector_derivative_part_circlepath: + "vector_derivative (part_circlepath z r s t) (at x) = + \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)" + using has_vector_derivative_part_circlepath vector_derivative_at by blast + +corollary vector_derivative_part_circlepath01: + "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> + \<Longrightarrow> vector_derivative (part_circlepath z r s t) (at x within {0..1}) = + \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)" + using has_vector_derivative_part_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" + apply (simp add: valid_path_def) + apply (rule C1_differentiable_imp_piecewise) + apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath + intro!: continuous_intros) + done + +lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" + by (simp add: valid_path_imp_path) + +proposition path_image_part_circlepath: + assumes "s \<le> t" + shows "path_image (part_circlepath z r s t) = {z + r * exp(\<i> * of_real x) | x. s \<le> x \<and> x \<le> t}" +proof - + { fix z::real + assume "0 \<le> z" "z \<le> 1" + with \<open>s \<le> t\<close> have "\<exists>x. (exp (\<i> * linepath s t z) = exp (\<i> * of_real x)) \<and> s \<le> x \<and> x \<le> t" + apply (rule_tac x="(1 - z) * s + z * t" in exI) + apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) + apply (rule conjI) + using mult_right_mono apply blast + using affine_ineq by (metis "mult.commute") + } + moreover + { fix z + assume "s \<le> z" "z \<le> t" + then have "z + of_real r * exp (\<i> * of_real z) \<in> (\<lambda>x. z + of_real r * exp (\<i> * linepath s t x)) ` {0..1}" + apply (rule_tac x="(z - s)/(t - s)" in image_eqI) + apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) + apply (auto simp: algebra_simps divide_simps) + done + } + ultimately show ?thesis + by (fastforce simp add: path_image_def part_circlepath_def) +qed + +corollary path_image_part_circlepath_subset: + "\<lbrakk>s \<le> t; 0 \<le> r\<rbrakk> \<Longrightarrow> path_image(part_circlepath z r s t) \<subseteq> sphere z r" +by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) + +proposition in_path_image_part_circlepath: + assumes "w \<in> path_image(part_circlepath z r s t)" "s \<le> t" "0 \<le> r" + shows "norm(w - z) = r" +proof - + have "w \<in> {c. dist z c = r}" + by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) + thus ?thesis + by (simp add: dist_norm norm_minus_commute) +qed + +proposition finite_bounded_log: "finite {z::complex. norm z \<le> b \<and> exp z = w}" +proof (cases "w = 0") + case True then show ?thesis by auto +next + case False + have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \<i>) \<le> b + cmod (Ln w)}" + apply (simp add: norm_mult finite_int_iff_bounded_le) + apply (rule_tac x="\<lfloor>(b + cmod (Ln w)) / (2*pi)\<rfloor>" in exI) + apply (auto simp: divide_simps le_floor_iff) + done + have [simp]: "\<And>P f. {z. P z \<and> (\<exists>n. z = f n)} = f ` {n. P (f n)}" + by blast + show ?thesis + apply (subst exp_Ln [OF False, symmetric]) + apply (simp add: exp_eq) + using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) + done +qed + +lemma finite_bounded_log2: + fixes a::complex + assumes "a \<noteq> 0" + shows "finite {z. norm z \<le> b \<and> exp(a*z) = w}" +proof - + have *: "finite ((\<lambda>z. z / a) ` {z. cmod z \<le> b * cmod a \<and> exp z = w})" + by (rule finite_imageI [OF finite_bounded_log]) + show ?thesis + by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) +qed + +proposition has_contour_integral_bound_part_circlepath_strong: + assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" + and "finite k" and le: "0 \<le> B" "0 < r" "s \<le> t" + and B: "\<And>x. x \<in> path_image(part_circlepath z r s t) - k \<Longrightarrow> norm(f x) \<le> B" + shows "cmod i \<le> B * r * (t - s)" +proof - + consider "s = t" | "s < t" using \<open>s \<le> t\<close> by linarith + then show ?thesis + proof cases + case 1 with fi [unfolded has_contour_integral] + have "i = 0" by (simp add: vector_derivative_part_circlepath) + with assms show ?thesis by simp + next + case 2 + have [simp]: "\<bar>r\<bar> = r" using \<open>r > 0\<close> by linarith + have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" + by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) + have "finite (part_circlepath z r s t -` {y} \<inter> {0..1})" if "y \<in> k" for y + proof - + define w where "w = (y - z)/of_real r / exp(\<i> * of_real s)" + have fin: "finite (of_real -` {z. cmod z \<le> 1 \<and> exp (\<i> * complex_of_real (t - s) * z) = w})" + apply (rule finite_vimageI [OF finite_bounded_log2]) + using \<open>s < t\<close> apply (auto simp: inj_of_real) + done + show ?thesis + apply (simp add: part_circlepath_def linepath_def vimage_def) + apply (rule finite_subset [OF _ fin]) + using le + apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) + done + qed + then have fin01: "finite ((part_circlepath z r s t) -` k \<inter> {0..1})" + by (rule finite_finite_vimage_IntI [OF \<open>finite k\<close>]) + have **: "((\<lambda>x. if (part_circlepath z r s t x) \<in> k then 0 + else f(part_circlepath z r s t x) * + vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" + apply (rule has_integral_spike + [where f = "\<lambda>x. f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)"]) + apply (rule negligible_finite [OF fin01]) + using fi has_contour_integral + apply auto + done + have *: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1; part_circlepath z r s t x \<notin> k\<rbrakk> \<Longrightarrow> cmod (f (part_circlepath z r s t x)) \<le> B" + by (auto intro!: B [unfolded path_image_def image_def, simplified]) + show ?thesis + apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) + using assms apply force + apply (simp add: norm_mult vector_derivative_part_circlepath) + using le * "2" \<open>r > 0\<close> by auto + qed +qed + +corollary has_contour_integral_bound_part_circlepath: + "\<lbrakk>(f has_contour_integral i) (part_circlepath z r s t); + 0 \<le> B; 0 < r; s \<le> t; + \<And>x. x \<in> path_image(part_circlepath z r s t) \<Longrightarrow> norm(f x) \<le> B\<rbrakk> + \<Longrightarrow> norm i \<le> B*r*(t - s)" + by (auto intro: has_contour_integral_bound_part_circlepath_strong) + +proposition contour_integrable_continuous_part_circlepath: + "continuous_on (path_image (part_circlepath z r s t)) f + \<Longrightarrow> f contour_integrable_on (part_circlepath z r s t)" + apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) + apply (rule integrable_continuous_real) + apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) + done + +proposition winding_number_part_circlepath_pos_less: + assumes "s < t" and no: "norm(w - z) < r" + shows "0 < Re (winding_number(part_circlepath z r s t) w)" +proof - + have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) + note valid_path_part_circlepath + moreover have " w \<notin> path_image (part_circlepath z r s t)" + using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) + moreover have "0 < r * (t - s) * (r - cmod (w - z))" + using assms by (metis \<open>0 < r\<close> diff_gt_0_iff_gt mult_pos_pos) + ultimately show ?thesis + apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) + apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) + apply (rule mult_left_mono)+ + using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] + apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) + using assms \<open>0 < r\<close> by auto +qed + +proposition simple_path_part_circlepath: + "simple_path(part_circlepath z r s t) \<longleftrightarrow> (r \<noteq> 0 \<and> s \<noteq> t \<and> \<bar>s - t\<bar> \<le> 2*pi)" +proof (cases "r = 0 \<or> s = t") + case True + then show ?thesis + apply (rule disjE) + apply (force simp: part_circlepath_def simple_path_def intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ + done +next + case False then have "r \<noteq> 0" "s \<noteq> t" by auto + have *: "\<And>x y z s t. \<i>*((1 - x) * s + x * t) = \<i>*(((1 - y) * s + y * t)) + z \<longleftrightarrow> \<i>*(x - y) * (t - s) = z" + by (simp add: algebra_simps) + have abs01: "\<And>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 + \<Longrightarrow> (x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0 \<longleftrightarrow> \<bar>x - y\<bar> \<in> {0,1})" + by auto + have abs_away: "\<And>P. (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. P \<bar>x - y\<bar>) \<longleftrightarrow> (\<forall>x::real. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> P x)" + by force + have **: "\<And>x y. (\<exists>n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \<longleftrightarrow> + (\<exists>n. \<bar>x - y\<bar> * (t - s) = 2 * (of_int n * pi))" + by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] + intro: exI [where x = "-n" for n]) + have 1: "\<forall>x. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> (\<exists>n. x * (t - s) = 2 * (real_of_int n * pi)) \<longrightarrow> x = 0 \<or> x = 1 \<Longrightarrow> \<bar>s - t\<bar> \<le> 2 * pi" + apply (rule ccontr) + apply (drule_tac x="2*pi / \<bar>t - s\<bar>" in spec) + using False + apply (simp add: abs_minus_commute divide_simps) + apply (frule_tac x=1 in spec) + apply (drule_tac x="-1" in spec) + apply (simp add:) + done + have 2: "\<bar>s - t\<bar> = \<bar>2 * (real_of_int n * pi) / x\<bar>" if "x \<noteq> 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n + proof - + have "t-s = 2 * (real_of_int n * pi)/x" + using that by (simp add: field_simps) + then show ?thesis by (metis abs_minus_commute) + qed + show ?thesis using False + apply (simp add: simple_path_def path_part_circlepath) + apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) + apply (subst abs_away) + apply (auto simp: 1) + apply (rule ccontr) + apply (auto simp: 2 divide_simps abs_mult dest: of_int_leD) + done +qed + +proposition arc_part_circlepath: + assumes "r \<noteq> 0" "s \<noteq> t" "\<bar>s - t\<bar> < 2*pi" + shows "arc (part_circlepath z r s t)" +proof - + have *: "x = y" if eq: "\<i> * (linepath s t x) = \<i> * (linepath s t y) + 2 * of_int n * complex_of_real pi * \<i>" + and x: "x \<in> {0..1}" and y: "y \<in> {0..1}" for x y n + proof - + have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" + by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_divide_cancel_left eq) + then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" + by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) + then have st: "x \<noteq> y \<Longrightarrow> (s-t) = (of_int n * (pi * 2) / (y-x))" + by (force simp: field_simps) + show ?thesis + apply (rule ccontr) + using assms x y + apply (simp add: st abs_mult field_simps) + using st + apply (auto simp: dest: of_int_lessD) + done + qed + show ?thesis + using assms + apply (simp add: arc_def) + apply (simp add: part_circlepath_def inj_on_def exp_eq) + apply (blast intro: *) + done +qed + + +subsection\<open>Special case of one complete circle\<close> + +definition circlepath :: "[complex, real, real] \<Rightarrow> complex" + where "circlepath z r \<equiv> part_circlepath z r 0 (2*pi)" + +lemma circlepath: "circlepath z r = (\<lambda>x. z + r * exp(2 * of_real pi * \<i> * of_real x))" + by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) + +lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" + by (simp add: circlepath_def) + +lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" + by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) + +lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" +proof - + have "z + of_real r * exp (2 * pi * \<i> * (x + 1 / 2)) = + z + of_real r * exp (2 * pi * \<i> * x + pi * \<i>)" + by (simp add: divide_simps) (simp add: algebra_simps) + also have "... = z - r * exp (2 * pi * \<i> * x)" + by (simp add: exp_add) + finally show ?thesis + by (simp add: circlepath path_image_def sphere_def dist_norm) +qed + +lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" + using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] + by (simp add: add.commute) + +lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" + using circlepath_add1 [of z r "x-1/2"] + by (simp add: add.commute) + +lemma path_image_circlepath_minus_subset: + "path_image (circlepath z (-r)) \<subseteq> path_image (circlepath z r)" + apply (simp add: path_image_def image_def circlepath_minus, clarify) + apply (case_tac "xa \<le> 1/2", force) + apply (force simp add: circlepath_add_half)+ + done + +lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" + using path_image_circlepath_minus_subset by fastforce + +proposition has_vector_derivative_circlepath [derivative_intros]: + "((circlepath z r) has_vector_derivative (2 * pi * \<i> * r * exp (2 * of_real pi * \<i> * of_real x))) + (at x within X)" + apply (simp add: circlepath_def scaleR_conv_of_real) + apply (rule derivative_eq_intros) + apply (simp add: algebra_simps) + done + +corollary vector_derivative_circlepath: + "vector_derivative (circlepath z r) (at x) = + 2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)" +using has_vector_derivative_circlepath vector_derivative_at by blast + +corollary vector_derivative_circlepath01: + "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> + \<Longrightarrow> vector_derivative (circlepath z r) (at x within {0..1}) = + 2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)" + using has_vector_derivative_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" + by (simp add: circlepath_def) + +lemma path_circlepath [simp]: "path (circlepath z r)" + by (simp add: valid_path_imp_path) + +lemma path_image_circlepath_nonneg: + assumes "0 \<le> r" shows "path_image (circlepath z r) = sphere z r" +proof - + have *: "x \<in> (\<lambda>u. z + (cmod (x - z)) * exp (\<i> * (of_real u * (of_real pi * 2)))) ` {0..1}" for x + proof (cases "x = z") + case True then show ?thesis by force + next + case False + define w where "w = x - z" + then have "w \<noteq> 0" by (simp add: False) + have **: "\<And>t. \<lbrakk>Re w = cos t * cmod w; Im w = sin t * cmod w\<rbrakk> \<Longrightarrow> w = of_real (cmod w) * exp (\<i> * t)" + using cis_conv_exp complex_eq_iff by auto + show ?thesis + apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) + apply (simp add: divide_simps \<open>w \<noteq> 0\<close> cmod_power2 [symmetric]) + apply (rule_tac x="t / (2*pi)" in image_eqI) + apply (simp add: divide_simps \<open>w \<noteq> 0\<close>) + using False ** + apply (auto simp: w_def) + done + qed + show ?thesis + unfolding circlepath path_image_def sphere_def dist_norm + by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) +qed + +proposition path_image_circlepath [simp]: + "path_image (circlepath z r) = sphere z \<bar>r\<bar>" + using path_image_circlepath_minus + by (force simp add: path_image_circlepath_nonneg abs_if) + +lemma has_contour_integral_bound_circlepath_strong: + "\<lbrakk>(f has_contour_integral i) (circlepath z r); + finite k; 0 \<le> B; 0 < r; + \<And>x. \<lbrakk>norm(x - z) = r; x \<notin> k\<rbrakk> \<Longrightarrow> norm(f x) \<le> B\<rbrakk> + \<Longrightarrow> norm i \<le> B*(2*pi*r)" + unfolding circlepath_def + by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) + +corollary has_contour_integral_bound_circlepath: + "\<lbrakk>(f has_contour_integral i) (circlepath z r); + 0 \<le> B; 0 < r; \<And>x. norm(x - z) = r \<Longrightarrow> norm(f x) \<le> B\<rbrakk> + \<Longrightarrow> norm i \<le> B*(2*pi*r)" + by (auto intro: has_contour_integral_bound_circlepath_strong) + +proposition contour_integrable_continuous_circlepath: + "continuous_on (path_image (circlepath z r)) f + \<Longrightarrow> f contour_integrable_on (circlepath z r)" + by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) + +lemma simple_path_circlepath: "simple_path(circlepath z r) \<longleftrightarrow> (r \<noteq> 0)" + by (simp add: circlepath_def simple_path_part_circlepath) + +lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \<Longrightarrow> w \<notin> path_image (circlepath z r)" + by (simp add: sphere_def dist_norm norm_minus_commute) + +proposition contour_integral_circlepath: + "0 < r \<Longrightarrow> contour_integral (circlepath z r) (\<lambda>w. 1 / (w - z)) = 2 * complex_of_real pi * \<i>" + apply (rule contour_integral_unique) + apply (simp add: has_contour_integral_def) + apply (subst has_integral_cong) + apply (simp add: vector_derivative_circlepath01) + using has_integral_const_real [of _ 0 1] + apply (force simp: circlepath) + done + +lemma winding_number_circlepath_centre: "0 < r \<Longrightarrow> winding_number (circlepath z r) z = 1" + apply (rule winding_number_unique_loop) + apply (simp_all add: sphere_def valid_path_imp_path) + apply (rule_tac x="circlepath z r" in exI) + apply (simp add: sphere_def contour_integral_circlepath) + done + +proposition winding_number_circlepath: + assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" +proof (cases "w = z") + case True then show ?thesis + using assms winding_number_circlepath_centre by auto +next + case False + have [simp]: "r > 0" + using assms le_less_trans norm_ge_zero by blast + define r' where "r' = norm(w - z)" + have "r' < r" + by (simp add: assms r'_def) + have disjo: "cball z r' \<inter> sphere z r = {}" + using \<open>r' < r\<close> by (force simp: cball_def sphere_def) + have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" + apply (rule winding_number_around_inside [where s = "cball z r'"]) + apply (simp_all add: disjo order.strict_implies_order winding_number_circlepath_centre) + apply (simp_all add: False r'_def dist_norm norm_minus_commute) + done + also have "... = 1" + by (simp add: winding_number_circlepath_centre) + finally show ?thesis . +qed + + +text\<open> Hence the Cauchy formula for points inside a circle.\<close> + +theorem Cauchy_integral_circlepath: + assumes "continuous_on (cball z r) f" "f holomorphic_on (ball z r)" "norm(w - z) < r" + shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w)) + (circlepath z r)" +proof - + have "r > 0" + using assms le_less_trans norm_ge_zero by blast + have "((\<lambda>u. f u / (u - w)) has_contour_integral (2 * pi) * \<i> * winding_number (circlepath z r) w * f w) + (circlepath z r)" + apply (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) + using assms \<open>r > 0\<close> + apply (simp_all add: dist_norm norm_minus_commute) + apply (metis at_within_interior dist_norm holomorphic_on_def interior_ball mem_ball norm_minus_commute) + apply (simp add: cball_def sphere_def dist_norm, clarify) + apply (simp add:) + by (metis dist_commute dist_norm less_irrefl) + then show ?thesis + by (simp add: winding_number_circlepath assms) +qed + +corollary Cauchy_integral_circlepath_simple: + assumes "f holomorphic_on cball z r" "norm(w - z) < r" + shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w)) + (circlepath z r)" +using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) + + +lemma no_bounded_connected_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s" + and nb: "\<And>z. bounded (connected_component_set (- s) z) \<longrightarrow> z \<in> s" + shows "winding_number g z = 0" +apply (rule winding_number_zero_in_outside) +apply (simp_all add: assms) +by (metis nb [of z] \<open>path_image g \<subseteq> s\<close> \<open>z \<notin> s\<close> contra_subsetD mem_Collect_eq outside outside_mono) + +lemma no_bounded_path_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s" + and nb: "\<And>z. bounded (path_component_set (- s) z) \<longrightarrow> z \<in> s" + shows "winding_number g z = 0" +apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) +by (simp add: bounded_subset nb path_component_subset_connected_component) + + +subsection\<open> Uniform convergence of path integral\<close> + +text\<open>Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\<close> + +proposition contour_integral_uniform_limit: + assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on \<gamma>) F" + and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image \<gamma>. norm(f n x - l x) < e) F" + and noleB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B" + and \<gamma>: "valid_path \<gamma>" + and [simp]: "~ (trivial_limit F)" + shows "l contour_integrable_on \<gamma>" "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F" +proof - + have "0 \<le> B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) + { fix e::real + assume "0 < e" + then have eB: "0 < e / (\<bar>B\<bar> + 1)" by simp + obtain a where fga: "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (f a (\<gamma> x) - l (\<gamma> x)) < e / (\<bar>B\<bar> + 1)" + and inta: "(\<lambda>t. f a (\<gamma> t) * vector_derivative \<gamma> (at t)) integrable_on {0..1}" + using eventually_happens [OF eventually_conj [OF ev_no [OF eB] ev_fint]] + by (fastforce simp: contour_integrable_on path_image_def) + have Ble: "B * e / (\<bar>B\<bar> + 1) \<le> e" + using \<open>0 \<le> B\<close> \<open>0 < e\<close> by (simp add: divide_simps) + have "\<exists>h. (\<forall>x\<in>{0..1}. cmod (l (\<gamma> x) * vector_derivative \<gamma> (at x) - h x) \<le> e) \<and> h integrable_on {0..1}" + apply (rule_tac x="\<lambda>x. f (a::'a) (\<gamma> x) * vector_derivative \<gamma> (at x)" in exI) + apply (intro inta conjI ballI) + apply (rule order_trans [OF _ Ble]) + apply (frule noleB) + apply (frule fga) + using \<open>0 \<le> B\<close> \<open>0 < e\<close> + apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) + apply (drule (1) mult_mono [OF less_imp_le]) + apply (simp_all add: mult_ac) + done + } + then show lintg: "l contour_integrable_on \<gamma>" + apply (simp add: contour_integrable_on) + apply (blast intro: integrable_uniform_limit_real) + done + { fix e::real + define B' where "B' = B + 1" + have B': "B' > 0" "B' > B" using \<open>0 \<le> B\<close> by (auto simp: B'_def) + assume "0 < e" + then have ev_no': "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image \<gamma>. 2 * cmod (f n x - l x) < e / B'" + using ev_no [of "e / B' / 2"] B' by (simp add: field_simps) + have ie: "integral {0..1::real} (\<lambda>x. e / 2) < e" using \<open>0 < e\<close> by simp + have *: "cmod (f x (\<gamma> t) * vector_derivative \<gamma> (at t) - l (\<gamma> t) * vector_derivative \<gamma> (at t)) \<le> e / 2" + if t: "t\<in>{0..1}" and leB': "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) < e / B'" for x t + proof - + have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) \<le> e * (B/ B')" + using mult_mono [OF less_imp_le [OF leB'] noleB] B' \<open>0 < e\<close> t by auto + also have "... < e" + by (simp add: B' \<open>0 < e\<close> mult_imp_div_pos_less) + finally have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) < e" . + then show ?thesis + by (simp add: left_diff_distrib [symmetric] norm_mult) + qed + have "\<forall>\<^sub>F x in F. dist (contour_integral \<gamma> (f x)) (contour_integral \<gamma> l) < e" + apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) + apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) + apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric], clarify) + apply (rule le_less_trans [OF integral_norm_bound_integral ie]) + apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) + apply (blast intro: *)+ + done + } + then show "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F" + by (rule tendstoI) +qed + +proposition contour_integral_uniform_limit_circlepath: + assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on (circlepath z r)) F" + and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image (circlepath z r). norm(f n x - l x) < e) F" + and [simp]: "~ (trivial_limit F)" "0 < r" + shows "l contour_integrable_on (circlepath z r)" "((\<lambda>n. contour_integral (circlepath z r) (f n)) \<longlongrightarrow> contour_integral (circlepath z r) l) F" +by (auto simp: vector_derivative_circlepath norm_mult intro: contour_integral_uniform_limit assms) + + +subsection\<open> General stepping result for derivative formulas.\<close> + +lemma sum_sqs_eq: + fixes x::"'a::idom" shows "x * x + y * y = x * (y * 2) \<Longrightarrow> y = x" + by algebra + +proposition Cauchy_next_derivative: + assumes "continuous_on (path_image \<gamma>) f'" + and leB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B" + and int: "\<And>w. w \<in> s - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f' u / (u - w)^k) has_contour_integral f w) \<gamma>" + and k: "k \<noteq> 0" + and "open s" + and \<gamma>: "valid_path \<gamma>" + and w: "w \<in> s - path_image \<gamma>" + shows "(\<lambda>u. f' u / (u - w)^(Suc k)) contour_integrable_on \<gamma>" + and "(f has_field_derivative (k * contour_integral \<gamma> (\<lambda>u. f' u/(u - w)^(Suc k)))) + (at w)" (is "?thes2") +proof - + have "open (s - path_image \<gamma>)" using \<open>open s\<close> closed_valid_path_image \<gamma> by blast + then obtain d where "d>0" and d: "ball w d \<subseteq> s - path_image \<gamma>" using w + using open_contains_ball by blast + have [simp]: "\<And>n. cmod (1 + of_nat n) = 1 + of_nat n" + by (metis norm_of_nat of_nat_Suc) + have 1: "\<forall>\<^sub>F n in at w. (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) + contour_integrable_on \<gamma>" + apply (simp add: eventually_at) + apply (rule_tac x=d in exI) + apply (simp add: \<open>d > 0\<close> dist_norm field_simps, clarify) + apply (rule contour_integrable_div [OF contour_integrable_diff]) + using int w d + apply (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ + done + have bim_g: "bounded (image f' (path_image \<gamma>))" + by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) + then obtain C where "C > 0" and C: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> cmod (f' (\<gamma> x)) \<le> C" + by (force simp: bounded_pos path_image_def) + have twom: "\<forall>\<^sub>F n in at w. + \<forall>x\<in>path_image \<gamma>. + cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" + if "0 < e" for e + proof - + have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" + if x: "x \<in> path_image \<gamma>" and "u \<noteq> w" and uwd: "cmod (u - w) < d/2" + and uw_less: "cmod (u - w) < e * (d / 2) ^ (k+2) / (1 + real k)" + for u x + proof - + define ff where [abs_def]: + "ff n w = + (if n = 0 then inverse(x - w)^k + else if n = 1 then k / (x - w)^(Suc k) + else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w + have km1: "\<And>z::complex. z \<noteq> 0 \<Longrightarrow> z ^ (k - Suc 0) = z ^ k / z" + by (simp add: field_simps) (metis Suc_pred \<open>k \<noteq> 0\<close> neq0_conv power_Suc) + have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d / 2))" + if "z \<in> ball w (d / 2)" "i \<le> 1" for i z + proof - + have "z \<notin> path_image \<gamma>" + using \<open>x \<in> path_image \<gamma>\<close> d that ball_divide_subset_numeral by blast + then have xz[simp]: "x \<noteq> z" using \<open>x \<in> path_image \<gamma>\<close> by blast + then have neq: "x * x + z * z \<noteq> x * (z * 2)" + by (blast intro: dest!: sum_sqs_eq) + with xz have "\<And>v. v \<noteq> 0 \<Longrightarrow> (x * x + z * z) * v \<noteq> (x * (z * 2) * v)" by auto + then have neqq: "\<And>v. v \<noteq> 0 \<Longrightarrow> x * (x * v) + z * (z * v) \<noteq> x * (z * (2 * v))" + by (simp add: algebra_simps) + show ?thesis using \<open>i \<le> 1\<close> + apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) + apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ + done + qed + { fix a::real and b::real assume ab: "a > 0" "b > 0" + then have "k * (1 + real k) * (1 / a) \<le> k * (1 + real k) * (4 / b) \<longleftrightarrow> b \<le> 4 * a" + apply (subst mult_le_cancel_left_pos) + using \<open>k \<noteq> 0\<close> + apply (auto simp: divide_simps) + done + with ab have "real k * (1 + real k) / a \<le> (real k * 4 + real k * real k * 4) / b \<longleftrightarrow> b \<le> 4 * a" + by (simp add: field_simps) + } note canc = this + have ff2: "cmod (ff (Suc 1) v) \<le> real (k * (k + 1)) / (d / 2) ^ (k + 2)" + if "v \<in> ball w (d / 2)" for v + proof - + have "d/2 \<le> cmod (x - v)" using d x that + apply (simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps, clarify) + apply (drule subsetD) + prefer 2 apply blast + apply (metis norm_minus_commute norm_triangle_half_r CollectI) + done + then have "d \<le> cmod (x - v) * 2" + by (simp add: divide_simps) + then have dpow_le: "d ^ (k+2) \<le> (cmod (x - v) * 2) ^ (k+2)" + using \<open>0 < d\<close> order_less_imp_le power_mono by blast + have "x \<noteq> v" using that + using \<open>x \<in> path_image \<gamma>\<close> ball_divide_subset_numeral d by fastforce + then show ?thesis + using \<open>d > 0\<close> + apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) + using dpow_le + apply (simp add: algebra_simps divide_simps mult_less_0_iff) + done + qed + have ub: "u \<in> ball w (d / 2)" + using uwd by (simp add: dist_commute dist_norm) + have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \<le> (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d / 2) ^ k))" + using complex_taylor [OF _ ff1 ff2 _ ub, of w, simplified] + by (simp add: ff_def \<open>0 < d\<close>) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \<le> (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)" + by (simp add: field_simps) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + / (cmod (u - w) * real k) + \<le> (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)" + using \<open>k \<noteq> 0\<close> \<open>u \<noteq> w\<close> by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) + also have "... < e" + using uw_less \<open>0 < d\<close> by (simp add: mult_ac divide_simps) + finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) + / cmod ((u - w) * real k) < e" + by (simp add: norm_mult) + have "x \<noteq> u" + using uwd \<open>0 < d\<close> x d by (force simp: dist_norm ball_def norm_minus_commute) + show ?thesis + apply (rule le_less_trans [OF _ e]) + using \<open>k \<noteq> 0\<close> \<open>x \<noteq> u\<close> \<open>u \<noteq> w\<close> + apply (simp add: field_simps norm_divide [symmetric]) + done + qed + show ?thesis + unfolding eventually_at + apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) + apply (force simp: \<open>d > 0\<close> dist_norm that simp del: power_Suc intro: *) + done + qed + have 2: "\<forall>\<^sub>F n in at w. + \<forall>x\<in>path_image \<gamma>. + cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" + if "0 < e" for e + proof - + have *: "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) - + f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) < e" + if ec: "cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) - + inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k) < e / C" + and x: "0 \<le> x" "x \<le> 1" + for u x + proof (cases "(f' (\<gamma> x)) = 0") + case True then show ?thesis by (simp add: \<open>0 < e\<close>) + next + case False + have "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) - + f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) = + cmod (f' (\<gamma> x) * ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) - + inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k))" + by (simp add: field_simps) + also have "... = cmod (f' (\<gamma> x)) * + cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) - + inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k)" + by (simp add: norm_mult) + also have "... < cmod (f' (\<gamma> x)) * (e/C)" + apply (rule mult_strict_left_mono [OF ec]) + using False by simp + also have "... \<le> e" using C + by (metis False \<open>0 < e\<close> frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) + finally show ?thesis . + qed + show ?thesis + using twom [OF divide_pos_pos [OF that \<open>C > 0\<close>]] unfolding path_image_def + by (force intro: * elim: eventually_mono) + qed + show "(\<lambda>u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \<gamma>" + by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto + have *: "(\<lambda>n. contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) + \<midarrow>w\<rightarrow> contour_integral \<gamma> (\<lambda>u. f' u / (u - w) ^ (Suc k))" + by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto + have **: "contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = + (f u - f w) / (u - w) / k" + if "dist u w < d" for u + apply (rule contour_integral_unique) + apply (simp add: diff_divide_distrib algebra_simps) + apply (rule has_contour_integral_diff; rule has_contour_integral_div; simp add: field_simps; rule int) + apply (metis contra_subsetD d dist_commute mem_ball that) + apply (rule w) + done + show ?thes2 + apply (simp add: DERIV_within_iff del: power_Suc) + apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \<open>0 < d\<close> ]) + apply (simp add: \<open>k \<noteq> 0\<close> **) + done +qed + +corollary Cauchy_next_derivative_circlepath: + assumes contf: "continuous_on (path_image (circlepath z r)) f" + and int: "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" + and k: "k \<noteq> 0" + and w: "w \<in> ball z r" + shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(g has_field_derivative (k * contour_integral (circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)))) (at w)" + (is "?thes2") +proof - + have "r > 0" using w + using ball_eq_empty by fastforce + have wim: "w \<in> ball z r - path_image (circlepath z r)" + using w by (auto simp: dist_norm) + show ?thes1 ?thes2 + by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \<bar>r\<bar>"]; + auto simp: vector_derivative_circlepath norm_mult)+ +qed + + +text\<open> In particular, the first derivative formula.\<close> + +proposition Cauchy_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \<in> ball z r" + shows "(\<lambda>u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(f has_field_derivative (1 / (2 * of_real pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u / (u - w)^2))) (at w)" + (is "?thes2") +proof - + have [simp]: "r \<ge> 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def) + have int: "\<And>w. dist z w < r \<Longrightarrow> + ((\<lambda>u. f u / (u - w)) has_contour_integral (\<lambda>x. 2 * of_real pi * \<i> * f x) w) (circlepath z r)" + by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) + show ?thes1 + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) + apply (blast intro: int) + done + have "((\<lambda>x. 2 * of_real pi * \<i> * f x) has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2)) (at w)" + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\<lambda>x. 2 * of_real pi * \<i> * f x", simplified]) + apply (blast intro: int) + done + then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2) / (2 * of_real pi * \<i>)) (at w)" + by (rule DERIV_cdivide [where f = "\<lambda>x. 2 * of_real pi * \<i> * f x" and c = "2 * of_real pi * \<i>", simplified]) + show ?thes2 + by simp (rule fder) +qed + +subsection\<open> Existence of all higher derivatives.\<close> + +proposition derivative_is_holomorphic: + assumes "open s" + and fder: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z)" + shows "f' holomorphic_on s" +proof - + have *: "\<exists>h. (f' has_field_derivative h) (at z)" if "z \<in> s" for z + proof - + obtain r where "r > 0" and r: "cball z r \<subseteq> s" + using open_contains_cball \<open>z \<in> s\<close> \<open>open s\<close> by blast + then have holf_cball: "f holomorphic_on cball z r" + apply (simp add: holomorphic_on_def) + using field_differentiable_at_within field_differentiable_def fder by blast + then have "continuous_on (path_image (circlepath z r)) f" + using \<open>r > 0\<close> by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) + then have contfpi: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1/(2 * of_real pi*\<i>) * f x)" + by (auto intro: continuous_intros)+ + have contf_cball: "continuous_on (cball z r) f" using holf_cball + by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) + have holf_ball: "f holomorphic_on ball z r" using holf_cball + using ball_subset_cball holomorphic_on_subset by blast + { fix w assume w: "w \<in> ball z r" + have intf: "(\<lambda>u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have fder': "(f has_field_derivative 1 / (2 * of_real pi * \<i>) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2)) + (at w)" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have f'_eq: "f' w = contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>)" + using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) + have "((\<lambda>u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \<i>)) has_contour_integral + contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>)) + (circlepath z r)" + by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) + then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral + contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>)) + (circlepath z r)" + by (simp add: algebra_simps) + then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" + by (simp add: f'_eq) + } note * = this + show ?thesis + apply (rule exI) + apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) + apply (simp_all add: \<open>0 < r\<close> * dist_norm) + done + qed + show ?thesis + by (simp add: holomorphic_on_open [OF \<open>open s\<close>] *) +qed + +lemma holomorphic_deriv [holomorphic_intros]: + "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv f) holomorphic_on s" +by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) + +lemma analytic_deriv: "f analytic_on s \<Longrightarrow> (deriv f) analytic_on s" + using analytic_on_holomorphic holomorphic_deriv by auto + +lemma holomorphic_higher_deriv [holomorphic_intros]: "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv ^^ n) f holomorphic_on s" + by (induction n) (auto simp: holomorphic_deriv) + +lemma analytic_higher_deriv: "f analytic_on s \<Longrightarrow> (deriv ^^ n) f analytic_on s" + unfolding analytic_on_def using holomorphic_higher_deriv by blast + +lemma has_field_derivative_higher_deriv: + "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk> + \<Longrightarrow> ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" +by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply + funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) + +lemma valid_path_compose_holomorphic: + assumes "valid_path g" and holo:"f holomorphic_on s" and "open s" "path_image g \<subseteq> s" + shows "valid_path (f o g)" +proof (rule valid_path_compose[OF \<open>valid_path g\<close>]) + fix x assume "x \<in> path_image g" + then show "\<exists>f'. (f has_field_derivative f') (at x)" + using holo holomorphic_on_open[OF \<open>open s\<close>] \<open>path_image g \<subseteq> s\<close> by auto +next + have "deriv f holomorphic_on s" + using holomorphic_deriv holo \<open>open s\<close> by auto + then show "continuous_on (path_image g) (deriv f)" + using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto +qed + + +subsection\<open> Morera's theorem.\<close> + +lemma Morera_local_triangle_ball: + assumes "\<And>z. z \<in> s + \<Longrightarrow> \<exists>e a. 0 < e \<and> z \<in> ball a e \<and> continuous_on (ball a e) f \<and> + (\<forall>b c. closed_segment b c \<subseteq> ball a e + \<longrightarrow> contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on s" +proof - + { fix z assume "z \<in> s" + with assms obtain e a where + "0 < e" and z: "z \<in> ball a e" and contf: "continuous_on (ball a e) f" + and 0: "\<And>b c. closed_segment b c \<subseteq> ball a e + \<Longrightarrow> contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by fastforce + have az: "dist a z < e" using mem_ball z by blast + have sb_ball: "ball z (e - dist a z) \<subseteq> ball a e" + by (simp add: dist_commute ball_subset_ball_iff) + have "\<exists>e>0. f holomorphic_on ball z e" + apply (rule_tac x="e - dist a z" in exI) + apply (simp add: az) + apply (rule holomorphic_on_subset [OF _ sb_ball]) + apply (rule derivative_is_holomorphic[OF open_ball]) + apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) + apply (simp_all add: 0 \<open>0 < e\<close>) + apply (meson \<open>0 < e\<close> centre_in_ball convex_ball convex_contains_segment mem_ball) + done + } + then show ?thesis + by (simp add: analytic_on_def) +qed + +lemma Morera_local_triangle: + assumes "\<And>z. z \<in> s + \<Longrightarrow> \<exists>t. open t \<and> z \<in> t \<and> continuous_on t f \<and> + (\<forall>a b c. convex hull {a,b,c} \<subseteq> t + \<longrightarrow> contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on s" +proof - + { fix z assume "z \<in> s" + with assms obtain t where + "open t" and z: "z \<in> t" and contf: "continuous_on t f" + and 0: "\<And>a b c. convex hull {a,b,c} \<subseteq> t + \<Longrightarrow> contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by force + then obtain e where "e>0" and e: "ball z e \<subseteq> t" + using open_contains_ball by blast + have [simp]: "continuous_on (ball z e) f" using contf + using continuous_on_subset e by blast + have "\<exists>e a. 0 < e \<and> + z \<in> ball a e \<and> + continuous_on (ball a e) f \<and> + (\<forall>b c. closed_segment b c \<subseteq> ball a e \<longrightarrow> + contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" + apply (rule_tac x=e in exI) + apply (rule_tac x=z in exI) + apply (simp add: \<open>e > 0\<close>, clarify) + apply (rule 0) + apply (meson z \<open>0 < e\<close> centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) + done + } + then show ?thesis + by (simp add: Morera_local_triangle_ball) +qed + +proposition Morera_triangle: + "\<lbrakk>continuous_on s f; open s; + \<And>a b c. convex hull {a,b,c} \<subseteq> s + \<longrightarrow> contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0\<rbrakk> + \<Longrightarrow> f analytic_on s" + using Morera_local_triangle by blast + + + +subsection\<open> Combining theorems for higher derivatives including Leibniz rule.\<close> + +lemma higher_deriv_linear [simp]: + "(deriv ^^ n) (\<lambda>w. c*w) = (\<lambda>z. if n = 0 then c*z else if n = 1 then c else 0)" + by (induction n) (auto simp: deriv_const deriv_linear) + +lemma higher_deriv_const [simp]: "(deriv ^^ n) (\<lambda>w. c) = (\<lambda>w. if n=0 then c else 0)" + by (induction n) (auto simp: deriv_const) + +lemma higher_deriv_ident [simp]: + "(deriv ^^ n) (\<lambda>w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" + apply (induction n, simp) + apply (metis higher_deriv_linear lambda_one) + done + +corollary higher_deriv_id [simp]: + "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" + by (simp add: id_def) + +lemma has_complex_derivative_funpow_1: + "\<lbrakk>(f has_field_derivative 1) (at z); f z = z\<rbrakk> \<Longrightarrow> (f^^n has_field_derivative 1) (at z)" + apply (induction n) + apply auto + apply (metis DERIV_ident DERIV_transform_at id_apply zero_less_one) + by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) + +proposition higher_deriv_uminus: + assumes "f holomorphic_on s" "open s" and z: "z \<in> s" + shows "(deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + show ?case + apply simp + apply (rule DERIV_imp_deriv) + apply (rule DERIV_transform_within_open [of "\<lambda>w. -((deriv ^^ n) f w)"]) + apply (rule derivative_eq_intros | rule * refl assms Suc)+ + apply (simp add: Suc) + done +qed + +proposition higher_deriv_add: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s" + shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + show ?case + apply simp + apply (rule DERIV_imp_deriv) + apply (rule DERIV_transform_within_open [of "\<lambda>w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) + apply (rule derivative_eq_intros | rule * refl assms Suc)+ + apply (simp add: Suc) + done +qed + +corollary higher_deriv_diff: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s" + shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" + apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add) + apply (subst higher_deriv_add) + using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus) + done + + +lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" + by (cases k) simp_all + +proposition higher_deriv_mult: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s" + shows "(deriv ^^ n) (\<lambda>w. f w * g w) z = + (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "\<And>n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have sumeq: "(\<Sum>i = 0..n. + of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = + g z * deriv ((deriv ^^ n) f) z + (\<Sum>i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" + apply (simp add: bb algebra_simps setsum.distrib) + apply (subst (4) setsum_Suc_reindex) + apply (auto simp: algebra_simps Suc_diff_le intro: setsum.cong) + done + show ?case + apply (simp only: funpow.simps o_apply) + apply (rule DERIV_imp_deriv) + apply (rule DERIV_transform_within_open + [of "\<lambda>w. (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) + apply (simp add: algebra_simps) + apply (rule DERIV_cong [OF DERIV_setsum]) + apply (rule DERIV_cmult) + apply (auto simp: intro: DERIV_mult * sumeq \<open>open s\<close> Suc.prems Suc.IH [symmetric]) + done +qed + + +proposition higher_deriv_transform_within_open: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s" + and fg: "\<And>w. w \<in> s \<Longrightarrow> f w = g w" + shows "(deriv ^^ i) f z = (deriv ^^ i) g z" +using z +by (induction i arbitrary: z) + (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) + +proposition higher_deriv_compose_linear: + fixes z::complex + assumes f: "f holomorphic_on t" and s: "open s" and t: "open t" and z: "z \<in> s" + and fg: "\<And>w. w \<in> s \<Longrightarrow> u * w \<in> t" + shows "(deriv ^^ n) (\<lambda>w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have holo0: "f holomorphic_on op * u ` s" + by (meson fg f holomorphic_on_subset image_subset_iff) + have holo1: "(\<lambda>w. f (u * w)) holomorphic_on s" + apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) + apply (rule holo0 holomorphic_intros)+ + done + have holo2: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on s" + apply (rule holomorphic_intros)+ + apply (rule holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def]) + apply (rule holomorphic_intros) + apply (rule holomorphic_on_subset [where s=t]) + apply (rule holomorphic_intros assms)+ + apply (blast intro: fg) + done + have "deriv ((deriv ^^ n) (\<lambda>w. f (u * w))) z = deriv (\<lambda>z. u^n * (deriv ^^ n) f (u*z)) z" + apply (rule complex_derivative_transform_within_open [OF _ holo2 s Suc.prems]) + apply (rule holomorphic_higher_deriv [OF holo1 s]) + apply (simp add: Suc.IH) + done + also have "... = u^n * deriv (\<lambda>z. (deriv ^^ n) f (u * z)) z" + apply (rule deriv_cmult) + apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) + apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and t=t, unfolded o_def]) + apply (simp add: analytic_on_linear) + apply (simp add: analytic_on_open f holomorphic_higher_deriv t) + apply (blast intro: fg) + done + also have "... = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" + apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "op*u", unfolded o_def]) + apply (rule derivative_intros) + using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv t apply blast + apply (simp add: deriv_linear) + done + finally show ?case + by simp +qed + +lemma higher_deriv_add_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \<and> g analytic_on {z}" + using assms by blast + with higher_deriv_add show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_diff_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \<and> g analytic_on {z}" + using assms by blast + with higher_deriv_diff show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_uminus_at: + "f analytic_on {z} \<Longrightarrow> (deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)" + using higher_deriv_uminus + by (auto simp: analytic_at) + +lemma higher_deriv_mult_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\<lambda>w. f w * g w) z = + (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +proof - + have "f analytic_on {z} \<and> g analytic_on {z}" + using assms by blast + with higher_deriv_mult show ?thesis + by (auto simp: analytic_at_two) +qed + + +text\<open> Nonexistence of isolated singularities and a stronger integral formula.\<close> + +proposition no_isolated_singularity: + fixes z::complex + assumes f: "continuous_on s f" and holf: "f holomorphic_on (s - k)" and s: "open s" and k: "finite k" + shows "f holomorphic_on s" +proof - + { fix z + assume "z \<in> s" and cdf: "\<And>x. x\<in>s - k \<Longrightarrow> f field_differentiable at x" + have "f field_differentiable at z" + proof (cases "z \<in> k") + case False then show ?thesis by (blast intro: cdf \<open>z \<in> s\<close>) + next + case True + with finite_set_avoid [OF k, of z] + obtain d where "d>0" and d: "\<And>x. \<lbrakk>x\<in>k; x \<noteq> z\<rbrakk> \<Longrightarrow> d \<le> dist z x" + by blast + obtain e where "e>0" and e: "ball z e \<subseteq> s" + using s \<open>z \<in> s\<close> by (force simp add: open_contains_ball) + have fde: "continuous_on (ball z (min d e)) f" + by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) + have "\<exists>g. \<forall>w \<in> ball z (min d e). (g has_field_derivative f w) (at w within ball z (min d e))" + apply (rule contour_integral_convex_primitive [OF convex_ball fde]) + apply (rule Cauchy_theorem_triangle_cofinite [OF _ k]) + apply (metis continuous_on_subset [OF fde] closed_segment_subset convex_ball starlike_convex_subset) + apply (rule cdf) + apply (metis Diff_iff Int_iff ball_min_Int bot_least contra_subsetD convex_ball e insert_subset + interior_mono interior_subset subset_hull) + done + then have "f holomorphic_on ball z (min d e)" + by (metis open_ball at_within_open derivative_is_holomorphic) + then show ?thesis + unfolding holomorphic_on_def + by (metis open_ball \<open>0 < d\<close> \<open>0 < e\<close> at_within_open centre_in_ball min_less_iff_conj) + qed + } + with holf s k show ?thesis + by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) +qed + +proposition Cauchy_integral_formula_convex: + assumes s: "convex s" and k: "finite k" and contf: "continuous_on s f" + and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)" + and z: "z \<in> interior s" and vpg: "valid_path \<gamma>" + and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>" + apply (rule Cauchy_integral_formula_weak [OF s finite.emptyI contf]) + apply (simp add: holomorphic_on_open [symmetric] field_differentiable_def) + using no_isolated_singularity [where s = "interior s"] + apply (metis k contf fcd holomorphic_on_open field_differentiable_def continuous_on_subset + has_field_derivative_at_within holomorphic_on_def interior_subset open_interior) + using assms + apply auto + done + + +text\<open> Formula for higher derivatives.\<close> + +proposition Cauchy_has_contour_integral_higher_derivative_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \<in> ball z r" + shows "((\<lambda>u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \<i>) / (fact k) * (deriv ^^ k) f w)) + (circlepath z r)" +using w +proof (induction k arbitrary: w) + case 0 then show ?case + using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) +next + case (Suc k) + have [simp]: "r > 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def less_imp_le) + obtain X where X: "((\<lambda>u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" + using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] + by (auto simp: contour_integrable_on_def) + then have con: "contour_integral (circlepath z r) ((\<lambda>u. f u / (u - w) ^ Suc (Suc k))) = X" + by (rule contour_integral_unique) + have "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + then have dnf_diff: "\<And>n. (deriv ^^ n) f field_differentiable (at w)" + by (force simp add: field_differentiable_def) + have "deriv (\<lambda>w. complex_of_real (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w) w = + of_nat (Suc k) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w) ^ Suc (Suc k))" + by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) + also have "... = of_nat (Suc k) * X" + by (simp only: con) + finally have "deriv (\<lambda>w. ((2 * pi) * \<i> / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . + then have "((2 * pi) * \<i> / (fact k)) * deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" + by (metis deriv_cmult dnf_diff) + then have "deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \<i> / (fact k))" + by (simp add: field_simps) + then show ?case + using of_nat_eq_0_iff X by fastforce +qed + +proposition Cauchy_higher_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \<in> ball z r" + shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(deriv ^^ k) f w = (fact k) / (2 * pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k))" + (is "?thes2") +proof - + have *: "((\<lambda>u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w) + (circlepath z r)" + using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] + by simp + show ?thes1 using * + using contour_integrable_on_def by blast + show ?thes2 + unfolding contour_integral_unique [OF *] by (simp add: divide_simps) +qed + +corollary Cauchy_contour_integral_circlepath: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r" + shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)) = (2 * pi * \<i>) * (deriv ^^ k) f w / (fact k)" +by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) + +corollary Cauchy_contour_integral_circlepath_2: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r" + shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^2) = (2 * pi * \<i>) * deriv f w" + using Cauchy_contour_integral_circlepath [OF assms, of 1] + by (simp add: power2_eq_square) + + +subsection\<open>A holomorphic function is analytic, i.e. has local power series.\<close> + +theorem holomorphic_power_series: + assumes holf: "f holomorphic_on ball z r" + and w: "w \<in> ball z r" + shows "((\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" +proof - + have fh': "f holomorphic_on cball z ((r + dist w z) / 2)" + apply (rule holomorphic_on_subset [OF holf]) + apply (clarsimp simp del: divide_const_simps) + apply (metis add.commute dist_commute le_less_trans mem_ball real_gt_half_sum w) + done + \<comment>\<open>Replacing @{term r} and the original (weak) premises\<close> + obtain r where "0 < r" and holfc: "f holomorphic_on cball z r" and w: "w \<in> ball z r" + apply (rule that [of "(r + dist w z) / 2"]) + apply (simp_all add: fh') + apply (metis add_0_iff ball_eq_empty dist_nz dist_self empty_iff not_less pos_add_strict w) + apply (metis add_less_cancel_right dist_commute mem_ball mult_2_right w) + done + then have holf: "f holomorphic_on ball z r" and contf: "continuous_on (cball z r) f" + using ball_subset_cball holomorphic_on_subset apply blast + by (simp add: holfc holomorphic_on_imp_continuous_on) + have cint: "\<And>k. (\<lambda>u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" + apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) + apply (simp add: \<open>0 < r\<close>) + done + obtain B where "0 < B" and B: "\<And>u. u \<in> cball z r \<Longrightarrow> norm(f u) \<le> B" + by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) + obtain k where k: "0 < k" "k \<le> r" and wz_eq: "norm(w - z) = r - k" + and kle: "\<And>u. norm(u - z) = r \<Longrightarrow> k \<le> norm(u - w)" + apply (rule_tac k = "r - dist z w" in that) + using w + apply (auto simp: dist_norm norm_minus_commute) + by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + have *: "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>path_image (circlepath z r). + norm ((\<Sum>k<n. (w - z) ^ k * (f x / (x - z) ^ Suc k)) - f x / (x - w)) < e" + if "0 < e" for e + proof - + have rr: "0 \<le> (r - k) / r" "(r - k) / r < 1" using k by auto + obtain n where n: "((r - k) / r) ^ n < e / B * k" + using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \<open>0 < e\<close> \<open>0 < B\<close> k by force + have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) - f u / (u - w)) < e" + if "n \<le> N" and r: "r = dist z u" for N u + proof - + have N: "((r - k) / r) ^ N < e / B * k" + apply (rule le_less_trans [OF power_decreasing n]) + using \<open>n \<le> N\<close> k by auto + have u [simp]: "(u \<noteq> z) \<and> (u \<noteq> w)" + using \<open>0 < r\<close> r w by auto + have wzu_not1: "(w - z) / (u - z) \<noteq> 1" + by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) + have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u) + = norm ((\<Sum>k<N. (((w - z) / (u - z)) ^ k)) * f u * (u - w) / (u - z) - f u)" + unfolding setsum_left_distrib setsum_divide_distrib power_divide by (simp add: algebra_simps) + also have "... = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" + using \<open>0 < B\<close> + apply (auto simp: geometric_sum [OF wzu_not1]) + apply (simp add: field_simps norm_mult [symmetric]) + done + also have "... = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" + using \<open>0 < r\<close> r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) + also have "... = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" + by (simp add: algebra_simps) + also have "... = norm (w - z) ^ N * norm (f u) / r ^ N" + by (simp add: norm_mult norm_power norm_minus_commute) + also have "... \<le> (((r - k)/r)^N) * B" + using \<open>0 < r\<close> w k + apply (simp add: divide_simps) + apply (rule mult_mono [OF power_mono]) + apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) + done + also have "... < e * k" + using \<open>0 < B\<close> N by (simp add: divide_simps) + also have "... \<le> e * norm (u - w)" + using r kle \<open>0 < e\<close> by (simp add: dist_commute dist_norm) + finally show ?thesis + by (simp add: divide_simps norm_divide del: power_Suc) + qed + with \<open>0 < r\<close> show ?thesis + by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def) + qed + have eq: "\<forall>\<^sub>F x in sequentially. + contour_integral (circlepath z r) (\<lambda>u. \<Sum>k<x. (w - z) ^ k * (f u / (u - z) ^ Suc k)) = + (\<Sum>k<x. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z) ^ k)" + apply (rule eventuallyI) + apply (subst contour_integral_setsum, simp) + using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) + apply (simp only: contour_integral_lmul cint algebra_simps) + done + have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k) + sums contour_integral (circlepath z r) (\<lambda>u. f u/(u - w))" + unfolding sums_def + apply (rule Lim_transform_eventually [OF eq]) + apply (rule contour_integral_uniform_limit_circlepath [OF eventuallyI *]) + apply (rule contour_integrable_setsum, simp) + apply (rule contour_integrable_lmul) + apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) + using \<open>0 < r\<close> + apply auto + done + then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k) + sums (2 * of_real pi * \<i> * f w)" + using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) + then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z)^k / (\<i> * (of_real pi * 2))) + sums ((2 * of_real pi * \<i> * f w) / (\<i> * (complex_of_real pi * 2)))" + by (rule sums_divide) + then have "(\<lambda>n. (w - z) ^ n * contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc n) / (\<i> * (of_real pi * 2))) + sums f w" + by (simp add: field_simps) + then show ?thesis + by (simp add: field_simps \<open>0 < r\<close> Cauchy_higher_derivative_integral_circlepath [OF contf holf]) +qed + + +subsection\<open>The Liouville theorem and the Fundamental Theorem of Algebra.\<close> + +text\<open> These weak Liouville versions don't even need the derivative formula.\<close> + +lemma Liouville_weak_0: + assumes holf: "f holomorphic_on UNIV" and inf: "(f \<longlongrightarrow> 0) at_infinity" + shows "f z = 0" +proof (rule ccontr) + assume fz: "f z \<noteq> 0" + with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] + obtain B where B: "\<And>x. B \<le> cmod x \<Longrightarrow> norm (f x) * 2 < cmod (f z)" + by (auto simp: dist_norm) + define R where "R = 1 + \<bar>B\<bar> + norm z" + have "R > 0" unfolding R_def + proof - + have "0 \<le> cmod z + \<bar>B\<bar>" + by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) + then show "0 < 1 + \<bar>B\<bar> + cmod z" + by linarith + qed + have *: "((\<lambda>u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \<i> * f z) (circlepath z R)" + apply (rule Cauchy_integral_circlepath) + using \<open>R > 0\<close> apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ + done + have "cmod (x - z) = R \<Longrightarrow> cmod (f x) * 2 \<le> cmod (f z)" for x + apply (simp add: R_def) + apply (rule less_imp_le) + apply (rule B) + using norm_triangle_ineq4 [of x z] + apply (auto simp:) + done + with \<open>R > 0\<close> fz show False + using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] + by (auto simp: norm_mult norm_divide divide_simps) +qed + +proposition Liouville_weak: + assumes "f holomorphic_on UNIV" and "(f \<longlongrightarrow> l) at_infinity" + shows "f z = l" + using Liouville_weak_0 [of "\<lambda>z. f z - l"] + by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero) + + +proposition Liouville_weak_inverse: + assumes "f holomorphic_on UNIV" and unbounded: "\<And>B. eventually (\<lambda>x. norm (f x) \<ge> B) at_infinity" + obtains z where "f z = 0" +proof - + { assume f: "\<And>z. f z \<noteq> 0" + have 1: "(\<lambda>x. 1 / f x) holomorphic_on UNIV" + by (simp add: holomorphic_on_divide holomorphic_on_const assms f) + have 2: "((\<lambda>x. 1 / f x) \<longlongrightarrow> 0) at_infinity" + apply (rule tendstoI [OF eventually_mono]) + apply (rule_tac B="2/e" in unbounded) + apply (simp add: dist_norm norm_divide divide_simps mult_ac) + done + have False + using Liouville_weak_0 [OF 1 2] f by simp + } + then show ?thesis + using that by blast +qed + + +text\<open> In particular we get the Fundamental Theorem of Algebra.\<close> + +theorem fundamental_theorem_of_algebra: + fixes a :: "nat \<Rightarrow> complex" + assumes "a 0 = 0 \<or> (\<exists>i \<in> {1..n}. a i \<noteq> 0)" + obtains z where "(\<Sum>i\<le>n. a i * z^i) = 0" +using assms +proof (elim disjE bexE) + assume "a 0 = 0" then show ?thesis + by (auto simp: that [of 0]) +next + fix i + assume i: "i \<in> {1..n}" and nz: "a i \<noteq> 0" + have 1: "(\<lambda>z. \<Sum>i\<le>n. a i * z^i) holomorphic_on UNIV" + by (rule holomorphic_intros)+ + show ?thesis + apply (rule Liouville_weak_inverse [OF 1]) + apply (rule polyfun_extremal) + apply (rule nz) + using i that + apply (auto simp:) + done +qed + + +subsection\<open> Weierstrass convergence theorem.\<close> + +proposition holomorphic_uniform_limit: + assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and> (f n) holomorphic_on ball z r) F" + and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F" + and F: "~ trivial_limit F" + obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" +proof (cases r "0::real" rule: linorder_cases) + case less then show ?thesis by (force simp add: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) +next + case equal then show ?thesis + by (force simp add: holomorphic_on_def continuous_on_sing intro: that) +next + case greater + have contg: "continuous_on (cball z r) g" + using cont + by (fastforce simp: eventually_conj_iff dist_norm intro: eventually_mono [OF lim] continuous_uniform_limit [OF F]) + have 1: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1 / (2 * complex_of_real pi * \<i>) * g x)" + apply (rule continuous_intros continuous_on_subset [OF contg])+ + using \<open>0 < r\<close> by auto + have 2: "((\<lambda>u. 1 / (2 * of_real pi * \<i>) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" + if w: "w \<in> ball z r" for w + proof - + define d where "d = (r - norm(w - z))" + have "0 < d" "d \<le> r" using w by (auto simp: norm_minus_commute d_def dist_norm) + have dle: "\<And>u. cmod (z - u) = r \<Longrightarrow> d \<le> cmod (u - w)" + unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + have ev_int: "\<forall>\<^sub>F n in F. (\<lambda>u. f n u / (u - w)) contour_integrable_on circlepath z r" + apply (rule eventually_mono [OF cont]) + using w + apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) + done + have ev_less: "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image (circlepath z r). cmod (f n x / (x - w) - g x / (x - w)) < e" + if "e > 0" for e + using greater \<open>0 < d\<close> \<open>0 < e\<close> + apply (simp add: norm_divide diff_divide_distrib [symmetric] divide_simps) + apply (rule_tac e1="e * d" in eventually_mono [OF lim]) + apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+ + done + have g_cint: "(\<lambda>u. g u/(u - w)) contour_integrable_on circlepath z r" + by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>]) + have cif_tends_cig: "((\<lambda>n. contour_integral(circlepath z r) (\<lambda>u. f n u / (u - w))) \<longlongrightarrow> contour_integral(circlepath z r) (\<lambda>u. g u/(u - w))) F" + by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>]) + have f_tends_cig: "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> contour_integral (circlepath z r) (\<lambda>u. g u / (u - w))) F" + apply (rule Lim_transform_eventually [where f = "\<lambda>n. contour_integral (circlepath z r) (\<lambda>u. f n u/(u - w))"]) + apply (rule eventually_mono [OF cont]) + apply (rule contour_integral_unique [OF Cauchy_integral_circlepath]) + using w + apply (auto simp: norm_minus_commute dist_norm cif_tends_cig) + done + have "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> 2 * of_real pi * \<i> * g w) F" + apply (rule tendsto_mult_left [OF tendstoI]) + apply (rule eventually_mono [OF lim], assumption) + using w + apply (force simp add: dist_norm) + done + then have "((\<lambda>u. g u / (u - w)) has_contour_integral 2 * of_real pi * \<i> * g w) (circlepath z r)" + using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w + by (force simp add: dist_norm) + then have "((\<lambda>u. g u / (2 * of_real pi * \<i> * (u - w))) has_contour_integral g w) (circlepath z r)" + using has_contour_integral_div [where c = "2 * of_real pi * \<i>"] + by (force simp add: field_simps) + then show ?thesis + by (simp add: dist_norm) + qed + show ?thesis + using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] + by (fastforce simp add: holomorphic_on_open contg intro: that) +qed + + +text\<open> Version showing that the limit is the limit of the derivatives.\<close> + +proposition has_complex_derivative_uniform_limit: + fixes z::complex + assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and> + (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" + and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F" + and F: "~ trivial_limit F" and "0 < r" + obtains g' where + "continuous_on (cball z r) g" + "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F" +proof - + let ?conint = "contour_integral (circlepath z r)" + have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" + by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] lim F]; + auto simp: holomorphic_on_open field_differentiable_def)+ + then obtain g' where g': "\<And>x. x \<in> ball z r \<Longrightarrow> (g has_field_derivative g' x) (at x)" + using DERIV_deriv_iff_has_field_derivative + by (fastforce simp add: holomorphic_on_open) + then have derg: "\<And>x. x \<in> ball z r \<Longrightarrow> deriv g x = g' x" + by (simp add: DERIV_imp_deriv) + have tends_f'n_g': "((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F" if w: "w \<in> ball z r" for w + proof - + have eq_f': "?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \<i>)" + if cont_fn: "continuous_on (cball z r) (f n)" + and fnd: "\<And>w. w \<in> ball z r \<Longrightarrow> (f n has_field_derivative f' n w) (at w)" for n + proof - + have hol_fn: "f n holomorphic_on ball z r" + using fnd by (force simp add: holomorphic_on_open) + have "(f n has_field_derivative 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)) (at w)" + by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) + then have f': "f' n w = 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)" + using DERIV_unique [OF fnd] w by blast + show ?thesis + by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] divide_simps) + qed + define d where "d = (r - norm(w - z))^2" + have "d > 0" + using w by (simp add: dist_commute dist_norm d_def) + have dle: "\<And>y. r = cmod (z - y) \<Longrightarrow> d \<le> cmod ((y - w)\<^sup>2)" + apply (simp add: d_def norm_power) + apply (rule power_mono) + apply (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + apply (metis diff_ge_0_iff_ge dist_commute dist_norm less_eq_real_def mem_ball w) + done + have 1: "\<forall>\<^sub>F n in F. (\<lambda>x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" + by (force simp add: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) + have 2: "0 < e \<Longrightarrow> \<forall>\<^sub>F n in F. \<forall>x \<in> path_image (circlepath z r). cmod (f n x / (x - w)\<^sup>2 - g x / (x - w)\<^sup>2) < e" for e + using \<open>r > 0\<close> + apply (simp add: diff_divide_distrib [symmetric] norm_divide divide_simps sphere_def) + apply (rule eventually_mono [OF lim, of "e*d"]) + apply (simp add: \<open>0 < d\<close>) + apply (force simp add: dist_norm dle intro: less_le_trans) + done + have "((\<lambda>n. contour_integral (circlepath z r) (\<lambda>x. f n x / (x - w)\<^sup>2)) + \<longlongrightarrow> contour_integral (circlepath z r) ((\<lambda>x. g x / (x - w)\<^sup>2))) F" + by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \<open>0 < r\<close>]) + then have tendsto_0: "((\<lambda>n. 1 / (2 * of_real pi * \<i>) * (?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2))) \<longlongrightarrow> 0) F" + using Lim_null by (force intro!: tendsto_mult_right_zero) + have "((\<lambda>n. f' n w - g' w) \<longlongrightarrow> 0) F" + apply (rule Lim_transform_eventually [OF _ tendsto_0]) + apply (force simp add: divide_simps intro: eq_f' eventually_mono [OF cont]) + done + then show ?thesis using Lim_null by blast + qed + obtain g' where "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F" + by (blast intro: tends_f'n_g' g' ) + then show ?thesis using g + using that by blast +qed + + +subsection\<open>Some more simple/convenient versions for applications.\<close> + +lemma holomorphic_uniform_sequence: + assumes s: "open s" + and hol_fn: "\<And>n. (f n) holomorphic_on s" + and to_g: "\<And>x. x \<in> s + \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and> + (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)" + shows "g holomorphic_on s" +proof - + have "\<exists>f'. (g has_field_derivative f') (at z)" if "z \<in> s" for z + proof - + obtain r where "0 < r" and r: "cball z r \<subseteq> s" + and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially" + using to_g [OF \<open>z \<in> s\<close>] by blast + have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and> f n holomorphic_on ball z r" + apply (intro eventuallyI conjI) + using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r apply blast + apply (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) + done + show ?thesis + apply (rule holomorphic_uniform_limit [OF *]) + using \<open>0 < r\<close> centre_in_ball fg + apply (auto simp: holomorphic_on_open) + done + qed + with s show ?thesis + by (simp add: holomorphic_on_open) +qed + +lemma has_complex_derivative_uniform_sequence: + fixes s :: "complex set" + assumes s: "open s" + and hfd: "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_field_derivative f' n x) (at x)" + and to_g: "\<And>x. x \<in> s + \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and> + (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)" + shows "\<exists>g'. \<forall>x \<in> s. (g has_field_derivative g' x) (at x) \<and> ((\<lambda>n. f' n x) \<longlongrightarrow> g' x) sequentially" +proof - + have y: "\<exists>y. (g has_field_derivative y) (at z) \<and> (\<lambda>n. f' n z) \<longlonglongrightarrow> y" if "z \<in> s" for z + proof - + obtain r where "0 < r" and r: "cball z r \<subseteq> s" + and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially" + using to_g [OF \<open>z \<in> s\<close>] by blast + have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and> + (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))" + apply (intro eventuallyI conjI) + apply (meson hfd holomorphic_on_imp_continuous_on holomorphic_on_open holomorphic_on_subset r s) + using ball_subset_cball hfd r apply blast + done + show ?thesis + apply (rule has_complex_derivative_uniform_limit [OF *, of g]) + using \<open>0 < r\<close> centre_in_ball fg + apply force+ + done + qed + show ?thesis + by (rule bchoice) (blast intro: y) +qed + + +subsection\<open>On analytic functions defined by a series.\<close> + +lemma series_and_derivative_comparison: + fixes s :: "complex set" + assumes s: "open s" + and h: "summable h" + and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + and to_g: "\<And>n x. \<lbrakk>N \<le> n; x \<in> s\<rbrakk> \<Longrightarrow> norm(f n x) \<le> h n" + obtains g g' where "\<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)" +proof - + obtain g where g: "\<And>e. e>0 \<Longrightarrow> \<exists>N. \<forall>n x. N \<le> n \<and> x \<in> s \<longrightarrow> dist (\<Sum>n<n. f n x) (g x) < e" + using series_comparison_uniform [OF h to_g, of N s] by force + have *: "\<exists>d>0. cball x d \<subseteq> s \<and> (\<forall>e>0. \<forall>\<^sub>F n in sequentially. \<forall>y\<in>cball x d. cmod ((\<Sum>a<n. f a y) - g y) < e)" + if "x \<in> s" for x + proof - + obtain d where "d>0" and d: "cball x d \<subseteq> s" + using open_contains_cball [of "s"] \<open>x \<in> s\<close> s by blast + then show ?thesis + apply (rule_tac x=d in exI) + apply (auto simp: dist_norm eventually_sequentially) + apply (metis g contra_subsetD dist_norm) + done + qed + have "(\<forall>x\<in>s. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> g x)" + using g by (force simp add: lim_sequentially) + moreover have "\<exists>g'. \<forall>x\<in>s. (g has_field_derivative g' x) (at x) \<and> (\<lambda>n. \<Sum>i<n. f' i x) \<longlonglongrightarrow> g' x" + by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_setsum)+ + ultimately show ?thesis + by (force simp add: sums_def conj_commute intro: that) +qed + +text\<open>A version where we only have local uniform/comparative convergence.\<close> + +lemma series_and_derivative_comparison_local: + fixes s :: "complex set" + assumes s: "open s" + and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + and to_g: "\<And>x. x \<in> s \<Longrightarrow> + \<exists>d h N. 0 < d \<and> summable h \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> norm(f n y) \<le> h n)" + shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)" +proof - + have "\<exists>y. (\<lambda>n. f n z) sums (\<Sum>n. f n z) \<and> (\<lambda>n. f' n z) sums y \<and> ((\<lambda>x. \<Sum>n. f n x) has_field_derivative y) (at z)" + if "z \<in> s" for z + proof - + obtain d h N where "0 < d" "summable h" and le_h: "\<And>n y. \<lbrakk>N \<le> n; y \<in> ball z d\<rbrakk> \<Longrightarrow> norm(f n y) \<le> h n" + using to_g \<open>z \<in> s\<close> by meson + then obtain r where "r>0" and r: "ball z r \<subseteq> ball z d \<inter> s" using \<open>z \<in> s\<close> s + by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) + have 1: "open (ball z d \<inter> s)" + by (simp add: open_Int s) + have 2: "\<And>n x. x \<in> ball z d \<inter> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + by (auto simp: hfd) + obtain g g' where gg': "\<forall>x \<in> ball z d \<inter> s. ((\<lambda>n. f n x) sums g x) \<and> + ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)" + by (auto intro: le_h series_and_derivative_comparison [OF 1 \<open>summable h\<close> hfd]) + then have "(\<lambda>n. f' n z) sums g' z" + by (meson \<open>0 < r\<close> centre_in_ball contra_subsetD r) + moreover have "(\<lambda>n. f n z) sums (\<Sum>n. f n z)" + by (metis summable_comparison_test' summable_sums centre_in_ball \<open>0 < d\<close> \<open>summable h\<close> le_h) + moreover have "((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' z) (at z)" + apply (rule_tac f=g in DERIV_transform_at [OF _ \<open>0 < r\<close>]) + apply (simp add: gg' \<open>z \<in> s\<close> \<open>0 < d\<close>) + apply (metis (full_types) contra_subsetD dist_commute gg' mem_ball r sums_unique) + done + ultimately show ?thesis by auto + qed + then show ?thesis + by (rule_tac x="\<lambda>x. suminf (\<lambda>n. f n x)" in exI) meson +qed + + +text\<open>Sometimes convenient to compare with a complex series of positive reals. (?)\<close> + +lemma series_and_derivative_comparison_complex: + fixes s :: "complex set" + assumes s: "open s" + and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + and to_g: "\<And>x. x \<in> s \<Longrightarrow> + \<exists>d h N. 0 < d \<and> summable h \<and> range h \<subseteq> nonneg_Reals \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> cmod(f n y) \<le> cmod (h n))" + shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)" +apply (rule series_and_derivative_comparison_local [OF s hfd], assumption) +apply (frule to_g) +apply (erule ex_forward) +apply (erule exE) +apply (rule_tac x="Re o h" in exI) +apply (erule ex_forward) +apply (simp add: summable_Re o_def ) +apply (elim conjE all_forward) +apply (simp add: nonneg_Reals_cmod_eq_Re image_subset_iff) +done + + +text\<open>In particular, a power series is analytic inside circle of convergence.\<close> + +lemma power_series_and_derivative_0: + fixes a :: "nat \<Rightarrow> complex" and r::real + assumes "summable (\<lambda>n. a n * r^n)" + shows "\<exists>g g'. \<forall>z. cmod z < r \<longrightarrow> + ((\<lambda>n. a n * z^n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * z^(n - 1)) sums g' z) \<and> (g has_field_derivative g' z) (at z)" +proof (cases "0 < r") + case True + have der: "\<And>n z. ((\<lambda>x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" + by (rule derivative_eq_intros | simp)+ + have y_le: "\<lbrakk>cmod (z - y) * 2 < r - cmod z\<rbrakk> \<Longrightarrow> cmod y \<le> cmod (of_real r + of_real (cmod z)) / 2" for z y + using \<open>r > 0\<close> + apply (auto simp: algebra_simps norm_mult norm_divide norm_power of_real_add [symmetric] simp del: of_real_add) + using norm_triangle_ineq2 [of y z] + apply (simp only: diff_le_eq norm_minus_commute mult_2) + done + have "summable (\<lambda>n. a n * complex_of_real r ^ n)" + using assms \<open>r > 0\<close> by simp + moreover have "\<And>z. cmod z < r \<Longrightarrow> cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" + using \<open>r > 0\<close> + by (simp add: of_real_add [symmetric] del: of_real_add) + ultimately have sum: "\<And>z. cmod z < r \<Longrightarrow> summable (\<lambda>n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" + by (rule power_series_conv_imp_absconv_weak) + have "\<exists>g g'. \<forall>z \<in> ball 0 r. (\<lambda>n. (a n) * z ^ n) sums g z \<and> + (\<lambda>n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \<and> (g has_field_derivative g' z) (at z)" + apply (rule series_and_derivative_comparison_complex [OF open_ball der]) + apply (rule_tac x="(r - norm z)/2" in exI) + apply (simp add: dist_norm) + apply (rule_tac x="\<lambda>n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) + using \<open>r > 0\<close> + apply (auto simp: sum nonneg_Reals_divide_I) + apply (rule_tac x=0 in exI) + apply (force simp: norm_mult norm_divide norm_power intro!: mult_left_mono power_mono y_le) + done + then show ?thesis + by (simp add: dist_0_norm ball_def) +next + case False then show ?thesis + apply (simp add: not_less) + using less_le_trans norm_not_less_zero by blast +qed + +proposition power_series_and_derivative: + fixes a :: "nat \<Rightarrow> complex" and r::real + assumes "summable (\<lambda>n. a n * r^n)" + obtains g g' where "\<forall>z \<in> ball w r. + ((\<lambda>n. a n * (z - w) ^ n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \<and> + (g has_field_derivative g' z) (at z)" + using power_series_and_derivative_0 [OF assms] + apply clarify + apply (rule_tac g="(\<lambda>z. g(z - w))" in that) + using DERIV_shift [where z="-w"] + apply (auto simp: norm_minus_commute Ball_def dist_norm) + done + +proposition power_series_holomorphic: + assumes "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>n. a n*(w - z)^n) sums f w)" + shows "f holomorphic_on ball z r" +proof - + have "\<exists>f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w + proof - + have inb: "z + complex_of_real ((dist z w + r) / 2) \<in> ball z r" + proof - + have wz: "cmod (w - z) < r" using w + by (auto simp: divide_simps dist_norm norm_minus_commute) + then have "0 \<le> r" + by (meson less_eq_real_def norm_ge_zero order_trans) + show ?thesis + using w by (simp add: dist_norm \<open>0\<le>r\<close> of_real_add [symmetric] del: of_real_add) + qed + have sum: "summable (\<lambda>n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" + using assms [OF inb] by (force simp add: summable_def dist_norm) + obtain g g' where gg': "\<And>u. u \<in> ball z ((cmod (z - w) + r) / 2) \<Longrightarrow> + (\<lambda>n. a n * (u - z) ^ n) sums g u \<and> + (\<lambda>n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \<and> (g has_field_derivative g' u) (at u)" + by (rule power_series_and_derivative [OF sum, of z]) fastforce + have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u + proof - + have less: "cmod (z - u) * 2 < cmod (z - w) + r" + using that dist_triangle2 [of z u w] + by (simp add: dist_norm [symmetric] algebra_simps) + show ?thesis + apply (rule sums_unique2 [of "\<lambda>n. a n*(u - z)^n"]) + using gg' [of u] less w + apply (auto simp: assms dist_norm) + done + qed + show ?thesis + apply (rule_tac x="g' w" in exI) + apply (rule DERIV_transform_at [where f=g and d="(r - norm(z - w))/2"]) + using w gg' [of w] + apply (auto simp: dist_norm) + done + qed + then show ?thesis by (simp add: holomorphic_on_open) +qed + +corollary holomorphic_iff_power_series: + "f holomorphic_on ball z r \<longleftrightarrow> + (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" + apply (intro iffI ballI) + using holomorphic_power_series apply force + apply (rule power_series_holomorphic [where a = "\<lambda>n. (deriv ^^ n) f z / (fact n)"]) + apply force + done + +corollary power_series_analytic: + "(\<And>w. w \<in> ball z r \<Longrightarrow> (\<lambda>n. a n*(w - z)^n) sums f w) \<Longrightarrow> f analytic_on ball z r" + by (force simp add: analytic_on_open intro!: power_series_holomorphic) + +corollary analytic_iff_power_series: + "f analytic_on ball z r \<longleftrightarrow> + (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" + by (simp add: analytic_on_open holomorphic_iff_power_series) + + +subsection\<open>Equality between holomorphic functions, on open ball then connected set.\<close> + +lemma holomorphic_fun_eq_on_ball: + "\<lbrakk>f holomorphic_on ball z r; g holomorphic_on ball z r; + w \<in> ball z r; + \<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z\<rbrakk> + \<Longrightarrow> f w = g w" + apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) + apply (auto simp: holomorphic_iff_power_series) + done + +lemma holomorphic_fun_eq_0_on_ball: + "\<lbrakk>f holomorphic_on ball z r; w \<in> ball z r; + \<And>n. (deriv ^^ n) f z = 0\<rbrakk> + \<Longrightarrow> f w = 0" + apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) + apply (auto simp: holomorphic_iff_power_series) + done + +lemma holomorphic_fun_eq_0_on_connected: + assumes holf: "f holomorphic_on s" and "open s" + and cons: "connected s" + and der: "\<And>n. (deriv ^^ n) f z = 0" + and "z \<in> s" "w \<in> s" + shows "f w = 0" +proof - + have *: "\<And>x e. \<lbrakk> \<forall>xa. (deriv ^^ xa) f x = 0; ball x e \<subseteq> s\<rbrakk> + \<Longrightarrow> ball x e \<subseteq> (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})" + apply auto + apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) + apply (rule holomorphic_on_subset [OF holf], simp_all) + by (metis funpow_add o_apply) + have 1: "openin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})" + apply (rule open_subset, force) + using \<open>open s\<close> + apply (simp add: open_contains_ball Ball_def) + apply (erule all_forward) + using "*" by auto blast+ + have 2: "closedin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})" + using assms + by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) + obtain e where "e>0" and e: "ball w e \<subseteq> s" using openE [OF \<open>open s\<close> \<open>w \<in> s\<close>] . + then have holfb: "f holomorphic_on ball w e" + using holf holomorphic_on_subset by blast + have 3: "(\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}) = s \<Longrightarrow> f w = 0" + using \<open>e>0\<close> e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) + show ?thesis + using cons der \<open>z \<in> s\<close> + apply (simp add: connected_clopen) + apply (drule_tac x="\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}" in spec) + apply (auto simp: 1 2 3) + done +qed + +lemma holomorphic_fun_eq_on_connected: + assumes "f holomorphic_on s" "g holomorphic_on s" and "open s" "connected s" + and "\<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z" + and "z \<in> s" "w \<in> s" + shows "f w = g w" + apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>x. f x - g x" s z, simplified]) + apply (intro assms holomorphic_intros) + using assms apply simp_all + apply (subst higher_deriv_diff, auto) + done + +lemma holomorphic_fun_eq_const_on_connected: + assumes holf: "f holomorphic_on s" and "open s" + and cons: "connected s" + and der: "\<And>n. 0 < n \<Longrightarrow> (deriv ^^ n) f z = 0" + and "z \<in> s" "w \<in> s" + shows "f w = f z" + apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>w. f w - f z" s z, simplified]) + apply (intro assms holomorphic_intros) + using assms apply simp_all + apply (subst higher_deriv_diff) + apply (intro holomorphic_intros | simp)+ + done + + +subsection\<open>Some basic lemmas about poles/singularities.\<close> + +lemma pole_lemma: + assumes holf: "f holomorphic_on s" and a: "a \<in> interior s" + shows "(\<lambda>z. if z = a then deriv f a + else (f z - f a) / (z - a)) holomorphic_on s" (is "?F holomorphic_on s") +proof - + have F1: "?F field_differentiable (at u within s)" if "u \<in> s" "u \<noteq> a" for u + proof - + have fcd: "f field_differentiable at u within s" + using holf holomorphic_on_def by (simp add: \<open>u \<in> s\<close>) + have cd: "(\<lambda>z. (f z - f a) / (z - a)) field_differentiable at u within s" + by (rule fcd derivative_intros | simp add: that)+ + have "0 < dist a u" using that dist_nz by blast + then show ?thesis + by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \<open>u \<in> s\<close>) + qed + have F2: "?F field_differentiable at a" if "0 < e" "ball a e \<subseteq> s" for e + proof - + have holfb: "f holomorphic_on ball a e" + by (rule holomorphic_on_subset [OF holf \<open>ball a e \<subseteq> s\<close>]) + have 2: "?F holomorphic_on ball a e - {a}" + apply (rule holomorphic_on_subset [where s = "s - {a}"]) + apply (simp add: holomorphic_on_def field_differentiable_def [symmetric]) + using mem_ball that + apply (auto intro: F1 field_differentiable_within_subset) + done + have "isCont (\<lambda>z. if z = a then deriv f a else (f z - f a) / (z - a)) x" + if "dist a x < e" for x + proof (cases "x=a") + case True then show ?thesis + using holfb \<open>0 < e\<close> + apply (simp add: holomorphic_on_open field_differentiable_def [symmetric]) + apply (drule_tac x=a in bspec) + apply (auto simp: DERIV_deriv_iff_field_differentiable [symmetric] continuous_at DERIV_iff2 + elim: rev_iffD1 [OF _ LIM_equal]) + done + next + case False with 2 that show ?thesis + by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) + qed + then have 1: "continuous_on (ball a e) ?F" + by (clarsimp simp: continuous_on_eq_continuous_at) + have "?F holomorphic_on ball a e" + by (auto intro: no_isolated_singularity [OF 1 2]) + with that show ?thesis + by (simp add: holomorphic_on_open field_differentiable_def [symmetric] + field_differentiable_at_within) + qed + show ?thesis + proof + fix x assume "x \<in> s" show "?F field_differentiable at x within s" + proof (cases "x=a") + case True then show ?thesis + using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) + next + case False with F1 \<open>x \<in> s\<close> + show ?thesis by blast + qed + qed +qed + +proposition pole_theorem: + assumes holg: "g holomorphic_on s" and a: "a \<in> interior s" + and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z" + shows "(\<lambda>z. if z = a then deriv g a + else f z - g a/(z - a)) holomorphic_on s" + using pole_lemma [OF holg a] + by (rule holomorphic_transform) (simp add: eq divide_simps) + +lemma pole_lemma_open: + assumes "f holomorphic_on s" "open s" + shows "(\<lambda>z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on s" +proof (cases "a \<in> s") + case True with assms interior_eq pole_lemma + show ?thesis by fastforce +next + case False with assms show ?thesis + apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) + apply (rule field_differentiable_transform_within [where f = "\<lambda>z. (f z - f a)/(z - a)" and d = 1]) + apply (rule derivative_intros | force)+ + done +qed + +proposition pole_theorem_open: + assumes holg: "g holomorphic_on s" and s: "open s" + and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z" + shows "(\<lambda>z. if z = a then deriv g a + else f z - g a/(z - a)) holomorphic_on s" + using pole_lemma_open [OF holg s] + by (rule holomorphic_transform) (auto simp: eq divide_simps) + +proposition pole_theorem_0: + assumes holg: "g holomorphic_on s" and a: "a \<in> interior s" + and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f holomorphic_on s" + using pole_theorem [OF holg a eq] + by (rule holomorphic_transform) (auto simp: eq divide_simps) + +proposition pole_theorem_open_0: + assumes holg: "g holomorphic_on s" and s: "open s" + and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f holomorphic_on s" + using pole_theorem_open [OF holg s eq] + by (rule holomorphic_transform) (auto simp: eq divide_simps) + +lemma pole_theorem_analytic: + assumes g: "g analytic_on s" + and eq: "\<And>z. z \<in> s + \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)" + shows "(\<lambda>z. if z = a then deriv g a + else f z - g a/(z - a)) analytic_on s" +using g +apply (simp add: analytic_on_def Ball_def) +apply (safe elim!: all_forward dest!: eq) +apply (rule_tac x="min d e" in exI, simp) +apply (rule pole_theorem_open) +apply (auto simp: holomorphic_on_subset subset_ball) +done + +lemma pole_theorem_analytic_0: + assumes g: "g analytic_on s" + and eq: "\<And>z. z \<in> s \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f analytic_on s" +proof - + have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f" + by auto + show ?thesis + using pole_theorem_analytic [OF g eq] by simp +qed + +lemma pole_theorem_analytic_open_superset: + assumes g: "g analytic_on s" and "s \<subseteq> t" "open t" + and eq: "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z" + shows "(\<lambda>z. if z = a then deriv g a + else f z - g a/(z - a)) analytic_on s" + apply (rule pole_theorem_analytic [OF g]) + apply (rule openE [OF \<open>open t\<close>]) + using assms eq by auto + +lemma pole_theorem_analytic_open_superset_0: + assumes g: "g analytic_on s" "s \<subseteq> t" "open t" "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f analytic_on s" +proof - + have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f" + by auto + have "(\<lambda>z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on s" + by (rule pole_theorem_analytic_open_superset [OF g]) + then show ?thesis by simp +qed + + + +subsection\<open>General, homology form of Cauchy's theorem.\<close> + +text\<open>Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\<close> + +lemma contour_integral_continuous_on_linepath_2D: + assumes "open u" and cont_dw: "\<And>w. w \<in> u \<Longrightarrow> F w contour_integrable_on (linepath a b)" + and cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). F x y)" + and abu: "closed_segment a b \<subseteq> u" + shows "continuous_on u (\<lambda>w. contour_integral (linepath a b) (F w))" +proof - + have *: "\<exists>d>0. \<forall>x'\<in>u. dist x' w < d \<longrightarrow> + dist (contour_integral (linepath a b) (F x')) + (contour_integral (linepath a b) (F w)) \<le> \<epsilon>" + if "w \<in> u" "0 < \<epsilon>" "a \<noteq> b" for w \<epsilon> + proof - + obtain \<delta> where "\<delta>>0" and \<delta>: "cball w \<delta> \<subseteq> u" using open_contains_cball \<open>open u\<close> \<open>w \<in> u\<close> by force + let ?TZ = "{(t,z) |t z. t \<in> cball w \<delta> \<and> z \<in> closed_segment a b}" + have "uniformly_continuous_on ?TZ (\<lambda>(x,y). F x y)" + apply (rule compact_uniformly_continuous) + apply (rule continuous_on_subset[OF cond_uu]) + using abu \<delta> + apply (auto simp: compact_Times simp del: mem_cball) + done + then obtain \<eta> where "\<eta>>0" + and \<eta>: "\<And>x x'. \<lbrakk>x\<in>?TZ; x'\<in>?TZ; dist x' x < \<eta>\<rbrakk> \<Longrightarrow> + dist ((\<lambda>(x,y). F x y) x') ((\<lambda>(x,y). F x y) x) < \<epsilon>/norm(b - a)" + apply (rule uniformly_continuous_onE [where e = "\<epsilon>/norm(b - a)"]) + using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> by auto + have \<eta>: "\<lbrakk>norm (w - x1) \<le> \<delta>; x2 \<in> closed_segment a b; + norm (w - x1') \<le> \<delta>; x2' \<in> closed_segment a b; norm ((x1', x2') - (x1, x2)) < \<eta>\<rbrakk> + \<Longrightarrow> norm (F x1' x2' - F x1 x2) \<le> \<epsilon> / cmod (b - a)" + for x1 x2 x1' x2' + using \<eta> [of "(x1,x2)" "(x1',x2')"] by (force simp add: dist_norm) + have le_ee: "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>" + if "x' \<in> u" "cmod (x' - w) < \<delta>" "cmod (x' - w) < \<eta>" for x' + proof - + have "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>/norm(b - a) * norm(b - a)" + apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \<eta>]) + apply (rule contour_integrable_diff [OF cont_dw cont_dw]) + using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> \<open>0 < \<delta>\<close> \<open>w \<in> u\<close> that + apply (auto simp: norm_minus_commute) + done + also have "... = \<epsilon>" using \<open>a \<noteq> b\<close> by simp + finally show ?thesis . + qed + show ?thesis + apply (rule_tac x="min \<delta> \<eta>" in exI) + using \<open>0 < \<delta>\<close> \<open>0 < \<eta>\<close> + apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \<open>w \<in> u\<close> intro: le_ee) + done + qed + show ?thesis + apply (rule continuous_onI) + apply (cases "a=b") + apply (auto intro: *) + done +qed + +text\<open>This version has @{term"polynomial_function \<gamma>"} as an additional assumption.\<close> +lemma Cauchy_integral_formula_global_weak: + assumes u: "open u" and holf: "f holomorphic_on u" + and z: "z \<in> u" and \<gamma>: "polynomial_function \<gamma>" + and pasz: "path_image \<gamma> \<subseteq> u - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + and zero: "\<And>w. w \<notin> u \<Longrightarrow> winding_number \<gamma> w = 0" + shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>" +proof - + obtain \<gamma>' where pf\<gamma>': "polynomial_function \<gamma>'" and \<gamma>': "\<And>x. (\<gamma> has_vector_derivative (\<gamma>' x)) (at x)" + using has_vector_derivative_polynomial_function [OF \<gamma>] by blast + then have "bounded(path_image \<gamma>')" + by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) + then obtain B where "B>0" and B: "\<And>x. x \<in> path_image \<gamma>' \<Longrightarrow> norm x \<le> B" + using bounded_pos by force + define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w + define v where "v = {w. w \<notin> path_image \<gamma> \<and> winding_number \<gamma> w = 0}" + have "path \<gamma>" "valid_path \<gamma>" using \<gamma> + by (auto simp: path_polynomial_function valid_path_polynomial_function) + then have ov: "open v" + by (simp add: v_def open_winding_number_levelsets loop) + have uv_Un: "u \<union> v = UNIV" + using pasz zero by (auto simp: v_def) + have conf: "continuous_on u f" + by (metis holf holomorphic_on_imp_continuous_on) + have hol_d: "(d y) holomorphic_on u" if "y \<in> u" for y + proof - + have *: "(\<lambda>c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on u" + by (simp add: holf pole_lemma_open u) + then have "isCont (\<lambda>x. if x = y then deriv f y else (f x - f y) / (x - y)) y" + using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that u by fastforce + then have "continuous_on u (d y)" + apply (simp add: d_def continuous_on_eq_continuous_at u, clarify) + using * holomorphic_on_def + by (meson field_differentiable_within_open field_differentiable_imp_continuous_at u) + moreover have "d y holomorphic_on u - {y}" + apply (simp add: d_def holomorphic_on_open u open_delete field_differentiable_def [symmetric]) + apply (intro ballI) + apply (rename_tac w) + apply (rule_tac d="dist w y" and f = "\<lambda>w. (f w - f y)/(w - y)" in field_differentiable_transform_within) + apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros) + using analytic_on_imp_differentiable_at analytic_on_open holf u apply blast + done + ultimately show ?thesis + by (rule no_isolated_singularity) (auto simp: u) + qed + have cint_fxy: "(\<lambda>x. (f x - f y) / (x - y)) contour_integrable_on \<gamma>" if "y \<notin> path_image \<gamma>" for y + apply (rule contour_integrable_holomorphic_simple [where s = "u-{y}"]) + using \<open>valid_path \<gamma>\<close> pasz + apply (auto simp: u open_delete) + apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] | + force simp add: that)+ + done + define h where + "h z = (if z \<in> u then contour_integral \<gamma> (d z) else contour_integral \<gamma> (\<lambda>w. f w/(w - z)))" for z + have U: "\<And>z. z \<in> u \<Longrightarrow> ((d z) has_contour_integral h z) \<gamma>" + apply (simp add: h_def) + apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]]) + using u pasz \<open>valid_path \<gamma>\<close> + apply (auto intro: holomorphic_on_imp_continuous_on hol_d) + done + have V: "((\<lambda>w. f w / (w - z)) has_contour_integral h z) \<gamma>" if z: "z \<in> v" for z + proof - + have 0: "0 = (f z) * 2 * of_real (2 * pi) * \<i> * winding_number \<gamma> z" + using v_def z by auto + then have "((\<lambda>x. 1 / (x - z)) has_contour_integral 0) \<gamma>" + using z v_def has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close>] by fastforce + then have "((\<lambda>x. f z * (1 / (x - z))) has_contour_integral 0) \<gamma>" + using has_contour_integral_lmul by fastforce + then have "((\<lambda>x. f z / (x - z)) has_contour_integral 0) \<gamma>" + by (simp add: divide_simps) + moreover have "((\<lambda>x. (f x - f z) / (x - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>" + using z + apply (auto simp: v_def) + apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) + done + ultimately have *: "((\<lambda>x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \<gamma> (d z))) \<gamma>" + by (rule has_contour_integral_add) + have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>" + if "z \<in> u" + using * by (auto simp: divide_simps has_contour_integral_eq) + moreover have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (\<lambda>w. f w / (w - z))) \<gamma>" + if "z \<notin> u" + apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]]) + using u pasz \<open>valid_path \<gamma>\<close> that + apply (auto intro: holomorphic_on_imp_continuous_on hol_d) + apply (rule continuous_intros conf holomorphic_intros holf | force)+ + done + ultimately show ?thesis + using z by (simp add: h_def) + qed + have znot: "z \<notin> path_image \<gamma>" + using pasz by blast + obtain d0 where "d0>0" and d0: "\<And>x y. x \<in> path_image \<gamma> \<Longrightarrow> y \<in> - u \<Longrightarrow> d0 \<le> dist x y" + using separate_compact_closed [of "path_image \<gamma>" "-u"] pasz u + by (fastforce simp add: \<open>path \<gamma>\<close> compact_path_image) + obtain dd where "0 < dd" and dd: "{y + k | y k. y \<in> path_image \<gamma> \<and> k \<in> ball 0 dd} \<subseteq> u" + apply (rule that [of "d0/2"]) + using \<open>0 < d0\<close> + apply (auto simp: dist_norm dest: d0) + done + have "\<And>x x'. \<lbrakk>x \<in> path_image \<gamma>; dist x x' * 2 < dd\<rbrakk> \<Longrightarrow> \<exists>y k. x' = y + k \<and> y \<in> path_image \<gamma> \<and> dist 0 k * 2 \<le> dd" + apply (rule_tac x=x in exI) + apply (rule_tac x="x'-x" in exI) + apply (force simp add: dist_norm) + done + then have 1: "path_image \<gamma> \<subseteq> interior {y + k |y k. y \<in> path_image \<gamma> \<and> k \<in> cball 0 (dd / 2)}" + apply (clarsimp simp add: mem_interior) + using \<open>0 < dd\<close> + apply (rule_tac x="dd/2" in exI, auto) + done + obtain t where "compact t" and subt: "path_image \<gamma> \<subseteq> interior t" and t: "t \<subseteq> u" + apply (rule that [OF _ 1]) + apply (fastforce simp add: \<open>valid_path \<gamma>\<close> compact_valid_path_image intro!: compact_sums) + apply (rule order_trans [OF _ dd]) + using \<open>0 < dd\<close> by fastforce + obtain L where "L>0" + and L: "\<And>f B. \<lbrakk>f holomorphic_on interior t; \<And>z. z\<in>interior t \<Longrightarrow> cmod (f z) \<le> B\<rbrakk> \<Longrightarrow> + cmod (contour_integral \<gamma> f) \<le> L * B" + using contour_integral_bound_exists [OF open_interior \<open>valid_path \<gamma>\<close> subt] + by blast + have "bounded(f ` t)" + by (meson \<open>compact t\<close> compact_continuous_image compact_imp_bounded conf continuous_on_subset t) + then obtain D where "D>0" and D: "\<And>x. x \<in> t \<Longrightarrow> norm (f x) \<le> D" + by (auto simp: bounded_pos) + obtain C where "C>0" and C: "\<And>x. x \<in> t \<Longrightarrow> norm x \<le> C" + using \<open>compact t\<close> bounded_pos compact_imp_bounded by force + have "dist (h y) 0 \<le> e" if "0 < e" and le: "D * L / e + C \<le> cmod y" for e y + proof - + have "D * L / e > 0" using \<open>D>0\<close> \<open>L>0\<close> \<open>e>0\<close> by simp + with le have ybig: "norm y > C" by force + with C have "y \<notin> t" by force + then have ynot: "y \<notin> path_image \<gamma>" + using subt interior_subset by blast + have [simp]: "winding_number \<gamma> y = 0" + apply (rule winding_number_zero_outside [of _ "cball 0 C"]) + using ybig interior_subset subt + apply (force simp add: loop \<open>path \<gamma>\<close> dist_norm intro!: C)+ + done + have [simp]: "h y = contour_integral \<gamma> (\<lambda>w. f w/(w - y))" + by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) + have holint: "(\<lambda>w. f w / (w - y)) holomorphic_on interior t" + apply (rule holomorphic_on_divide) + using holf holomorphic_on_subset interior_subset t apply blast + apply (rule holomorphic_intros)+ + using \<open>y \<notin> t\<close> interior_subset by auto + have leD: "cmod (f z / (z - y)) \<le> D * (e / L / D)" if z: "z \<in> interior t" for z + proof - + have "D * L / e + cmod z \<le> cmod y" + using le C [of z] z using interior_subset by force + then have DL2: "D * L / e \<le> cmod (z - y)" + using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) + have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" + by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) + also have "... \<le> D * (e / L / D)" + apply (rule mult_mono) + using that D interior_subset apply blast + using \<open>L>0\<close> \<open>e>0\<close> \<open>D>0\<close> DL2 + apply (auto simp: norm_divide divide_simps algebra_simps) + done + finally show ?thesis . + qed + have "dist (h y) 0 = cmod (contour_integral \<gamma> (\<lambda>w. f w / (w - y)))" + by (simp add: dist_norm) + also have "... \<le> L * (D * (e / L / D))" + by (rule L [OF holint leD]) + also have "... = e" + using \<open>L>0\<close> \<open>0 < D\<close> by auto + finally show ?thesis . + qed + then have "(h \<longlongrightarrow> 0) at_infinity" + by (meson Lim_at_infinityI) + moreover have "h holomorphic_on UNIV" + proof - + have con_ff: "continuous (at (x,z)) (\<lambda>(x,y). (f y - f x) / (y - x))" + if "x \<in> u" "z \<in> u" "x \<noteq> z" for x z + using that conf + apply (simp add: split_def continuous_on_eq_continuous_at u) + apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ + done + have con_fstsnd: "continuous_on UNIV (\<lambda>x. (fst x - snd x) ::complex)" + by (rule continuous_intros)+ + have open_uu_Id: "open (u \<times> u - Id)" + apply (rule open_Diff) + apply (simp add: open_Times u) + using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] + apply (auto simp: Id_fstsnd_eq algebra_simps) + done + have con_derf: "continuous (at z) (deriv f)" if "z \<in> u" for z + apply (rule continuous_on_interior [of u]) + apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on u) + by (simp add: interior_open that u) + have tendsto_f': "((\<lambda>(x,y). if y = x then deriv f (x) + else (f (y) - f (x)) / (y - x)) \<longlongrightarrow> deriv f x) + (at (x, x) within u \<times> u)" if "x \<in> u" for x + proof (rule Lim_withinI) + fix e::real assume "0 < e" + obtain k1 where "k1>0" and k1: "\<And>x'. norm (x' - x) \<le> k1 \<Longrightarrow> norm (deriv f x' - deriv f x) < e" + using \<open>0 < e\<close> continuous_within_E [OF con_derf [OF \<open>x \<in> u\<close>]] + by (metis UNIV_I dist_norm) + obtain k2 where "k2>0" and k2: "ball x k2 \<subseteq> u" by (blast intro: openE [OF u] \<open>x \<in> u\<close>) + have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e" + if "z' \<noteq> x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" + for x' z' + proof - + have cs_less: "w \<in> closed_segment x' z' \<Longrightarrow> cmod (w - x) \<le> norm (x'-x, z'-x)" for w + apply (drule segment_furthest_le [where y=x]) + by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) + have derf_le: "w \<in> closed_segment x' z' \<Longrightarrow> z' \<noteq> x' \<Longrightarrow> cmod (deriv f w - deriv f x) \<le> e" for w + by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) + have f_has_der: "\<And>x. x \<in> u \<Longrightarrow> (f has_field_derivative deriv f x) (at x within u)" + by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def u) + have "closed_segment x' z' \<subseteq> u" + by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) + then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" + using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp + then have *: "((\<lambda>x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" + by (rule has_contour_integral_div) + have "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e/norm(z' - x') * norm(z' - x')" + apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) + using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] + \<open>e > 0\<close> \<open>z' \<noteq> x'\<close> + apply (auto simp: norm_divide divide_simps derf_le) + done + also have "... \<le> e" using \<open>0 < e\<close> by simp + finally show ?thesis . + qed + show "\<exists>d>0. \<forall>xa\<in>u \<times> u. + 0 < dist xa (x, x) \<and> dist xa (x, x) < d \<longrightarrow> + dist (case xa of (x, y) \<Rightarrow> if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \<le> e" + apply (rule_tac x="min k1 k2" in exI) + using \<open>k1>0\<close> \<open>k2>0\<close> \<open>e>0\<close> + apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) + done + qed + have con_pa_f: "continuous_on (path_image \<gamma>) f" + by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt t) + have le_B: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (vector_derivative \<gamma> (at t)) \<le> B" + apply (rule B) + using \<gamma>' using path_image_def vector_derivative_at by fastforce + have f_has_cint: "\<And>w. w \<in> v - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f u / (u - w) ^ 1) has_contour_integral h w) \<gamma>" + by (simp add: V) + have cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). d x y)" + apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') + apply (simp add: tendsto_within_open_NO_MATCH open_Times u, clarify) + apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\<lambda>(x,y). (f y - f x) / (y - x))"]) + using con_ff + apply (auto simp: continuous_within) + done + have hol_dw: "(\<lambda>z. d z w) holomorphic_on u" if "w \<in> u" for w + proof - + have "continuous_on u ((\<lambda>(x,y). d x y) o (\<lambda>z. (w,z)))" + by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ + then have *: "continuous_on u (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z))" + by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) + have **: "\<And>x. \<lbrakk>x \<in> u; x \<noteq> w\<rbrakk> \<Longrightarrow> (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" + apply (rule_tac f = "\<lambda>x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) + apply (rule u derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp add: dist_commute)+ + done + show ?thesis + unfolding d_def + apply (rule no_isolated_singularity [OF * _ u, where k = "{w}"]) + apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff u **) + done + qed + { fix a b + assume abu: "closed_segment a b \<subseteq> u" + then have "\<And>w. w \<in> u \<Longrightarrow> (\<lambda>z. d z w) contour_integrable_on (linepath a b)" + by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) + then have cont_cint_d: "continuous_on u (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" + apply (rule contour_integral_continuous_on_linepath_2D [OF \<open>open u\<close> _ _ abu]) + apply (auto simp: intro: continuous_on_swap_args cond_uu) + done + have cont_cint_d\<gamma>: "continuous_on {0..1} ((\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) o \<gamma>)" + apply (rule continuous_on_compose) + using \<open>path \<gamma>\<close> path_def pasz + apply (auto intro!: continuous_on_subset [OF cont_cint_d]) + apply (force simp add: path_image_def) + done + have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>" + apply (simp add: contour_integrable_on) + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF cont_cint_d\<gamma> [unfolded o_def]]) + using pf\<gamma>' + by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \<gamma>']) + have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\<lambda>z. contour_integral \<gamma> (d z))" + using abu by (force simp add: h_def intro: contour_integral_eq) + also have "... = contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" + apply (rule contour_integral_swap) + apply (rule continuous_on_subset [OF cond_uu]) + using abu pasz \<open>valid_path \<gamma>\<close> + apply (auto intro!: continuous_intros) + by (metis \<gamma>' continuous_on_eq path_def path_polynomial_function pf\<gamma>' vector_derivative_at) + finally have cint_h_eq: + "contour_integral (linepath a b) h = + contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" . + note cint_cint cint_h_eq + } note cint_h = this + have conthu: "continuous_on u h" + proof (simp add: continuous_on_sequentially, clarify) + fix a x + assume x: "x \<in> u" and au: "\<forall>n. a n \<in> u" and ax: "a \<longlonglongrightarrow> x" + then have A1: "\<forall>\<^sub>F n in sequentially. d (a n) contour_integrable_on \<gamma>" + by (meson U contour_integrable_on_def eventuallyI) + obtain dd where "dd>0" and dd: "cball x dd \<subseteq> u" using open_contains_cball u x by force + have A2: "\<forall>\<^sub>F n in sequentially. \<forall>xa\<in>path_image \<gamma>. cmod (d (a n) xa - d x xa) < ee" if "0 < ee" for ee + proof - + let ?ddpa = "{(w,z) |w z. w \<in> cball x dd \<and> z \<in> path_image \<gamma>}" + have "uniformly_continuous_on ?ddpa (\<lambda>(x,y). d x y)" + apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) + using dd pasz \<open>valid_path \<gamma>\<close> + apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) + done + then obtain kk where "kk>0" + and kk: "\<And>x x'. \<lbrakk>x\<in>?ddpa; x'\<in>?ddpa; dist x' x < kk\<rbrakk> \<Longrightarrow> + dist ((\<lambda>(x,y). d x y) x') ((\<lambda>(x,y). d x y) x) < ee" + apply (rule uniformly_continuous_onE [where e = ee]) + using \<open>0 < ee\<close> by auto + have kk: "\<lbrakk>norm (w - x) \<le> dd; z \<in> path_image \<gamma>; norm ((w, z) - (x, z)) < kk\<rbrakk> \<Longrightarrow> norm (d w z - d x z) < ee" + for w z + using \<open>dd>0\<close> kk [of "(x,z)" "(w,z)"] by (force simp add: norm_minus_commute dist_norm) + show ?thesis + using ax unfolding lim_sequentially eventually_sequentially + apply (drule_tac x="min dd kk" in spec) + using \<open>dd > 0\<close> \<open>kk > 0\<close> + apply (fastforce simp: kk dist_norm) + done + qed + have tendsto_hx: "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> h x" + apply (simp add: contour_integral_unique [OF U, symmetric] x) + apply (rule contour_integral_uniform_limit [OF A1 A2 le_B]) + apply (auto simp: \<open>valid_path \<gamma>\<close>) + done + then show "(h \<circ> a) \<longlonglongrightarrow> h x" + by (simp add: h_def x au o_def) + qed + show ?thesis + proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) + fix z0 + consider "z0 \<in> v" | "z0 \<in> u" using uv_Un by blast + then show "h field_differentiable at z0" + proof cases + assume "z0 \<in> v" then show ?thesis + using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \<open>valid_path \<gamma>\<close> + by (auto simp: field_differentiable_def v_def) + next + assume "z0 \<in> u" then + obtain e where "e>0" and e: "ball z0 e \<subseteq> u" by (blast intro: openE [OF u]) + have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" + if abc_subset: "convex hull {a, b, c} \<subseteq> ball z0 e" for a b c + proof - + have *: "\<And>x1 x2 z. z \<in> u \<Longrightarrow> closed_segment x1 x2 \<subseteq> u \<Longrightarrow> (\<lambda>w. d w z) contour_integrable_on linepath x1 x2" + using hol_dw holomorphic_on_imp_continuous_on u + by (auto intro!: contour_integrable_holomorphic_simple) + have abc: "closed_segment a b \<subseteq> u" "closed_segment b c \<subseteq> u" "closed_segment c a \<subseteq> u" + using that e segments_subset_convex_hull by fastforce+ + have eq0: "\<And>w. w \<in> u \<Longrightarrow> contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\<lambda>z. d z w) = 0" + apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) + apply (rule holomorphic_on_subset [OF hol_dw]) + using e abc_subset by auto + have "contour_integral \<gamma> + (\<lambda>x. contour_integral (linepath a b) (\<lambda>z. d z x) + + (contour_integral (linepath b c) (\<lambda>z. d z x) + + contour_integral (linepath c a) (\<lambda>z. d z x))) = 0" + apply (rule contour_integral_eq_0) + using abc pasz u + apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ + done + then show ?thesis + by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) + qed + show ?thesis + using e \<open>e > 0\<close> + by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic + Morera_triangle continuous_on_subset [OF conthu] *) + qed + qed + qed + ultimately have [simp]: "h z = 0" for z + by (meson Liouville_weak) + have "((\<lambda>w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z) \<gamma>" + by (rule has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close> znot]) + then have "((\<lambda>w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>" + by (metis mult.commute has_contour_integral_lmul) + then have 1: "((\<lambda>w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>" + by (simp add: divide_simps) + moreover have 2: "((\<lambda>w. (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>" + using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\<lambda>w. (f w - f z)/(w - z)"]) + show ?thesis + using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) +qed + + +theorem Cauchy_integral_formula_global: + assumes s: "open s" and holf: "f holomorphic_on s" + and z: "z \<in> s" and vpg: "valid_path \<gamma>" + and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0" + shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>" +proof - + have "path \<gamma>" using vpg by (blast intro: valid_path_imp_path) + have hols: "(\<lambda>w. f w / (w - z)) holomorphic_on s - {z}" "(\<lambda>w. 1 / (w - z)) holomorphic_on s - {z}" + by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ + then have cint_fw: "(\<lambda>w. f w / (w - z)) contour_integrable_on \<gamma>" + by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete s vpg pasz) + obtain d where "d>0" + and d: "\<And>g h. \<lbrakk>valid_path g; valid_path h; \<forall>t\<in>{0..1}. cmod (g t - \<gamma> t) < d \<and> cmod (h t - \<gamma> t) < d; + pathstart h = pathstart g \<and> pathfinish h = pathfinish g\<rbrakk> + \<Longrightarrow> path_image h \<subseteq> s - {z} \<and> (\<forall>f. f holomorphic_on s - {z} \<longrightarrow> contour_integral h f = contour_integral g f)" + using contour_integral_nearby_ends [OF _ \<open>path \<gamma>\<close> pasz] s by (simp add: open_Diff) metis + obtain p where polyp: "polynomial_function p" + and ps: "pathstart p = pathstart \<gamma>" and pf: "pathfinish p = pathfinish \<gamma>" and led: "\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < d" + using path_approx_polynomial_function [OF \<open>path \<gamma>\<close> \<open>d > 0\<close>] by blast + then have ploop: "pathfinish p = pathstart p" using loop by auto + have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast + have [simp]: "z \<notin> path_image \<gamma>" using pasz by blast + have paps: "path_image p \<subseteq> s - {z}" and cint_eq: "(\<And>f. f holomorphic_on s - {z} \<Longrightarrow> contour_integral p f = contour_integral \<gamma> f)" + using pf ps led d [OF vpg vpp] \<open>d > 0\<close> by auto + have wn_eq: "winding_number p z = winding_number \<gamma> z" + using vpp paps + by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) + have "winding_number p w = winding_number \<gamma> w" if "w \<notin> s" for w + proof - + have hol: "(\<lambda>v. 1 / (v - w)) holomorphic_on s - {z}" + using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) + have "w \<notin> path_image p" "w \<notin> path_image \<gamma>" using paps pasz that by auto + then show ?thesis + using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) + qed + then have wn0: "\<And>w. w \<notin> s \<Longrightarrow> winding_number p w = 0" + by (simp add: zero) + show ?thesis + using Cauchy_integral_formula_global_weak [OF s holf z polyp paps ploop wn0] hols + by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) +qed + +theorem Cauchy_theorem_global: + assumes s: "open s" and holf: "f holomorphic_on s" + and vpg: "valid_path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" + and pas: "path_image \<gamma> \<subseteq> s" + and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0" + shows "(f has_contour_integral 0) \<gamma>" +proof - + obtain z where "z \<in> s" and znot: "z \<notin> path_image \<gamma>" + proof - + have "compact (path_image \<gamma>)" + using compact_valid_path_image vpg by blast + then have "path_image \<gamma> \<noteq> s" + by (metis (no_types) compact_open path_image_nonempty s) + with pas show ?thesis by (blast intro: that) + qed + then have pasz: "path_image \<gamma> \<subseteq> s - {z}" using pas by blast + have hol: "(\<lambda>w. (w - z) * f w) holomorphic_on s" + by (rule holomorphic_intros holf)+ + show ?thesis + using Cauchy_integral_formula_global [OF s hol \<open>z \<in> s\<close> vpg pasz loop zero] + by (auto simp: znot elim!: has_contour_integral_eq) +qed + +corollary Cauchy_theorem_global_outside: + assumes "open s" "f holomorphic_on s" "valid_path \<gamma>" "pathfinish \<gamma> = pathstart \<gamma>" "path_image \<gamma> \<subseteq> s" + "\<And>w. w \<notin> s \<Longrightarrow> w \<in> outside(path_image \<gamma>)" + shows "(f has_contour_integral 0) \<gamma>" +by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) + + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Complete_Measure.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,307 @@ +(* Title: HOL/Analysis/Complete_Measure.thy + Author: Robert Himmelmann, Johannes Hoelzl, TU Muenchen +*) + +theory Complete_Measure + imports Bochner_Integration +begin + +definition + "split_completion M A p = (if A \<in> sets M then p = (A, {}) else + \<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and> fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets M)" + +definition + "main_part M A = fst (Eps (split_completion M A))" + +definition + "null_part M A = snd (Eps (split_completion M A))" + +definition completion :: "'a measure \<Rightarrow> 'a measure" where + "completion M = measure_of (space M) { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' } + (emeasure M \<circ> main_part M)" + +lemma completion_into_space: + "{ S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' } \<subseteq> Pow (space M)" + using sets.sets_into_space by auto + +lemma space_completion[simp]: "space (completion M) = space M" + unfolding completion_def using space_measure_of[OF completion_into_space] by simp + +lemma completionI: + assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M" + shows "A \<in> { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }" + using assms by auto + +lemma completionE: + assumes "A \<in> { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }" + obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M" + using assms by auto + +lemma sigma_algebra_completion: + "sigma_algebra (space M) { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }" + (is "sigma_algebra _ ?A") + unfolding sigma_algebra_iff2 +proof (intro conjI ballI allI impI) + show "?A \<subseteq> Pow (space M)" + using sets.sets_into_space by auto +next + show "{} \<in> ?A" by auto +next + let ?C = "space M" + fix A assume "A \<in> ?A" from completionE[OF this] guess S N N' . + then show "space M - A \<in> ?A" + by (intro completionI[of _ "(?C - S) \<inter> (?C - N')" "(?C - S) \<inter> N' \<inter> (?C - N)"]) auto +next + fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> ?A" + then have "\<forall>n. \<exists>S N N'. A n = S \<union> N \<and> S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N'" + by (auto simp: image_subset_iff) + from choice[OF this] guess S .. + from choice[OF this] guess N .. + from choice[OF this] guess N' .. + then show "UNION UNIV A \<in> ?A" + using null_sets_UN[of N'] + by (intro completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"]) auto +qed + +lemma sets_completion: + "sets (completion M) = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }" + using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion] by (simp add: completion_def) + +lemma sets_completionE: + assumes "A \<in> sets (completion M)" + obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M" + using assms unfolding sets_completion by auto + +lemma sets_completionI: + assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M" + shows "A \<in> sets (completion M)" + using assms unfolding sets_completion by auto + +lemma sets_completionI_sets[intro, simp]: + "A \<in> sets M \<Longrightarrow> A \<in> sets (completion M)" + unfolding sets_completion by force + +lemma null_sets_completion: + assumes "N' \<in> null_sets M" "N \<subseteq> N'" shows "N \<in> sets (completion M)" + using assms by (intro sets_completionI[of N "{}" N N']) auto + +lemma split_completion: + assumes "A \<in> sets (completion M)" + shows "split_completion M A (main_part M A, null_part M A)" +proof cases + assume "A \<in> sets M" then show ?thesis + by (simp add: split_completion_def[abs_def] main_part_def null_part_def) +next + assume nA: "A \<notin> sets M" + show ?thesis + unfolding main_part_def null_part_def if_not_P[OF nA] + proof (rule someI2_ex) + from assms[THEN sets_completionE] guess S N N' . note A = this + let ?P = "(S, N - S)" + show "\<exists>p. split_completion M A p" + unfolding split_completion_def if_not_P[OF nA] using A + proof (intro exI conjI) + show "A = fst ?P \<union> snd ?P" using A by auto + show "snd ?P \<subseteq> N'" using A by auto + qed auto + qed auto +qed + +lemma + assumes "S \<in> sets (completion M)" + shows main_part_sets[intro, simp]: "main_part M S \<in> sets M" + and main_part_null_part_Un[simp]: "main_part M S \<union> null_part M S = S" + and main_part_null_part_Int[simp]: "main_part M S \<inter> null_part M S = {}" + using split_completion[OF assms] + by (auto simp: split_completion_def split: if_split_asm) + +lemma main_part[simp]: "S \<in> sets M \<Longrightarrow> main_part M S = S" + using split_completion[of S M] + by (auto simp: split_completion_def split: if_split_asm) + +lemma null_part: + assumes "S \<in> sets (completion M)" shows "\<exists>N. N\<in>null_sets M \<and> null_part M S \<subseteq> N" + using split_completion[OF assms] by (auto simp: split_completion_def split: if_split_asm) + +lemma null_part_sets[intro, simp]: + assumes "S \<in> sets M" shows "null_part M S \<in> sets M" "emeasure M (null_part M S) = 0" +proof - + have S: "S \<in> sets (completion M)" using assms by auto + have "S - main_part M S \<in> sets M" using assms by auto + moreover + from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S] + have "S - main_part M S = null_part M S" by auto + ultimately show sets: "null_part M S \<in> sets M" by auto + from null_part[OF S] guess N .. + with emeasure_eq_0[of N _ "null_part M S"] sets + show "emeasure M (null_part M S) = 0" by auto +qed + +lemma emeasure_main_part_UN: + fixes S :: "nat \<Rightarrow> 'a set" + assumes "range S \<subseteq> sets (completion M)" + shows "emeasure M (main_part M (\<Union>i. (S i))) = emeasure M (\<Union>i. main_part M (S i))" +proof - + have S: "\<And>i. S i \<in> sets (completion M)" using assms by auto + then have UN: "(\<Union>i. S i) \<in> sets (completion M)" by auto + have "\<forall>i. \<exists>N. N \<in> null_sets M \<and> null_part M (S i) \<subseteq> N" + using null_part[OF S] by auto + from choice[OF this] guess N .. note N = this + then have UN_N: "(\<Union>i. N i) \<in> null_sets M" by (intro null_sets_UN) auto + have "(\<Union>i. S i) \<in> sets (completion M)" using S by auto + from null_part[OF this] guess N' .. note N' = this + let ?N = "(\<Union>i. N i) \<union> N'" + have null_set: "?N \<in> null_sets M" using N' UN_N by (intro null_sets.Un) auto + have "main_part M (\<Union>i. S i) \<union> ?N = (main_part M (\<Union>i. S i) \<union> null_part M (\<Union>i. S i)) \<union> ?N" + using N' by auto + also have "\<dots> = (\<Union>i. main_part M (S i) \<union> null_part M (S i)) \<union> ?N" + unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto + also have "\<dots> = (\<Union>i. main_part M (S i)) \<union> ?N" + using N by auto + finally have *: "main_part M (\<Union>i. S i) \<union> ?N = (\<Union>i. main_part M (S i)) \<union> ?N" . + have "emeasure M (main_part M (\<Union>i. S i)) = emeasure M (main_part M (\<Union>i. S i) \<union> ?N)" + using null_set UN by (intro emeasure_Un_null_set[symmetric]) auto + also have "\<dots> = emeasure M ((\<Union>i. main_part M (S i)) \<union> ?N)" + unfolding * .. + also have "\<dots> = emeasure M (\<Union>i. main_part M (S i))" + using null_set S by (intro emeasure_Un_null_set) auto + finally show ?thesis . +qed + +lemma emeasure_completion[simp]: + assumes S: "S \<in> sets (completion M)" shows "emeasure (completion M) S = emeasure M (main_part M S)" +proof (subst emeasure_measure_of[OF completion_def completion_into_space]) + let ?\<mu> = "emeasure M \<circ> main_part M" + show "S \<in> sets (completion M)" "?\<mu> S = emeasure M (main_part M S) " using S by simp_all + show "positive (sets (completion M)) ?\<mu>" + by (simp add: positive_def) + show "countably_additive (sets (completion M)) ?\<mu>" + proof (intro countably_additiveI) + fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets (completion M)" "disjoint_family A" + have "disjoint_family (\<lambda>i. main_part M (A i))" + proof (intro disjoint_family_on_bisimulation[OF A(2)]) + fix n m assume "A n \<inter> A m = {}" + then have "(main_part M (A n) \<union> null_part M (A n)) \<inter> (main_part M (A m) \<union> null_part M (A m)) = {}" + using A by (subst (1 2) main_part_null_part_Un) auto + then show "main_part M (A n) \<inter> main_part M (A m) = {}" by auto + qed + then have "(\<Sum>n. emeasure M (main_part M (A n))) = emeasure M (\<Union>i. main_part M (A i))" + using A by (auto intro!: suminf_emeasure) + then show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (UNION UNIV A)" + by (simp add: completion_def emeasure_main_part_UN[OF A(1)]) + qed +qed + +lemma emeasure_completion_UN: + "range S \<subseteq> sets (completion M) \<Longrightarrow> + emeasure (completion M) (\<Union>i::nat. (S i)) = emeasure M (\<Union>i. main_part M (S i))" + by (subst emeasure_completion) (auto simp add: emeasure_main_part_UN) + +lemma emeasure_completion_Un: + assumes S: "S \<in> sets (completion M)" and T: "T \<in> sets (completion M)" + shows "emeasure (completion M) (S \<union> T) = emeasure M (main_part M S \<union> main_part M T)" +proof (subst emeasure_completion) + have UN: "(\<Union>i. binary (main_part M S) (main_part M T) i) = (\<Union>i. main_part M (binary S T i))" + unfolding binary_def by (auto split: if_split_asm) + show "emeasure M (main_part M (S \<union> T)) = emeasure M (main_part M S \<union> main_part M T)" + using emeasure_main_part_UN[of "binary S T" M] assms + by (simp add: range_binary_eq, simp add: Un_range_binary UN) +qed (auto intro: S T) + +lemma sets_completionI_sub: + assumes N: "N' \<in> null_sets M" "N \<subseteq> N'" + shows "N \<in> sets (completion M)" + using assms by (intro sets_completionI[of _ "{}" N N']) auto + +lemma completion_ex_simple_function: + assumes f: "simple_function (completion M) f" + shows "\<exists>f'. simple_function M f' \<and> (AE x in M. f x = f' x)" +proof - + let ?F = "\<lambda>x. f -` {x} \<inter> space M" + have F: "\<And>x. ?F x \<in> sets (completion M)" and fin: "finite (f`space M)" + using simple_functionD[OF f] simple_functionD[OF f] by simp_all + have "\<forall>x. \<exists>N. N \<in> null_sets M \<and> null_part M (?F x) \<subseteq> N" + using F null_part by auto + from choice[OF this] obtain N where + N: "\<And>x. null_part M (?F x) \<subseteq> N x" "\<And>x. N x \<in> null_sets M" by auto + let ?N = "\<Union>x\<in>f`space M. N x" + let ?f' = "\<lambda>x. if x \<in> ?N then undefined else f x" + have sets: "?N \<in> null_sets M" using N fin by (intro null_sets.finite_UN) auto + show ?thesis unfolding simple_function_def + proof (safe intro!: exI[of _ ?f']) + have "?f' ` space M \<subseteq> f`space M \<union> {undefined}" by auto + from finite_subset[OF this] simple_functionD(1)[OF f] + show "finite (?f' ` space M)" by auto + next + fix x assume "x \<in> space M" + have "?f' -` {?f' x} \<inter> space M = + (if x \<in> ?N then ?F undefined \<union> ?N + else if f x = undefined then ?F (f x) \<union> ?N + else ?F (f x) - ?N)" + using N(2) sets.sets_into_space by (auto split: if_split_asm simp: null_sets_def) + moreover { fix y have "?F y \<union> ?N \<in> sets M" + proof cases + assume y: "y \<in> f`space M" + have "?F y \<union> ?N = (main_part M (?F y) \<union> null_part M (?F y)) \<union> ?N" + using main_part_null_part_Un[OF F] by auto + also have "\<dots> = main_part M (?F y) \<union> ?N" + using y N by auto + finally show ?thesis + using F sets by auto + next + assume "y \<notin> f`space M" then have "?F y = {}" by auto + then show ?thesis using sets by auto + qed } + moreover { + have "?F (f x) - ?N = main_part M (?F (f x)) \<union> null_part M (?F (f x)) - ?N" + using main_part_null_part_Un[OF F] by auto + also have "\<dots> = main_part M (?F (f x)) - ?N" + using N \<open>x \<in> space M\<close> by auto + finally have "?F (f x) - ?N \<in> sets M" + using F sets by auto } + ultimately show "?f' -` {?f' x} \<inter> space M \<in> sets M" by auto + next + show "AE x in M. f x = ?f' x" + by (rule AE_I', rule sets) auto + qed +qed + +lemma completion_ex_borel_measurable: + fixes g :: "'a \<Rightarrow> ennreal" + assumes g: "g \<in> borel_measurable (completion M)" + shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)" +proof - + from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this + from this(1)[THEN completion_ex_simple_function] + have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x in M. f i x = f' x)" .. + from this[THEN choice] obtain f' where + sf: "\<And>i. simple_function M (f' i)" and + AE: "\<forall>i. AE x in M. f i x = f' i x" by auto + show ?thesis + proof (intro bexI) + from AE[unfolded AE_all_countable[symmetric]] + show "AE x in M. g x = (SUP i. f' i x)" (is "AE x in M. g x = ?f x") + proof (elim AE_mp, safe intro!: AE_I2) + fix x assume eq: "\<forall>i. f i x = f' i x" + moreover have "g x = (SUP i. f i x)" + unfolding f by (auto split: split_max) + ultimately show "g x = ?f x" by auto + qed + show "?f \<in> borel_measurable M" + using sf[THEN borel_measurable_simple_function] by auto + qed +qed + +lemma null_sets_completionI: "N \<in> null_sets M \<Longrightarrow> N \<in> null_sets (completion M)" + by (auto simp: null_sets_def) + +lemma AE_completion: "(AE x in M. P x) \<Longrightarrow> (AE x in completion M. P x)" + unfolding eventually_ae_filter by (auto intro: null_sets_completionI) + +lemma null_sets_completion_iff: "N \<in> sets M \<Longrightarrow> N \<in> null_sets (completion M) \<longleftrightarrow> N \<in> null_sets M" + by (auto simp: null_sets_def) + +lemma AE_completion_iff: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in completion M. P x)" + by (simp add: AE_iff_null null_sets_completion_iff) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,1328 @@ +(* Author: John Harrison, Marco Maggesi, Graziano Gentili, Gianni Ciolli, Valentina Bruno + Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014) +*) + +section \<open>Complex Analysis Basics\<close> + +theory Complex_Analysis_Basics +imports Cartesian_Euclidean_Space "~~/src/HOL/Library/Nonpos_Ints" +begin + + +subsection\<open>General lemmas\<close> + +lemma nonneg_Reals_cmod_eq_Re: "z \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> norm z = Re z" + by (simp add: complex_nonneg_Reals_iff cmod_eq_Re) + +lemma has_derivative_mult_right: + fixes c:: "'a :: real_normed_algebra" + shows "((op * c) has_derivative (op * c)) F" +by (rule has_derivative_mult_right [OF has_derivative_id]) + +lemma has_derivative_of_real[derivative_intros, simp]: + "(f has_derivative f') F \<Longrightarrow> ((\<lambda>x. of_real (f x)) has_derivative (\<lambda>x. of_real (f' x))) F" + using bounded_linear.has_derivative[OF bounded_linear_of_real] . + +lemma has_vector_derivative_real_complex: + "DERIV f (of_real a) :> f' \<Longrightarrow> ((\<lambda>x. f (of_real x)) has_vector_derivative f') (at a within s)" + using has_derivative_compose[of of_real of_real a _ f "op * f'"] + by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) + +lemma fact_cancel: + fixes c :: "'a::real_field" + shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)" + by (simp add: of_nat_mult del: of_nat_Suc times_nat.simps) + +lemma bilinear_times: + fixes c::"'a::real_algebra" shows "bilinear (\<lambda>x y::'a. x*y)" + by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI) + +lemma linear_cnj: "linear cnj" + using bounded_linear.linear[OF bounded_linear_cnj] . + +lemma tendsto_Re_upper: + assumes "~ (trivial_limit F)" + "(f \<longlongrightarrow> l) F" + "eventually (\<lambda>x. Re(f x) \<le> b) F" + shows "Re(l) \<le> b" + by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Re) + +lemma tendsto_Re_lower: + assumes "~ (trivial_limit F)" + "(f \<longlongrightarrow> l) F" + "eventually (\<lambda>x. b \<le> Re(f x)) F" + shows "b \<le> Re(l)" + by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Re) + +lemma tendsto_Im_upper: + assumes "~ (trivial_limit F)" + "(f \<longlongrightarrow> l) F" + "eventually (\<lambda>x. Im(f x) \<le> b) F" + shows "Im(l) \<le> b" + by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Im) + +lemma tendsto_Im_lower: + assumes "~ (trivial_limit F)" + "(f \<longlongrightarrow> l) F" + "eventually (\<lambda>x. b \<le> Im(f x)) F" + shows "b \<le> Im(l)" + by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Im) + +lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = op * 0" + by auto + +lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = op * 1" + by auto + +lemma continuous_mult_left: + fixes c::"'a::real_normed_algebra" + shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)" +by (rule continuous_mult [OF continuous_const]) + +lemma continuous_mult_right: + fixes c::"'a::real_normed_algebra" + shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)" +by (rule continuous_mult [OF _ continuous_const]) + +lemma continuous_on_mult_left: + fixes c::"'a::real_normed_algebra" + shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)" +by (rule continuous_on_mult [OF continuous_on_const]) + +lemma continuous_on_mult_right: + fixes c::"'a::real_normed_algebra" + shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)" +by (rule continuous_on_mult [OF _ continuous_on_const]) + +lemma uniformly_continuous_on_cmul_right [continuous_intros]: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" + shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. f x * c)" + using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . + +lemma uniformly_continuous_on_cmul_left[continuous_intros]: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" + assumes "uniformly_continuous_on s f" + shows "uniformly_continuous_on s (\<lambda>x. c * f x)" +by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right) + +lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm" + by (rule continuous_norm [OF continuous_ident]) + +lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" + by (intro continuous_on_id continuous_on_norm) + +subsection\<open>DERIV stuff\<close> + +lemma DERIV_zero_connected_constant: + fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a" + assumes "connected s" + and "open s" + and "finite k" + and "continuous_on s f" + and "\<forall>x\<in>(s - k). DERIV f x :> 0" + obtains c where "\<And>x. x \<in> s \<Longrightarrow> 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) + +lemma DERIV_zero_constant: + fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" + shows "\<lbrakk>convex s; + \<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)\<rbrakk> + \<Longrightarrow> \<exists>c. \<forall>x \<in> s. f(x) = c" + by (auto simp: has_field_derivative_def lambda_zero intro: has_derivative_zero_constant) + +lemma DERIV_zero_unique: + fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" + assumes "convex s" + and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)" + and "a \<in> s" + and "x \<in> s" + shows "f x = f a" + by (rule has_derivative_zero_unique [OF assms(1) _ assms(4,3)]) + (metis d0 has_field_derivative_imp_has_derivative lambda_zero) + +lemma DERIV_zero_connected_unique: + fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" + assumes "connected s" + and "open s" + and d0: "\<And>x. x\<in>s \<Longrightarrow> DERIV f x :> 0" + and "a \<in> s" + and "x \<in> s" + shows "f x = f a" + by (rule has_derivative_zero_unique_connected [OF assms(2,1) _ assms(5,4)]) + (metis has_field_derivative_def lambda_zero d0) + +lemma DERIV_transform_within: + assumes "(f has_field_derivative f') (at a within s)" + and "0 < d" "a \<in> s" + and "\<And>x. x\<in>s \<Longrightarrow> dist x a < d \<Longrightarrow> f x = g x" + shows "(g has_field_derivative f') (at a within s)" + using assms unfolding has_field_derivative_def + by (blast intro: has_derivative_transform_within) + +lemma DERIV_transform_within_open: + assumes "DERIV f a :> f'" + and "open s" "a \<in> s" + and "\<And>x. x\<in>s \<Longrightarrow> f x = g x" + shows "DERIV g a :> f'" + using assms unfolding has_field_derivative_def +by (metis has_derivative_transform_within_open) + +lemma DERIV_transform_at: + assumes "DERIV f a :> f'" + and "0 < d" + and "\<And>x. dist x a < d \<Longrightarrow> f x = g x" + shows "DERIV g a :> f'" + by (blast intro: assms DERIV_transform_within) + +(*generalising DERIV_isconst_all, which requires type real (using the ordering)*) +lemma DERIV_zero_UNIV_unique: + fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" + shows "(\<And>x. DERIV f x :> 0) \<Longrightarrow> f x = f a" +by (metis DERIV_zero_unique UNIV_I convex_UNIV) + +subsection \<open>Some limit theorems about real part of real series etc.\<close> + +(*MOVE? But not to Finite_Cartesian_Product*) +lemma sums_vec_nth : + assumes "f sums a" + shows "(\<lambda>x. f x $ i) sums a $ i" +using assms unfolding sums_def +by (auto dest: tendsto_vec_nth [where i=i]) + +lemma summable_vec_nth : + assumes "summable f" + shows "summable (\<lambda>x. f x $ i)" +using assms unfolding summable_def +by (blast intro: sums_vec_nth) + +subsection \<open>Complex number lemmas\<close> + +lemma + shows open_halfspace_Re_lt: "open {z. Re(z) < b}" + and open_halfspace_Re_gt: "open {z. Re(z) > b}" + and closed_halfspace_Re_ge: "closed {z. Re(z) \<ge> b}" + and closed_halfspace_Re_le: "closed {z. Re(z) \<le> b}" + and closed_halfspace_Re_eq: "closed {z. Re(z) = b}" + and open_halfspace_Im_lt: "open {z. Im(z) < b}" + and open_halfspace_Im_gt: "open {z. Im(z) > b}" + and closed_halfspace_Im_ge: "closed {z. Im(z) \<ge> b}" + and closed_halfspace_Im_le: "closed {z. Im(z) \<le> b}" + and closed_halfspace_Im_eq: "closed {z. Im(z) = b}" + by (intro open_Collect_less closed_Collect_le closed_Collect_eq continuous_on_Re + continuous_on_Im continuous_on_id continuous_on_const)+ + +lemma closed_complex_Reals: "closed (\<real> :: complex set)" +proof - + have "(\<real> :: complex set) = {z. Im z = 0}" + by (auto simp: complex_is_Real_iff) + then show ?thesis + by (metis closed_halfspace_Im_eq) +qed + +lemma closed_Real_halfspace_Re_le: "closed (\<real> \<inter> {w. Re w \<le> x})" + by (simp add: closed_Int closed_complex_Reals closed_halfspace_Re_le) + +corollary closed_nonpos_Reals_complex [simp]: "closed (\<real>\<^sub>\<le>\<^sub>0 :: complex set)" +proof - + have "\<real>\<^sub>\<le>\<^sub>0 = \<real> \<inter> {z. Re(z) \<le> 0}" + using complex_nonpos_Reals_iff complex_is_Real_iff by auto + then show ?thesis + by (metis closed_Real_halfspace_Re_le) +qed + +lemma closed_Real_halfspace_Re_ge: "closed (\<real> \<inter> {w. x \<le> Re(w)})" + using closed_halfspace_Re_ge + by (simp add: closed_Int closed_complex_Reals) + +corollary closed_nonneg_Reals_complex [simp]: "closed (\<real>\<^sub>\<ge>\<^sub>0 :: complex set)" +proof - + have "\<real>\<^sub>\<ge>\<^sub>0 = \<real> \<inter> {z. Re(z) \<ge> 0}" + using complex_nonneg_Reals_iff complex_is_Real_iff by auto + then show ?thesis + by (metis closed_Real_halfspace_Re_ge) +qed + +lemma closed_real_abs_le: "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}" +proof - + have "{w \<in> \<real>. \<bar>Re w\<bar> \<le> r} = (\<real> \<inter> {w. Re w \<le> r}) \<inter> (\<real> \<inter> {w. Re w \<ge> -r})" + by auto + then show "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}" + by (simp add: closed_Int closed_Real_halfspace_Re_ge closed_Real_halfspace_Re_le) +qed + +lemma real_lim: + fixes l::complex + assumes "(f \<longlongrightarrow> l) F" and "~(trivial_limit F)" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>" + shows "l \<in> \<real>" +proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)]) + show "eventually (\<lambda>x. f x \<in> \<real>) F" + using assms(3, 4) by (auto intro: eventually_mono) +qed + +lemma real_lim_sequentially: + fixes l::complex + shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>" +by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially) + +lemma real_series: + fixes l::complex + shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>" +unfolding sums_def +by (metis real_lim_sequentially setsum_in_Reals) + +lemma Lim_null_comparison_Re: + assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g \<longlongrightarrow> 0) F" shows "(f \<longlongrightarrow> 0) F" + by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp + +subsection\<open>Holomorphic functions\<close> + +definition field_differentiable :: "['a \<Rightarrow> 'a::real_normed_field, 'a filter] \<Rightarrow> bool" + (infixr "(field'_differentiable)" 50) + where "f field_differentiable F \<equiv> \<exists>f'. (f has_field_derivative f') F" + +lemma field_differentiable_derivI: + "f field_differentiable (at x) \<Longrightarrow> (f has_field_derivative deriv f x) (at x)" +by (simp add: field_differentiable_def DERIV_deriv_iff_has_field_derivative) + +lemma field_differentiable_imp_continuous_at: + "f field_differentiable (at x within s) \<Longrightarrow> continuous (at x within s) f" + by (metis DERIV_continuous field_differentiable_def) + +lemma field_differentiable_within_subset: + "\<lbrakk>f field_differentiable (at x within s); t \<subseteq> s\<rbrakk> + \<Longrightarrow> f field_differentiable (at x within t)" + by (metis DERIV_subset field_differentiable_def) + +lemma field_differentiable_at_within: + "\<lbrakk>f field_differentiable (at x)\<rbrakk> + \<Longrightarrow> f field_differentiable (at x within s)" + unfolding field_differentiable_def + by (metis DERIV_subset top_greatest) + +lemma field_differentiable_linear [simp,derivative_intros]: "(op * c) field_differentiable F" +proof - + show ?thesis + unfolding field_differentiable_def has_field_derivative_def mult_commute_abs + by (force intro: has_derivative_mult_right) +qed + +lemma field_differentiable_const [simp,derivative_intros]: "(\<lambda>z. c) field_differentiable F" + unfolding field_differentiable_def has_field_derivative_def + by (rule exI [where x=0]) + (metis has_derivative_const lambda_zero) + +lemma field_differentiable_ident [simp,derivative_intros]: "(\<lambda>z. z) field_differentiable F" + unfolding field_differentiable_def has_field_derivative_def + by (rule exI [where x=1]) + (simp add: lambda_one [symmetric]) + +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 \<Longrightarrow> (\<lambda>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 "(\<lambda>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]: + "op + c field_differentiable F" + by (simp add: field_differentiable_add) + +lemma field_differentiable_setsum [derivative_intros]: + "(\<And>i. i \<in> I \<Longrightarrow> (f i) field_differentiable F) \<Longrightarrow> (\<lambda>z. \<Sum>i\<in>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 "(\<lambda>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 \<noteq> 0" + shows "(\<lambda>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 "(\<lambda>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 \<noteq> 0" + shows "(\<lambda>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 "(\<lambda>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 \<Longrightarrow> + x \<in> s \<Longrightarrow> + (\<And>x'. x' \<in> s \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x') \<Longrightarrow> + f field_differentiable (at x within s) + \<Longrightarrow> 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 \<Longrightarrow> g field_differentiable at (f z) + \<Longrightarrow> (g o f) field_differentiable at z" +by (metis field_differentiable_at_within field_differentiable_compose_within) + +lemma field_differentiable_within_open: + "\<lbrakk>a \<in> s; open s\<rbrakk> \<Longrightarrow> f field_differentiable at a within s \<longleftrightarrow> + f field_differentiable at a" + unfolding field_differentiable_def + by (metis at_within_open) + +subsection\<open>Caratheodory characterization\<close> + +lemma field_differentiable_caratheodory_at: + "f field_differentiable (at z) \<longleftrightarrow> + (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> 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) \<longleftrightarrow> + (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z within s) g)" + using DERIV_caratheodory_within [of f] + by (simp add: field_differentiable_def has_field_derivative_def) + +subsection\<open>Holomorphic\<close> + +definition holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool" + (infixl "(holomorphic'_on)" 50) + where "f holomorphic_on s \<equiv> \<forall>x\<in>s. f field_differentiable (at x within s)" + +named_theorems holomorphic_intros "structural introduction rules for holomorphic_on" + +lemma holomorphic_onI [intro?]: "(\<And>x. x \<in> s \<Longrightarrow> f field_differentiable (at x within s)) \<Longrightarrow> f holomorphic_on s" + by (simp add: holomorphic_on_def) + +lemma holomorphic_onD [dest?]: "\<lbrakk>f holomorphic_on s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x within s)" + by (simp add: holomorphic_on_def) + +lemma holomorphic_on_imp_differentiable_at: + "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x)" +using at_within_open holomorphic_on_def by fastforce + +lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}" + by (simp add: holomorphic_on_def) + +lemma holomorphic_on_open: + "open s \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>f'. DERIV f x :> f')" + by (auto simp: holomorphic_on_def field_differentiable_def has_field_derivative_def at_within_open [of _ s]) + +lemma holomorphic_on_imp_continuous_on: + "f holomorphic_on s \<Longrightarrow> continuous_on s f" + by (metis field_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def) + +lemma holomorphic_on_subset [elim]: + "f holomorphic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f holomorphic_on t" + unfolding holomorphic_on_def + by (metis field_differentiable_within_subset subsetD) + +lemma holomorphic_transform: "\<lbrakk>f holomorphic_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g holomorphic_on s" + by (metis field_differentiable_transform_within linordered_field_no_ub holomorphic_on_def) + +lemma holomorphic_cong: "s = t ==> (\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> g holomorphic_on t" + by (metis holomorphic_transform) + +lemma holomorphic_on_linear [simp, holomorphic_intros]: "(op * c) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_linear) + +lemma holomorphic_on_const [simp, holomorphic_intros]: "(\<lambda>z. c) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_const) + +lemma holomorphic_on_ident [simp, holomorphic_intros]: "(\<lambda>x. x) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_ident) + +lemma holomorphic_on_id [simp, holomorphic_intros]: "id holomorphic_on s" + unfolding id_def by (rule holomorphic_on_ident) + +lemma holomorphic_on_compose: + "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) \<Longrightarrow> (g o f) holomorphic_on s" + using field_differentiable_compose_within[of f _ s g] + by (auto simp: holomorphic_on_def) + +lemma holomorphic_on_compose_gen: + "f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g o f) holomorphic_on s" + by (metis holomorphic_on_compose holomorphic_on_subset) + +lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s" + by (metis field_differentiable_minus holomorphic_on_def) + +lemma holomorphic_on_add [holomorphic_intros]: + "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_add) + +lemma holomorphic_on_diff [holomorphic_intros]: + "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_diff) + +lemma holomorphic_on_mult [holomorphic_intros]: + "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_mult) + +lemma holomorphic_on_inverse [holomorphic_intros]: + "\<lbrakk>f holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. inverse (f z)) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_inverse) + +lemma holomorphic_on_divide [holomorphic_intros]: + "\<lbrakk>f holomorphic_on s; g holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. f z / g z) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_divide) + +lemma holomorphic_on_power [holomorphic_intros]: + "f holomorphic_on s \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_power) + +lemma holomorphic_on_setsum [holomorphic_intros]: + "(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) holomorphic_on s" + unfolding holomorphic_on_def by (metis field_differentiable_setsum) + +lemma DERIV_deriv_iff_field_differentiable: + "DERIV f x :> deriv f x \<longleftrightarrow> f field_differentiable at x" + unfolding field_differentiable_def by (metis DERIV_imp_deriv) + +lemma holomorphic_derivI: + "\<lbrakk>f holomorphic_on S; open S; x \<in> S\<rbrakk> + \<Longrightarrow> (f has_field_derivative deriv f x) (at x within T)" +by (metis DERIV_deriv_iff_field_differentiable at_within_open holomorphic_on_def has_field_derivative_at_within) + +lemma complex_derivative_chain: + "f field_differentiable at x \<Longrightarrow> g field_differentiable at (f x) + \<Longrightarrow> 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 (\<lambda>w. c * w) = (\<lambda>z. c)" + by (metis DERIV_imp_deriv DERIV_cmult_Id) + +lemma deriv_ident [simp]: "deriv (\<lambda>w. w) = (\<lambda>z. 1)" + by (metis DERIV_imp_deriv DERIV_ident) + +lemma deriv_id [simp]: "deriv id = (\<lambda>z. 1)" + by (simp add: id_def) + +lemma deriv_const [simp]: "deriv (\<lambda>w. c) = (\<lambda>z. 0)" + by (metis DERIV_imp_deriv DERIV_const) + +lemma deriv_add [simp]: + "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> + \<Longrightarrow> deriv (\<lambda>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]: + "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> + \<Longrightarrow> deriv (\<lambda>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]: + "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> + \<Longrightarrow> deriv (\<lambda>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 [simp]: + "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. c * f w) z = c * deriv f z" + unfolding DERIV_deriv_iff_field_differentiable[symmetric] + by (auto intro!: DERIV_imp_deriv derivative_eq_intros) + +lemma deriv_cmult_right [simp]: + "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c" + unfolding DERIV_deriv_iff_field_differentiable[symmetric] + by (auto intro!: DERIV_imp_deriv derivative_eq_intros) + +lemma deriv_cdivide_right [simp]: + "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w / c) z = deriv f z / c" + unfolding Fields.field_class.field_divide_inverse + by (blast intro: deriv_cmult_right) + +lemma complex_derivative_transform_within_open: + "\<lbrakk>f holomorphic_on s; g holomorphic_on s; open s; z \<in> s; \<And>w. w \<in> s \<Longrightarrow> f w = g w\<rbrakk> + \<Longrightarrow> deriv f z = deriv g z" + unfolding holomorphic_on_def + by (rule DERIV_imp_deriv) + (metis DERIV_deriv_iff_field_differentiable DERIV_transform_within_open at_within_open) + +lemma deriv_compose_linear: + "f field_differentiable at (c * z) \<Longrightarrow> deriv (\<lambda>w. f (c * w)) z = c * deriv f (c * z)" +apply (rule DERIV_imp_deriv) +apply (simp add: DERIV_deriv_iff_field_differentiable [symmetric]) +apply (drule DERIV_chain' [of "times c" c z UNIV f "deriv f (c * z)", OF DERIV_cmult_Id]) +apply (simp add: algebra_simps) +done + +lemma nonzero_deriv_nonconstant: + assumes df: "DERIV f \<xi> :> df" and S: "open S" "\<xi> \<in> S" and "df \<noteq> 0" + shows "\<not> f constant_on S" +unfolding constant_on_def +by (metis \<open>df \<noteq> 0\<close> DERIV_transform_within_open [OF df S] DERIV_const DERIV_unique) + +lemma holomorphic_nonconstant: + assumes holf: "f holomorphic_on S" and "open S" "\<xi> \<in> S" "deriv f \<xi> \<noteq> 0" + shows "\<not> f constant_on S" + apply (rule nonzero_deriv_nonconstant [of f "deriv f \<xi>" \<xi> S]) + using assms + apply (auto simp: holomorphic_derivI) + done + +subsection\<open>Analyticity on a set\<close> + +definition analytic_on (infixl "(analytic'_on)" 50) + where + "f analytic_on s \<equiv> \<forall>x \<in> s. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)" + +lemma analytic_imp_holomorphic: "f analytic_on s \<Longrightarrow> f holomorphic_on s" + by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def) + (metis centre_in_ball field_differentiable_at_within) + +lemma analytic_on_open: "open s \<Longrightarrow> f analytic_on s \<longleftrightarrow> f holomorphic_on s" +apply (auto simp: analytic_imp_holomorphic) +apply (auto simp: analytic_on_def holomorphic_on_def) +by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball) + +lemma analytic_on_imp_differentiable_at: + "f analytic_on s \<Longrightarrow> x \<in> s \<Longrightarrow> f field_differentiable (at x)" + apply (auto simp: analytic_on_def holomorphic_on_def) +by (metis Topology_Euclidean_Space.open_ball centre_in_ball field_differentiable_within_open) + +lemma analytic_on_subset: "f analytic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f analytic_on t" + by (auto simp: analytic_on_def) + +lemma analytic_on_Un: "f analytic_on (s \<union> t) \<longleftrightarrow> f analytic_on s \<and> f analytic_on t" + by (auto simp: analytic_on_def) + +lemma analytic_on_Union: "f analytic_on (\<Union>s) \<longleftrightarrow> (\<forall>t \<in> s. f analytic_on t)" + by (auto simp: analytic_on_def) + +lemma analytic_on_UN: "f analytic_on (\<Union>i\<in>I. s i) \<longleftrightarrow> (\<forall>i\<in>I. f analytic_on (s i))" + by (auto simp: analytic_on_def) + +lemma analytic_on_holomorphic: + "f analytic_on s \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f holomorphic_on t)" + (is "?lhs = ?rhs") +proof - + have "?lhs \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t)" + proof safe + assume "f analytic_on s" + then show "\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t" + apply (simp add: analytic_on_def) + apply (rule exI [where x="\<Union>{u. open u \<and> f analytic_on u}"], auto) + apply (metis Topology_Euclidean_Space.open_ball analytic_on_open centre_in_ball) + by (metis analytic_on_def) + next + fix t + assume "open t" "s \<subseteq> t" "f analytic_on t" + then show "f analytic_on s" + by (metis analytic_on_subset) + qed + also have "... \<longleftrightarrow> ?rhs" + by (auto simp: analytic_on_open) + finally show ?thesis . +qed + +lemma analytic_on_linear: "(op * c) analytic_on s" + by (auto simp add: analytic_on_holomorphic holomorphic_on_linear) + +lemma analytic_on_const: "(\<lambda>z. c) analytic_on s" + by (metis analytic_on_def holomorphic_on_const zero_less_one) + +lemma analytic_on_ident: "(\<lambda>x. x) analytic_on s" + by (simp add: analytic_on_def holomorphic_on_ident gt_ex) + +lemma analytic_on_id: "id analytic_on s" + unfolding id_def by (rule analytic_on_ident) + +lemma analytic_on_compose: + assumes f: "f analytic_on s" + and g: "g analytic_on (f ` s)" + shows "(g o f) analytic_on s" +unfolding analytic_on_def +proof (intro ballI) + fix x + assume x: "x \<in> s" + then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f + by (metis analytic_on_def) + obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g + by (metis analytic_on_def g image_eqI x) + have "isCont f x" + by (metis analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at f x) + with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'" + by (auto simp: continuous_at_ball) + have "g \<circ> f holomorphic_on ball x (min d e)" + apply (rule holomorphic_on_compose) + apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) + by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball) + then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e" + by (metis d e min_less_iff_conj) +qed + +lemma analytic_on_compose_gen: + "f analytic_on s \<Longrightarrow> g analytic_on t \<Longrightarrow> (\<And>z. z \<in> s \<Longrightarrow> f z \<in> t) + \<Longrightarrow> g o f analytic_on s" +by (metis analytic_on_compose analytic_on_subset image_subset_iff) + +lemma analytic_on_neg: + "f analytic_on s \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on s" +by (metis analytic_on_holomorphic holomorphic_on_minus) + +lemma analytic_on_add: + assumes f: "f analytic_on s" + and g: "g analytic_on s" + shows "(\<lambda>z. f z + g z) analytic_on s" +unfolding analytic_on_def +proof (intro ballI) + fix z + assume z: "z \<in> s" + then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f + by (metis analytic_on_def) + obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g + by (metis analytic_on_def g z) + have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')" + apply (rule holomorphic_on_add) + apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) + by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) + then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e" + by (metis e e' min_less_iff_conj) +qed + +lemma analytic_on_diff: + assumes f: "f analytic_on s" + and g: "g analytic_on s" + shows "(\<lambda>z. f z - g z) analytic_on s" +unfolding analytic_on_def +proof (intro ballI) + fix z + assume z: "z \<in> s" + then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f + by (metis analytic_on_def) + obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g + by (metis analytic_on_def g z) + have "(\<lambda>z. f z - g z) holomorphic_on ball z (min e e')" + apply (rule holomorphic_on_diff) + apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) + by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) + then show "\<exists>e>0. (\<lambda>z. f z - g z) holomorphic_on ball z e" + by (metis e e' min_less_iff_conj) +qed + +lemma analytic_on_mult: + assumes f: "f analytic_on s" + and g: "g analytic_on s" + shows "(\<lambda>z. f z * g z) analytic_on s" +unfolding analytic_on_def +proof (intro ballI) + fix z + assume z: "z \<in> s" + then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f + by (metis analytic_on_def) + obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g + by (metis analytic_on_def g z) + have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')" + apply (rule holomorphic_on_mult) + apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) + by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) + then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e" + by (metis e e' min_less_iff_conj) +qed + +lemma analytic_on_inverse: + assumes f: "f analytic_on s" + and nz: "(\<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0)" + shows "(\<lambda>z. inverse (f z)) analytic_on s" +unfolding analytic_on_def +proof (intro ballI) + fix z + assume z: "z \<in> s" + then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f + by (metis analytic_on_def) + have "continuous_on (ball z e) f" + by (metis fh holomorphic_on_imp_continuous_on) + then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0" + by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz) + have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')" + apply (rule holomorphic_on_inverse) + apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball) + by (metis nz' mem_ball min_less_iff_conj) + then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e" + by (metis e e' min_less_iff_conj) +qed + +lemma analytic_on_divide: + assumes f: "f analytic_on s" + and g: "g analytic_on s" + and nz: "(\<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0)" + shows "(\<lambda>z. f z / g z) analytic_on s" +unfolding divide_inverse +by (metis analytic_on_inverse analytic_on_mult f g nz) + +lemma analytic_on_power: + "f analytic_on s \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on s" +by (induct n) (auto simp: analytic_on_const analytic_on_mult) + +lemma analytic_on_setsum: + "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) analytic_on s" + by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add) + +lemma deriv_left_inverse: + assumes "f holomorphic_on S" and "g holomorphic_on T" + and "open S" and "open T" + and "f ` S \<subseteq> T" + and [simp]: "\<And>z. z \<in> S \<Longrightarrow> g (f z) = z" + and "w \<in> S" + shows "deriv f w * deriv g (f w) = 1" +proof - + have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w" + by (simp add: algebra_simps) + also have "... = deriv (g o f) w" + using assms + by (metis analytic_on_imp_differentiable_at analytic_on_open complex_derivative_chain image_subset_iff) + also have "... = deriv id w" + apply (rule complex_derivative_transform_within_open [where s=S]) + apply (rule assms holomorphic_on_compose_gen holomorphic_intros)+ + apply simp + done + also have "... = 1" + by simp + finally show ?thesis . +qed + +subsection\<open>analyticity at a point\<close> + +lemma analytic_at_ball: + "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)" +by (metis analytic_on_def singleton_iff) + +lemma analytic_at: + "f analytic_on {z} \<longleftrightarrow> (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s)" +by (metis analytic_on_holomorphic empty_subsetI insert_subset) + +lemma analytic_on_analytic_at: + "f analytic_on s \<longleftrightarrow> (\<forall>z \<in> s. f analytic_on {z})" +by (metis analytic_at_ball analytic_on_def) + +lemma analytic_at_two: + "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow> + (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s \<and> g holomorphic_on s)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain s t + where st: "open s" "z \<in> s" "f holomorphic_on s" + "open t" "z \<in> t" "g holomorphic_on t" + by (auto simp: analytic_at) + show ?rhs + apply (rule_tac x="s \<inter> t" in exI) + using st + apply (auto simp: Diff_subset holomorphic_on_subset) + done +next + assume ?rhs + then show ?lhs + by (force simp add: analytic_at) +qed + +subsection\<open>Combining theorems for derivative with ``analytic at'' hypotheses\<close> + +lemma + assumes "f analytic_on {z}" "g analytic_on {z}" + shows complex_derivative_add_at: "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z" + and complex_derivative_diff_at: "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z" + and complex_derivative_mult_at: "deriv (\<lambda>w. f w * g w) z = + f z * deriv g z + deriv f z * g z" +proof - + obtain s where s: "open s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s" + using assms by (metis analytic_at_two) + show "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z" + apply (rule DERIV_imp_deriv [OF DERIV_add]) + using s + apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) + done + show "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z" + apply (rule DERIV_imp_deriv [OF DERIV_diff]) + using s + apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) + done + show "deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z" + apply (rule DERIV_imp_deriv [OF DERIV_mult']) + using s + apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) + done +qed + +lemma deriv_cmult_at: + "f analytic_on {z} \<Longrightarrow> deriv (\<lambda>w. c * f w) z = c * deriv f z" +by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) + +lemma deriv_cmult_right_at: + "f analytic_on {z} \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c" +by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) + +subsection\<open>Complex differentiation of sequences and series\<close> + +(* TODO: Could probably be simplified using Uniform_Limit *) +lemma has_complex_derivative_sequence: + fixes s :: "complex set" + assumes cvs: "convex s" + and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" + and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s \<longrightarrow> norm (f' n x - g' x) \<le> e" + and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially" + shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> + (g has_field_derivative (g' x)) (at x within s)" +proof - + from assms obtain x l where x: "x \<in> s" and tf: "((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially" + by blast + { fix e::real assume e: "e > 0" + then obtain N where N: "\<forall>n\<ge>N. \<forall>x. x \<in> s \<longrightarrow> cmod (f' n x - g' x) \<le> e" + by (metis conv) + have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h" + proof (rule exI [of _ N], clarify) + fix n y h + assume "N \<le> n" "y \<in> s" + then have "cmod (f' n y - g' y) \<le> e" + by (metis N) + then have "cmod h * cmod (f' n y - g' y) \<le> cmod h * e" + by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) + then show "cmod (f' n y * h - g' y * h) \<le> e * cmod h" + by (simp add: norm_mult [symmetric] field_simps) + qed + } note ** = this + show ?thesis + unfolding has_field_derivative_def + proof (rule has_derivative_sequence [OF cvs _ _ x]) + show "\<forall>n. \<forall>x\<in>s. (f n has_derivative (op * (f' n x))) (at x within s)" + by (metis has_field_derivative_def df) + next show "(\<lambda>n. f n x) \<longlonglongrightarrow> l" + by (rule tf) + next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h" + by (blast intro: **) + qed +qed + +lemma has_complex_derivative_series: + fixes s :: "complex set" + assumes cvs: "convex s" + and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" + and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s + \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" + and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) sums l)" + shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((g has_field_derivative g' x) (at x within s))" +proof - + from assms obtain x l where x: "x \<in> s" and sf: "((\<lambda>n. f n x) sums l)" + by blast + { fix e::real assume e: "e > 0" + then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> s + \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" + by (metis conv) + have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h" + proof (rule exI [of _ N], clarify) + fix n y h + assume "N \<le> n" "y \<in> s" + then have "cmod ((\<Sum>i<n. f' i y) - g' y) \<le> e" + by (metis N) + then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e" + by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) + then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h" + by (simp add: norm_mult [symmetric] field_simps setsum_right_distrib) + qed + } note ** = this + show ?thesis + unfolding has_field_derivative_def + proof (rule has_derivative_series [OF cvs _ _ x]) + fix n x + assume "x \<in> s" + then show "((f n) has_derivative (\<lambda>z. z * f' n x)) (at x within s)" + by (metis df has_field_derivative_def mult_commute_abs) + next show " ((\<lambda>n. f n x) sums l)" + by (rule sf) + next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h" + by (blast intro: **) + qed +qed + + +lemma field_differentiable_series: + fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex" + assumes "convex s" "open s" + assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" + assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" and x: "x \<in> s" + shows "summable (\<lambda>n. f n x)" and "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)" +proof - + from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" + unfolding uniformly_convergent_on_def by blast + from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open) + have "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (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: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x" + "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast + from g[OF x] show "summable (\<lambda>n. f n x)" by (auto simp: summable_def) + from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)" + by (simp add: has_field_derivative_def s) + have "((\<lambda>x. \<Sum>n. f n x) has_derivative op * (g' x)) (at x)" + by (rule has_derivative_transform_within_open[OF g' \<open>open s\<close> x]) + (insert g, auto simp: sums_iff) + thus "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)" unfolding differentiable_def + by (auto simp: summable_def field_differentiable_def has_field_derivative_def) +qed + +lemma field_differentiable_series': + fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex" + assumes "convex s" "open s" + assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" + assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" + shows "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x0)" + using field_differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+ + +subsection\<open>Bound theorem\<close> + +lemma field_differentiable_bound: + fixes s :: "complex set" + assumes cvs: "convex s" + and df: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z within s)" + and dn: "\<And>z. z \<in> s \<Longrightarrow> norm (f' z) \<le> B" + and "x \<in> s" "y \<in> s" + shows "norm(f x - f y) \<le> B * norm(x - y)" + apply (rule differentiable_bound [OF cvs]) + apply (rule ballI, erule df [unfolded has_field_derivative_def]) + apply (rule ballI, rule onorm_le, simp add: norm_mult mult_right_mono dn) + apply fact + apply fact + done + +subsection\<open>Inverse function theorem for complex derivatives\<close> + +lemma has_complex_derivative_inverse_basic: + fixes f :: "complex \<Rightarrow> complex" + shows "DERIV f (g y) :> f' \<Longrightarrow> + f' \<noteq> 0 \<Longrightarrow> + continuous (at y) g \<Longrightarrow> + open t \<Longrightarrow> + y \<in> t \<Longrightarrow> + (\<And>z. z \<in> t \<Longrightarrow> f (g z) = z) + \<Longrightarrow> DERIV g y :> inverse (f')" + unfolding has_field_derivative_def + apply (rule has_derivative_inverse_basic) + apply (auto simp: bounded_linear_mult_right) + done + +(*Used only once, in Multivariate/cauchy.ml. *) +lemma has_complex_derivative_inverse_strong: + fixes f :: "complex \<Rightarrow> complex" + shows "DERIV f x :> f' \<Longrightarrow> + f' \<noteq> 0 \<Longrightarrow> + open s \<Longrightarrow> + x \<in> s \<Longrightarrow> + continuous_on s f \<Longrightarrow> + (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) + \<Longrightarrow> DERIV g (f x) :> inverse (f')" + unfolding has_field_derivative_def + apply (rule has_derivative_inverse_strong [of s x f g ]) + by auto + +lemma has_complex_derivative_inverse_strong_x: + fixes f :: "complex \<Rightarrow> complex" + shows "DERIV f (g y) :> f' \<Longrightarrow> + f' \<noteq> 0 \<Longrightarrow> + open s \<Longrightarrow> + continuous_on s f \<Longrightarrow> + g y \<in> s \<Longrightarrow> f(g y) = y \<Longrightarrow> + (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) + \<Longrightarrow> DERIV g y :> inverse (f')" + unfolding has_field_derivative_def + apply (rule has_derivative_inverse_strong_x [of s g y f]) + by auto + +subsection \<open>Taylor on Complex Numbers\<close> + +lemma setsum_Suc_reindex: + fixes f :: "nat \<Rightarrow> 'a::ab_group_add" + shows "setsum f {0..n} = f 0 - f (Suc n) + setsum (\<lambda>i. f (Suc i)) {0..n}" +by (induct n) auto + +lemma complex_taylor: + assumes s: "convex s" + and f: "\<And>i x. x \<in> s \<Longrightarrow> i \<le> n \<Longrightarrow> (f i has_field_derivative f (Suc i) x) (at x within s)" + and B: "\<And>x. x \<in> s \<Longrightarrow> cmod (f (Suc n) x) \<le> B" + and w: "w \<in> s" + and z: "z \<in> s" + shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i))) + \<le> B * cmod(z - w)^(Suc n) / fact n" +proof - + have wzs: "closed_segment w z \<subseteq> s" using assms + by (metis convex_contains_segment) + { fix u + assume "u \<in> closed_segment w z" + then have "u \<in> s" + by (metis wzs subsetD) + have "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) + + f (Suc i) u * (z-u)^i / (fact i)) = + f (Suc n) u * (z-u) ^ n / (fact n)" + proof (induction n) + case 0 show ?case by simp + next + case (Suc n) + have "(\<Sum>i\<le>Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) + + f (Suc i) u * (z-u) ^ i / (fact i)) = + f (Suc n) u * (z-u) ^ n / (fact n) + + f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) - + f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))" + using Suc by simp + also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))" + proof - + have "(fact(Suc n)) * + (f(Suc n) u *(z-u) ^ n / (fact n) + + f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / (fact(Suc n)) - + f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / (fact(Suc n))) = + ((fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / (fact n) + + ((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) - + ((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))" + by (simp add: algebra_simps del: fact_Suc) + also have "... = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) + + (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - + (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" + by (simp del: fact_Suc) + also have "... = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) + + (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - + (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" + by (simp only: fact_Suc of_nat_mult ac_simps) simp + also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)" + by (simp add: algebra_simps) + finally show ?thesis + by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact_Suc) + qed + finally show ?case . + qed + then have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / (fact i))) + has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n)) + (at u within s)" + apply (intro derivative_eq_intros) + apply (blast intro: assms \<open>u \<in> s\<close>) + apply (rule refl)+ + apply (auto simp: field_simps) + done + } note sum_deriv = this + { fix u + assume u: "u \<in> closed_segment w z" + then have us: "u \<in> s" + by (metis wzs subsetD) + have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> cmod (f (Suc n) u) * cmod (u - z) ^ n" + by (metis norm_minus_commute order_refl) + also have "... \<le> cmod (f (Suc n) u) * cmod (z - w) ^ n" + by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u]) + also have "... \<le> B * cmod (z - w) ^ n" + by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us]) + finally have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> B * cmod (z - w) ^ n" . + } note cmod_bound = this + have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)) = (\<Sum>i\<le>n. (f i z / (fact i)) * 0 ^ i)" + by simp + also have "\<dots> = f 0 z / (fact 0)" + by (subst setsum_zero_power) simp + finally have "cmod (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i))) + \<le> cmod ((\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)) - + (\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)))" + by (simp add: norm_minus_commute) + also have "... \<le> B * cmod (z - w) ^ n / (fact n) * cmod (w - z)" + apply (rule field_differentiable_bound + [where f' = "\<lambda>w. f (Suc n) w * (z - w)^n / (fact n)" + and s = "closed_segment w z", OF convex_closed_segment]) + apply (auto simp: ends_in_segment DERIV_subset [OF sum_deriv wzs] + norm_divide norm_mult norm_power divide_le_cancel cmod_bound) + done + also have "... \<le> B * cmod (z - w) ^ Suc n / (fact n)" + by (simp add: algebra_simps norm_minus_commute) + finally show ?thesis . +qed + +text\<open>Something more like the traditional MVT for real components\<close> + +lemma complex_mvt_line: + assumes "\<And>u. u \<in> closed_segment w z \<Longrightarrow> (f has_field_derivative f'(u)) (at u)" + shows "\<exists>u. u \<in> closed_segment w z \<and> Re(f z) - Re(f w) = Re(f'(u) * (z - w))" +proof - + have twz: "\<And>t. (1 - t) *\<^sub>R w + t *\<^sub>R z = w + t *\<^sub>R (z - w)" + by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) + note assms[unfolded has_field_derivative_def, derivative_intros] + show ?thesis + apply (cut_tac mvt_simple + [of 0 1 "Re o f o (\<lambda>t. (1 - t) *\<^sub>R w + t *\<^sub>R z)" + "\<lambda>u. Re o (\<lambda>h. f'((1 - u) *\<^sub>R w + u *\<^sub>R z) * h) o (\<lambda>t. t *\<^sub>R (z - w))"]) + apply auto + apply (rule_tac x="(1 - x) *\<^sub>R w + x *\<^sub>R z" in exI) + apply (auto simp: closed_segment_def twz) [] + apply (intro derivative_eq_intros has_derivative_at_within, simp_all) + apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib) + apply (force simp: twz closed_segment_def) + done +qed + +lemma complex_taylor_mvt: + assumes "\<And>i x. \<lbrakk>x \<in> closed_segment w z; i \<le> n\<rbrakk> \<Longrightarrow> ((f i) has_field_derivative f (Suc i) x) (at x)" + shows "\<exists>u. u \<in> closed_segment w z \<and> + Re (f 0 z) = + Re ((\<Sum>i = 0..n. f i w * (z - w) ^ i / (fact i)) + + (f (Suc n) u * (z-u)^n / (fact n)) * (z - w))" +proof - + { fix u + assume u: "u \<in> closed_segment w z" + have "(\<Sum>i = 0..n. + (f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / + (fact i)) = + f (Suc 0) u - + (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / + (fact (Suc n)) + + (\<Sum>i = 0..n. + (f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) / + (fact (Suc i)))" + by (subst setsum_Suc_reindex) simp + also have "... = f (Suc 0) u - + (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / + (fact (Suc n)) + + (\<Sum>i = 0..n. + f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) - + f (Suc i) u * (z-u) ^ i / (fact i))" + by (simp only: diff_divide_distrib fact_cancel ac_simps) + also have "... = f (Suc 0) u - + (f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) / + (fact (Suc n)) + + f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u" + by (subst setsum_Suc_diff) auto + also have "... = f (Suc n) u * (z-u) ^ n / (fact n)" + by (simp only: algebra_simps diff_divide_distrib fact_cancel) + finally have "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i + - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = + f (Suc n) u * (z - u) ^ n / (fact n)" . + then have "((\<lambda>u. \<Sum>i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative + f (Suc n) u * (z - u) ^ n / (fact n)) (at u)" + apply (intro derivative_eq_intros)+ + apply (force intro: u assms) + apply (rule refl)+ + apply (auto simp: ac_simps) + done + } + then show ?thesis + apply (cut_tac complex_mvt_line [of w z "\<lambda>u. \<Sum>i = 0..n. f i u * (z-u) ^ i / (fact i)" + "\<lambda>u. (f (Suc n) u * (z-u)^n / (fact n))"]) + apply (auto simp add: intro: open_closed_segment) + done +qed + + +subsection \<open>Polynomal function extremal theorem, from HOL Light\<close> + +lemma polyfun_extremal_lemma: (*COMPLEX_POLYFUN_EXTREMAL_LEMMA in HOL Light*) + fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" + assumes "0 < e" + shows "\<exists>M. \<forall>z. M \<le> norm(z) \<longrightarrow> norm (\<Sum>i\<le>n. c(i) * z^i) \<le> e * norm(z) ^ (Suc n)" +proof (induct n) + case 0 with assms + show ?case + apply (rule_tac x="norm (c 0) / e" in exI) + apply (auto simp: field_simps) + done +next + case (Suc n) + obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" + using Suc assms by blast + show ?case + proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) + fix z::'a + assume z1: "M \<le> norm z" and "1 + norm (c (Suc n)) / e \<le> norm z" + then have z2: "e + norm (c (Suc n)) \<le> e * norm z" + using assms by (simp add: field_simps) + have "norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" + using M [OF z1] by simp + then have "norm (\<Sum>i\<le>n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" + by simp + then have "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" + by (blast intro: norm_triangle_le elim: ) + also have "... \<le> (e + norm (c (Suc n))) * norm z ^ Suc n" + by (simp add: norm_power norm_mult algebra_simps) + also have "... \<le> (e * norm z) * norm z ^ Suc n" + by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) + finally show "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc (Suc n)" + by simp + qed +qed + +lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) + fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" + assumes k: "c k \<noteq> 0" "1\<le>k" and kn: "k\<le>n" + shows "eventually (\<lambda>z. norm (\<Sum>i\<le>n. c(i) * z^i) \<ge> B) at_infinity" +using kn +proof (induction n) + case 0 + then show ?case + using k by simp +next + case (Suc m) + let ?even = ?case + show ?even + proof (cases "c (Suc m) = 0") + case True + then show ?even using Suc k + by auto (metis antisym_conv less_eq_Suc_le not_le) + next + case False + then obtain M where M: + "\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>m. c i * z^i) \<le> norm (c (Suc m)) / 2 * norm z ^ Suc m" + using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc + by auto + have "\<exists>b. \<forall>z. b \<le> norm z \<longrightarrow> B \<le> norm (\<Sum>i\<le>Suc m. c i * z^i)" + proof (rule exI [where x="max M (max 1 (\<bar>B\<bar> / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) + fix z::'a + assume z1: "M \<le> norm z" "1 \<le> norm z" + and "\<bar>B\<bar> * 2 / norm (c (Suc m)) \<le> norm z" + then have z2: "\<bar>B\<bar> \<le> norm (c (Suc m)) * norm z / 2" + using False by (simp add: field_simps) + have nz: "norm z \<le> norm z ^ Suc m" + by (metis \<open>1 \<le> norm z\<close> One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) + have *: "\<And>y x. norm (c (Suc m)) * norm z / 2 \<le> norm y - norm x \<Longrightarrow> B \<le> norm (x + y)" + by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) + have "norm z * norm (c (Suc m)) + 2 * norm (\<Sum>i\<le>m. c i * z^i) + \<le> norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" + using M [of z] Suc z1 by auto + also have "... \<le> 2 * (norm (c (Suc m)) * norm z ^ Suc m)" + using nz by (simp add: mult_mono del: power_Suc) + finally show "B \<le> norm ((\<Sum>i\<le>m. c i * z^i) + c (Suc m) * z ^ Suc m)" + using Suc.IH + apply (auto simp: eventually_at_infinity) + apply (rule *) + apply (simp add: field_simps norm_mult norm_power) + done + qed + then show ?even + by (simp add: eventually_at_infinity) + qed +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,3231 @@ +section \<open>Complex Transcendental Functions\<close> + +text\<open>By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\<close> + +theory Complex_Transcendental +imports + Complex_Analysis_Basics + Summation_Tests +begin + +(* TODO: Figure out what to do with Möbius transformations *) +definition "moebius a b c d = (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))" + +lemma moebius_inverse: + assumes "a * d \<noteq> b * c" "c * z + d \<noteq> 0" + shows "moebius d (-b) (-c) a (moebius a b c d z) = z" +proof - + from assms have "(-c) * moebius a b c d z + a \<noteq> 0" unfolding moebius_def + by (simp add: field_simps) + with assms show ?thesis + unfolding moebius_def by (simp add: moebius_def divide_simps) (simp add: algebra_simps)? +qed + +lemma moebius_inverse': + assumes "a * d \<noteq> b * c" "c * z - a \<noteq> 0" + shows "moebius a b c d (moebius d (-b) (-c) a z) = z" + using assms moebius_inverse[of d a "-b" "-c" z] + by (auto simp: algebra_simps) + +lemma cmod_add_real_less: + assumes "Im z \<noteq> 0" "r\<noteq>0" + shows "cmod (z + r) < cmod z + \<bar>r\<bar>" +proof (cases z) + case (Complex x y) + have "r * x / \<bar>r\<bar> < sqrt (x*x + y*y)" + apply (rule real_less_rsqrt) + using assms + apply (simp add: Complex power2_eq_square) + using not_real_square_gt_zero by blast + then show ?thesis using assms Complex + apply (auto simp: cmod_def) + apply (rule power2_less_imp_less, auto) + apply (simp add: power2_eq_square field_simps) + done +qed + +lemma cmod_diff_real_less: "Im z \<noteq> 0 \<Longrightarrow> x\<noteq>0 \<Longrightarrow> cmod (z - x) < cmod z + \<bar>x\<bar>" + using cmod_add_real_less [of z "-x"] + by simp + +lemma cmod_square_less_1_plus: + assumes "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1" + shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)" + using assms + apply (cases "Im z = 0 \<or> Re z = 0") + using abs_square_less_1 + apply (force simp add: Re_power2 Im_power2 cmod_def) + using cmod_diff_real_less [of "1 - z\<^sup>2" "1"] + apply (simp add: norm_power Im_power2) + done + +subsection\<open>The Exponential Function is Differentiable and Continuous\<close> + +lemma field_differentiable_within_exp: "exp field_differentiable (at z within s)" + using DERIV_exp field_differentiable_at_within field_differentiable_def by blast + +lemma continuous_within_exp: + fixes z::"'a::{real_normed_field,banach}" + shows "continuous (at z within s) exp" +by (simp add: continuous_at_imp_continuous_within) + +lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s" + by (simp add: field_differentiable_within_exp holomorphic_on_def) + +subsection\<open>Euler and de Moivre formulas.\<close> + +text\<open>The sine series times @{term i}\<close> +lemma sin_ii_eq: "(\<lambda>n. (\<i> * sin_coeff n) * z^n) sums (\<i> * sin z)" +proof - + have "(\<lambda>n. \<i> * sin_coeff n *\<^sub>R z^n) sums (\<i> * sin z)" + using sin_converges sums_mult by blast + then show ?thesis + by (simp add: scaleR_conv_of_real field_simps) +qed + +theorem exp_Euler: "exp(\<i> * z) = cos(z) + \<i> * sin(z)" +proof - + have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) + = (\<lambda>n. (\<i> * z) ^ n /\<^sub>R (fact n))" + proof + fix n + show "(cos_coeff n + \<i> * sin_coeff n) * z^n = (\<i> * z) ^ n /\<^sub>R (fact n)" + by (auto simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE) + qed + also have "... sums (exp (\<i> * z))" + by (rule exp_converges) + finally have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) sums (exp (\<i> * z))" . + moreover have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) sums (cos z + \<i> * sin z)" + using sums_add [OF cos_converges [of z] sin_ii_eq [of z]] + by (simp add: field_simps scaleR_conv_of_real) + ultimately show ?thesis + using sums_unique2 by blast +qed + +corollary exp_minus_Euler: "exp(-(\<i> * z)) = cos(z) - \<i> * sin(z)" + using exp_Euler [of "-z"] + by simp + +lemma sin_exp_eq: "sin z = (exp(\<i> * z) - exp(-(\<i> * z))) / (2*\<i>)" + by (simp add: exp_Euler exp_minus_Euler) + +lemma sin_exp_eq': "sin z = \<i> * (exp(-(\<i> * z)) - exp(\<i> * z)) / 2" + by (simp add: exp_Euler exp_minus_Euler) + +lemma cos_exp_eq: "cos z = (exp(\<i> * z) + exp(-(\<i> * z))) / 2" + by (simp add: exp_Euler exp_minus_Euler) + +subsection\<open>Relationships between real and complex trig functions\<close> + +lemma real_sin_eq [simp]: + fixes x::real + shows "Re(sin(of_real x)) = sin x" + by (simp add: sin_of_real) + +lemma real_cos_eq [simp]: + fixes x::real + shows "Re(cos(of_real x)) = cos x" + by (simp add: cos_of_real) + +lemma DeMoivre: "(cos z + \<i> * sin z) ^ n = cos(n * z) + \<i> * sin(n * z)" + apply (simp add: exp_Euler [symmetric]) + by (metis exp_of_nat_mult mult.left_commute) + +lemma exp_cnj: + fixes z::complex + shows "cnj (exp z) = exp (cnj z)" +proof - + have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) = (\<lambda>n. (cnj z)^n /\<^sub>R (fact n))" + by auto + also have "... sums (exp (cnj z))" + by (rule exp_converges) + finally have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) sums (exp (cnj z))" . + moreover have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) sums (cnj (exp z))" + by (metis exp_converges sums_cnj) + ultimately show ?thesis + using sums_unique2 + by blast +qed + +lemma cnj_sin: "cnj(sin z) = sin(cnj z)" + by (simp add: sin_exp_eq exp_cnj field_simps) + +lemma cnj_cos: "cnj(cos z) = cos(cnj z)" + by (simp add: cos_exp_eq exp_cnj field_simps) + +lemma field_differentiable_at_sin: "sin field_differentiable at z" + using DERIV_sin field_differentiable_def by blast + +lemma field_differentiable_within_sin: "sin field_differentiable (at z within s)" + by (simp add: field_differentiable_at_sin field_differentiable_at_within) + +lemma field_differentiable_at_cos: "cos field_differentiable at z" + using DERIV_cos field_differentiable_def by blast + +lemma field_differentiable_within_cos: "cos field_differentiable (at z within s)" + by (simp add: field_differentiable_at_cos field_differentiable_at_within) + +lemma holomorphic_on_sin: "sin holomorphic_on s" + by (simp add: field_differentiable_within_sin holomorphic_on_def) + +lemma holomorphic_on_cos: "cos holomorphic_on s" + by (simp add: field_differentiable_within_cos holomorphic_on_def) + +subsection\<open>Get a nice real/imaginary separation in Euler's formula.\<close> + +lemma Euler: "exp(z) = of_real(exp(Re z)) * + (of_real(cos(Im z)) + \<i> * of_real(sin(Im z)))" +by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real) + +lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2" + by (simp add: sin_exp_eq field_simps Re_divide Im_exp) + +lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2" + by (simp add: sin_exp_eq field_simps Im_divide Re_exp) + +lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2" + by (simp add: cos_exp_eq field_simps Re_divide Re_exp) + +lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2" + by (simp add: cos_exp_eq field_simps Im_divide Im_exp) + +lemma Re_sin_pos: "0 < Re z \<Longrightarrow> Re z < pi \<Longrightarrow> Re (sin z) > 0" + by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero) + +lemma Im_sin_nonneg: "Re z = 0 \<Longrightarrow> 0 \<le> Im z \<Longrightarrow> 0 \<le> Im (sin z)" + by (simp add: Re_sin Im_sin algebra_simps) + +lemma Im_sin_nonneg2: "Re z = pi \<Longrightarrow> Im z \<le> 0 \<Longrightarrow> 0 \<le> Im (sin z)" + by (simp add: Re_sin Im_sin algebra_simps) + +subsection\<open>More on the Polar Representation of Complex Numbers\<close> + +lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)" + by (simp add: exp_add exp_Euler exp_of_real sin_of_real cos_of_real) + +lemma exp_eq_1: "exp z = 1 \<longleftrightarrow> Re(z) = 0 \<and> (\<exists>n::int. Im(z) = of_int (2 * n) * pi)" +apply auto +apply (metis exp_eq_one_iff norm_exp_eq_Re norm_one) +apply (metis Re_exp cos_one_2pi_int mult.commute mult.left_neutral norm_exp_eq_Re norm_one one_complex.simps(1)) +by (metis Im_exp Re_exp complex_Re_Im_cancel_iff cos_one_2pi_int sin_double Re_complex_of_real complex_Re_numeral exp_zero mult.assoc mult.left_commute mult_eq_0_iff mult_numeral_1 numeral_One of_real_0 sin_zero_iff_int2) + +lemma exp_eq: "exp w = exp z \<longleftrightarrow> (\<exists>n::int. w = z + (of_int (2 * n) * pi) * \<i>)" + (is "?lhs = ?rhs") +proof - + have "exp w = exp z \<longleftrightarrow> exp (w-z) = 1" + by (simp add: exp_diff) + also have "... \<longleftrightarrow> (Re w = Re z \<and> (\<exists>n::int. Im w - Im z = of_int (2 * n) * pi))" + by (simp add: exp_eq_1) + also have "... \<longleftrightarrow> ?rhs" + by (auto simp: algebra_simps intro!: complex_eqI) + finally show ?thesis . +qed + +lemma exp_complex_eqI: "\<bar>Im w - Im z\<bar> < 2*pi \<Longrightarrow> exp w = exp z \<Longrightarrow> w = z" + by (auto simp: exp_eq abs_mult) + +lemma exp_integer_2pi: + assumes "n \<in> \<int>" + shows "exp((2 * n * pi) * \<i>) = 1" +proof - + have "exp((2 * n * pi) * \<i>) = exp 0" + using assms + by (simp only: Ints_def exp_eq) auto + also have "... = 1" + by simp + finally show ?thesis . +qed + +lemma sin_cos_eq_iff: "sin y = sin x \<and> cos y = cos x \<longleftrightarrow> (\<exists>n::int. y = x + 2 * n * pi)" +proof - + { assume "sin y = sin x" "cos y = cos x" + then have "cos (y-x) = 1" + using cos_add [of y "-x"] by simp + then have "\<exists>n::int. y-x = n * 2 * pi" + using cos_one_2pi_int by blast } + then show ?thesis + apply (auto simp: sin_add cos_add) + apply (metis add.commute diff_add_cancel mult.commute) + done +qed + +lemma exp_i_ne_1: + assumes "0 < x" "x < 2*pi" + shows "exp(\<i> * of_real x) \<noteq> 1" +proof + assume "exp (\<i> * of_real x) = 1" + then have "exp (\<i> * of_real x) = exp 0" + by simp + then obtain n where "\<i> * of_real x = (of_int (2 * n) * pi) * \<i>" + by (simp only: Ints_def exp_eq) auto + then have "of_real x = (of_int (2 * n) * pi)" + by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real) + then have "x = (of_int (2 * n) * pi)" + by simp + then show False using assms + by (cases n) (auto simp: zero_less_mult_iff mult_less_0_iff) +qed + +lemma sin_eq_0: + fixes z::complex + shows "sin z = 0 \<longleftrightarrow> (\<exists>n::int. z = of_real(n * pi))" + by (simp add: sin_exp_eq exp_eq of_real_numeral) + +lemma cos_eq_0: + fixes z::complex + shows "cos z = 0 \<longleftrightarrow> (\<exists>n::int. z = of_real(n * pi) + of_real pi/2)" + using sin_eq_0 [of "z - of_real pi/2"] + by (simp add: sin_diff algebra_simps) + +lemma cos_eq_1: + fixes z::complex + shows "cos z = 1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi))" +proof - + have "cos z = cos (2*(z/2))" + by simp + also have "... = 1 - 2 * sin (z/2) ^ 2" + by (simp only: cos_double_sin) + finally have [simp]: "cos z = 1 \<longleftrightarrow> sin (z/2) = 0" + by simp + show ?thesis + by (auto simp: sin_eq_0 of_real_numeral) +qed + +lemma csin_eq_1: + fixes z::complex + shows "sin z = 1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi) + of_real pi/2)" + using cos_eq_1 [of "z - of_real pi/2"] + by (simp add: cos_diff algebra_simps) + +lemma csin_eq_minus1: + fixes z::complex + shows "sin z = -1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi) + 3/2*pi)" + (is "_ = ?rhs") +proof - + have "sin z = -1 \<longleftrightarrow> sin (-z) = 1" + by (simp add: equation_minus_iff) + also have "... \<longleftrightarrow> (\<exists>n::int. -z = of_real(2 * n * pi) + of_real pi/2)" + by (simp only: csin_eq_1) + also have "... \<longleftrightarrow> (\<exists>n::int. z = - of_real(2 * n * pi) - of_real pi/2)" + apply (rule iff_exI) + by (metis (no_types) is_num_normalize(8) minus_minus of_real_def real_vector.scale_minus_left uminus_add_conv_diff) + also have "... = ?rhs" + apply (auto simp: of_real_numeral) + apply (rule_tac [2] x="-(x+1)" in exI) + apply (rule_tac x="-(x+1)" in exI) + apply (simp_all add: algebra_simps) + done + finally show ?thesis . +qed + +lemma ccos_eq_minus1: + fixes z::complex + shows "cos z = -1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi) + pi)" + using csin_eq_1 [of "z - of_real pi/2"] + apply (simp add: sin_diff) + apply (simp add: algebra_simps of_real_numeral equation_minus_iff) + done + +lemma sin_eq_1: "sin x = 1 \<longleftrightarrow> (\<exists>n::int. x = (2 * n + 1 / 2) * pi)" + (is "_ = ?rhs") +proof - + have "sin x = 1 \<longleftrightarrow> sin (complex_of_real x) = 1" + by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real) + also have "... \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)" + by (simp only: csin_eq_1) + also have "... \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + of_real pi/2)" + apply (rule iff_exI) + apply (auto simp: algebra_simps of_real_numeral) + apply (rule injD [OF inj_of_real [where 'a = complex]]) + apply (auto simp: of_real_numeral) + done + also have "... = ?rhs" + by (auto simp: algebra_simps) + finally show ?thesis . +qed + +lemma sin_eq_minus1: "sin x = -1 \<longleftrightarrow> (\<exists>n::int. x = (2*n + 3/2) * pi)" (is "_ = ?rhs") +proof - + have "sin x = -1 \<longleftrightarrow> sin (complex_of_real x) = -1" + by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real) + also have "... \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)" + by (simp only: csin_eq_minus1) + also have "... \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + 3/2*pi)" + apply (rule iff_exI) + apply (auto simp: algebra_simps) + apply (rule injD [OF inj_of_real [where 'a = complex]], auto) + done + also have "... = ?rhs" + by (auto simp: algebra_simps) + finally show ?thesis . +qed + +lemma cos_eq_minus1: "cos x = -1 \<longleftrightarrow> (\<exists>n::int. x = (2*n + 1) * pi)" + (is "_ = ?rhs") +proof - + have "cos x = -1 \<longleftrightarrow> cos (complex_of_real x) = -1" + by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real) + also have "... \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + pi)" + by (simp only: ccos_eq_minus1) + also have "... \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + pi)" + apply (rule iff_exI) + apply (auto simp: algebra_simps) + apply (rule injD [OF inj_of_real [where 'a = complex]], auto) + done + also have "... = ?rhs" + by (auto simp: algebra_simps) + finally show ?thesis . +qed + +lemma dist_exp_ii_1: "norm(exp(\<i> * of_real t) - 1) = 2 * \<bar>sin(t / 2)\<bar>" + apply (simp add: exp_Euler cmod_def power2_diff sin_of_real cos_of_real algebra_simps) + using cos_double_sin [of "t/2"] + apply (simp add: real_sqrt_mult) + done + +lemma sinh_complex: + fixes z :: complex + shows "(exp z - inverse (exp z)) / 2 = -\<i> * sin(\<i> * z)" + by (simp add: sin_exp_eq divide_simps exp_minus of_real_numeral) + +lemma sin_ii_times: + fixes z :: complex + shows "sin(\<i> * z) = \<i> * ((exp z - inverse (exp z)) / 2)" + using sinh_complex by auto + +lemma sinh_real: + fixes x :: real + shows "of_real((exp x - inverse (exp x)) / 2) = -\<i> * sin(\<i> * of_real x)" + by (simp add: exp_of_real sin_ii_times of_real_numeral) + +lemma cosh_complex: + fixes z :: complex + shows "(exp z + inverse (exp z)) / 2 = cos(\<i> * z)" + by (simp add: cos_exp_eq divide_simps exp_minus of_real_numeral exp_of_real) + +lemma cosh_real: + fixes x :: real + shows "of_real((exp x + inverse (exp x)) / 2) = cos(\<i> * of_real x)" + by (simp add: cos_exp_eq divide_simps exp_minus of_real_numeral exp_of_real) + +lemmas cos_ii_times = cosh_complex [symmetric] + +lemma norm_cos_squared: + "norm(cos z) ^ 2 = cos(Re z) ^ 2 + (exp(Im z) - inverse(exp(Im z))) ^ 2 / 4" + apply (cases z) + apply (simp add: cos_add cmod_power2 cos_of_real sin_of_real) + apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) + apply (simp only: left_diff_distrib [symmetric] power_mult_distrib) + apply (simp add: sin_squared_eq) + apply (simp add: power2_eq_square algebra_simps divide_simps) + done + +lemma norm_sin_squared: + "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4" + apply (cases z) + apply (simp add: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double) + apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) + apply (simp only: left_diff_distrib [symmetric] power_mult_distrib) + apply (simp add: cos_squared_eq) + apply (simp add: power2_eq_square algebra_simps divide_simps) + done + +lemma exp_uminus_Im: "exp (- Im z) \<le> exp (cmod z)" + using abs_Im_le_cmod linear order_trans by fastforce + +lemma norm_cos_le: + fixes z::complex + shows "norm(cos z) \<le> exp(norm z)" +proof - + have "Im z \<le> cmod z" + using abs_Im_le_cmod abs_le_D1 by auto + with exp_uminus_Im show ?thesis + apply (simp add: cos_exp_eq norm_divide) + apply (rule order_trans [OF norm_triangle_ineq], simp) + apply (metis add_mono exp_le_cancel_iff mult_2_right) + done +qed + +lemma norm_cos_plus1_le: + fixes z::complex + shows "norm(1 + cos z) \<le> 2 * exp(norm z)" +proof - + have mono: "\<And>u w z::real. (1 \<le> w | 1 \<le> z) \<Longrightarrow> (w \<le> u & z \<le> u) \<Longrightarrow> 2 + w + z \<le> 4 * u" + by arith + have *: "Im z \<le> cmod z" + using abs_Im_le_cmod abs_le_D1 by auto + have triangle3: "\<And>x y z. norm(x + y + z) \<le> norm(x) + norm(y) + norm(z)" + by (simp add: norm_add_rule_thm) + have "norm(1 + cos z) = cmod (1 + (exp (\<i> * z) + exp (- (\<i> * z))) / 2)" + by (simp add: cos_exp_eq) + also have "... = cmod ((2 + exp (\<i> * z) + exp (- (\<i> * z))) / 2)" + by (simp add: field_simps) + also have "... = cmod (2 + exp (\<i> * z) + exp (- (\<i> * z))) / 2" + by (simp add: norm_divide) + finally show ?thesis + apply (rule ssubst, simp) + apply (rule order_trans [OF triangle3], simp) + using exp_uminus_Im * + apply (auto intro: mono) + done +qed + +subsection\<open>Taylor series for complex exponential, sine and cosine.\<close> + +declare power_Suc [simp del] + +lemma Taylor_exp: + "norm(exp z - (\<Sum>k\<le>n. z ^ k / (fact k))) \<le> exp\<bar>Re z\<bar> * (norm z) ^ (Suc n) / (fact n)" +proof (rule complex_taylor [of _ n "\<lambda>k. exp" "exp\<bar>Re z\<bar>" 0 z, simplified]) + show "convex (closed_segment 0 z)" + by (rule convex_closed_segment [of 0 z]) +next + fix k x + assume "x \<in> closed_segment 0 z" "k \<le> n" + show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" + using DERIV_exp DERIV_subset by blast +next + fix x + assume "x \<in> closed_segment 0 z" + then show "Re x \<le> \<bar>Re z\<bar>" + apply (auto simp: closed_segment_def scaleR_conv_of_real) + by (meson abs_ge_self abs_ge_zero linear mult_left_le_one_le mult_nonneg_nonpos order_trans) +next + show "0 \<in> closed_segment 0 z" + by (auto simp: closed_segment_def) +next + show "z \<in> closed_segment 0 z" + apply (simp add: closed_segment_def scaleR_conv_of_real) + using of_real_1 zero_le_one by blast +qed + +lemma + assumes "0 \<le> u" "u \<le> 1" + shows cmod_sin_le_exp: "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" + and cmod_cos_le_exp: "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" +proof - + have mono: "\<And>u w z::real. w \<le> u \<Longrightarrow> z \<le> u \<Longrightarrow> w + z \<le> u*2" + by arith + show "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" using assms + apply (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide) + apply (rule order_trans [OF norm_triangle_ineq4]) + apply (rule mono) + apply (auto simp: abs_if mult_left_le_one_le) + apply (meson mult_nonneg_nonneg neg_le_0_iff_le not_le order_trans) + apply (meson less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans) + done + show "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" using assms + apply (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide) + apply (rule order_trans [OF norm_triangle_ineq]) + apply (rule mono) + apply (auto simp: abs_if mult_left_le_one_le) + apply (meson mult_nonneg_nonneg neg_le_0_iff_le not_le order_trans) + apply (meson less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans) + done +qed + +lemma Taylor_sin: + "norm(sin z - (\<Sum>k\<le>n. complex_of_real (sin_coeff k) * z ^ k)) + \<le> exp\<bar>Im z\<bar> * (norm z) ^ (Suc n) / (fact n)" +proof - + have mono: "\<And>u w z::real. w \<le> u \<Longrightarrow> z \<le> u \<Longrightarrow> w + z \<le> u*2" + by arith + have *: "cmod (sin z - + (\<Sum>i\<le>n. (-1) ^ (i div 2) * (if even i then sin 0 else cos 0) * z ^ i / (fact i))) + \<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact n)" + proof (rule complex_taylor [of "closed_segment 0 z" n + "\<lambda>k x. (-1)^(k div 2) * (if even k then sin x else cos x)" + "exp\<bar>Im z\<bar>" 0 z, simplified]) + fix k x + show "((\<lambda>x. (- 1) ^ (k div 2) * (if even k then sin x else cos x)) has_field_derivative + (- 1) ^ (Suc k div 2) * (if odd k then sin x else cos x)) + (at x within closed_segment 0 z)" + apply (auto simp: power_Suc) + apply (intro derivative_eq_intros | simp)+ + done + next + fix x + assume "x \<in> closed_segment 0 z" + then show "cmod ((- 1) ^ (Suc n div 2) * (if odd n then sin x else cos x)) \<le> exp \<bar>Im z\<bar>" + by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) + qed + have **: "\<And>k. complex_of_real (sin_coeff k) * z ^ k + = (-1)^(k div 2) * (if even k then sin 0 else cos 0) * z^k / of_nat (fact k)" + by (auto simp: sin_coeff_def elim!: oddE) + show ?thesis + apply (rule order_trans [OF _ *]) + apply (simp add: **) + done +qed + +lemma Taylor_cos: + "norm(cos z - (\<Sum>k\<le>n. complex_of_real (cos_coeff k) * z ^ k)) + \<le> exp\<bar>Im z\<bar> * (norm z) ^ Suc n / (fact n)" +proof - + have mono: "\<And>u w z::real. w \<le> u \<Longrightarrow> z \<le> u \<Longrightarrow> w + z \<le> u*2" + by arith + have *: "cmod (cos z - + (\<Sum>i\<le>n. (-1) ^ (Suc i div 2) * (if even i then cos 0 else sin 0) * z ^ i / (fact i))) + \<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact n)" + proof (rule complex_taylor [of "closed_segment 0 z" n "\<lambda>k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" "exp\<bar>Im z\<bar>" 0 z, +simplified]) + fix k x + assume "x \<in> closed_segment 0 z" "k \<le> n" + show "((\<lambda>x. (- 1) ^ (Suc k div 2) * (if even k then cos x else sin x)) has_field_derivative + (- 1) ^ Suc (k div 2) * (if odd k then cos x else sin x)) + (at x within closed_segment 0 z)" + apply (auto simp: power_Suc) + apply (intro derivative_eq_intros | simp)+ + done + next + fix x + assume "x \<in> closed_segment 0 z" + then show "cmod ((- 1) ^ Suc (n div 2) * (if odd n then cos x else sin x)) \<le> exp \<bar>Im z\<bar>" + by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) + qed + have **: "\<And>k. complex_of_real (cos_coeff k) * z ^ k + = (-1)^(Suc k div 2) * (if even k then cos 0 else sin 0) * z^k / of_nat (fact k)" + by (auto simp: cos_coeff_def elim!: evenE) + show ?thesis + apply (rule order_trans [OF _ *]) + apply (simp add: **) + done +qed + +declare power_Suc [simp] + +text\<open>32-bit Approximation to e\<close> +lemma e_approx_32: "\<bar>exp(1) - 5837465777 / 2147483648\<bar> \<le> (inverse(2 ^ 32)::real)" + using Taylor_exp [of 1 14] exp_le + apply (simp add: setsum_left_distrib in_Reals_norm Re_exp atMost_nat_numeral fact_numeral) + apply (simp only: pos_le_divide_eq [symmetric], linarith) + done + +lemma e_less_3: "exp 1 < (3::real)" + using e_approx_32 + by (simp add: abs_if split: if_split_asm) + +lemma ln3_gt_1: "ln 3 > (1::real)" + by (metis e_less_3 exp_less_cancel_iff exp_ln_iff less_trans ln_exp) + + +subsection\<open>The argument of a complex number\<close> + +definition Arg :: "complex \<Rightarrow> real" where + "Arg z \<equiv> if z = 0 then 0 + else THE t. 0 \<le> t \<and> t < 2*pi \<and> + z = of_real(norm z) * exp(\<i> * of_real t)" + +lemma Arg_0 [simp]: "Arg(0) = 0" + by (simp add: Arg_def) + +lemma Arg_unique_lemma: + assumes z: "z = of_real(norm z) * exp(\<i> * of_real t)" + and z': "z = of_real(norm z) * exp(\<i> * of_real t')" + and t: "0 \<le> t" "t < 2*pi" + and t': "0 \<le> t'" "t' < 2*pi" + and nz: "z \<noteq> 0" + shows "t' = t" +proof - + have [dest]: "\<And>x y z::real. x\<ge>0 \<Longrightarrow> x+y < z \<Longrightarrow> y<z" + by arith + have "of_real (cmod z) * exp (\<i> * of_real t') = of_real (cmod z) * exp (\<i> * of_real t)" + by (metis z z') + then have "exp (\<i> * of_real t') = exp (\<i> * of_real t)" + by (metis nz mult_left_cancel mult_zero_left z) + then have "sin t' = sin t \<and> cos t' = cos t" + apply (simp add: exp_Euler sin_of_real cos_of_real) + by (metis Complex_eq complex.sel) + then obtain n::int where n: "t' = t + 2 * n * pi" + by (auto simp: sin_cos_eq_iff) + then have "n=0" + apply (rule_tac z=n in int_cases) + using t t' + apply (auto simp: mult_less_0_iff algebra_simps) + done + then show "t' = t" + by (simp add: n) +qed + +lemma Arg: "0 \<le> Arg z & Arg z < 2*pi & z = of_real(norm z) * exp(\<i> * of_real(Arg z))" +proof (cases "z=0") + case True then show ?thesis + by (simp add: Arg_def) +next + case False + obtain t where t: "0 \<le> t" "t < 2*pi" + and ReIm: "Re z / cmod z = cos t" "Im z / cmod z = sin t" + using sincos_total_2pi [OF complex_unit_circle [OF False]] + by blast + have z: "z = of_real(norm z) * exp(\<i> * of_real t)" + apply (rule complex_eqI) + using t False ReIm + apply (auto simp: exp_Euler sin_of_real cos_of_real divide_simps) + done + show ?thesis + apply (simp add: Arg_def False) + apply (rule theI [where a=t]) + using t z False + apply (auto intro: Arg_unique_lemma) + done +qed + +corollary + shows Arg_ge_0: "0 \<le> Arg z" + and Arg_lt_2pi: "Arg z < 2*pi" + and Arg_eq: "z = of_real(norm z) * exp(\<i> * of_real(Arg z))" + using Arg by auto + +lemma complex_norm_eq_1_exp: "norm z = 1 \<longleftrightarrow> (\<exists>t. z = exp(\<i> * of_real t))" + using Arg [of z] by auto + +lemma Arg_unique: "\<lbrakk>of_real r * exp(\<i> * of_real a) = z; 0 < r; 0 \<le> a; a < 2*pi\<rbrakk> \<Longrightarrow> Arg z = a" + apply (rule Arg_unique_lemma [OF _ Arg_eq]) + using Arg [of z] + apply (auto simp: norm_mult) + done + +lemma Arg_minus: "z \<noteq> 0 \<Longrightarrow> Arg (-z) = (if Arg z < pi then Arg z + pi else Arg z - pi)" + apply (rule Arg_unique [of "norm z"]) + apply (rule complex_eqI) + using Arg_ge_0 [of z] Arg_eq [of z] Arg_lt_2pi [of z] Arg_eq [of z] + apply auto + apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric]) + apply (metis Re_rcis Im_rcis rcis_def)+ + done + +lemma Arg_times_of_real [simp]: "0 < r \<Longrightarrow> Arg (of_real r * z) = Arg z" + apply (cases "z=0", simp) + apply (rule Arg_unique [of "r * norm z"]) + using Arg + apply auto + done + +lemma Arg_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg (z * of_real r) = Arg z" + by (metis Arg_times_of_real mult.commute) + +lemma Arg_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg (z / of_real r) = Arg z" + by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) + +lemma Arg_le_pi: "Arg z \<le> pi \<longleftrightarrow> 0 \<le> Im z" +proof (cases "z=0") + case True then show ?thesis + by simp +next + case False + have "0 \<le> Im z \<longleftrightarrow> 0 \<le> Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))" + by (metis Arg_eq) + also have "... = (0 \<le> Im (exp (\<i> * complex_of_real (Arg z))))" + using False + by (simp add: zero_le_mult_iff) + also have "... \<longleftrightarrow> Arg z \<le> pi" + by (simp add: Im_exp) (metis Arg_ge_0 Arg_lt_2pi sin_lt_zero sin_ge_zero not_le) + finally show ?thesis + by blast +qed + +lemma Arg_lt_pi: "0 < Arg z \<and> Arg z < pi \<longleftrightarrow> 0 < Im z" +proof (cases "z=0") + case True then show ?thesis + by simp +next + case False + have "0 < Im z \<longleftrightarrow> 0 < Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))" + by (metis Arg_eq) + also have "... = (0 < Im (exp (\<i> * complex_of_real (Arg z))))" + using False + by (simp add: zero_less_mult_iff) + also have "... \<longleftrightarrow> 0 < Arg z \<and> Arg z < pi" + using Arg_ge_0 Arg_lt_2pi sin_le_zero sin_gt_zero + apply (auto simp: Im_exp) + using le_less apply fastforce + using not_le by blast + finally show ?thesis + by blast +qed + +lemma Arg_eq_0: "Arg z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z" +proof (cases "z=0") + case True then show ?thesis + by simp +next + case False + have "z \<in> \<real> \<and> 0 \<le> Re z \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))" + by (metis Arg_eq) + also have "... \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg z)))" + using False + by (simp add: zero_le_mult_iff) + also have "... \<longleftrightarrow> Arg z = 0" + apply (auto simp: Re_exp) + apply (metis Arg_lt_pi Arg_ge_0 Arg_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl) + using Arg_eq [of z] + apply (auto simp: Reals_def) + done + finally show ?thesis + by blast +qed + +corollary Arg_gt_0: + assumes "z \<in> \<real> \<Longrightarrow> Re z < 0" + shows "Arg z > 0" + using Arg_eq_0 Arg_ge_0 assms dual_order.strict_iff_order by fastforce + +lemma Arg_of_real: "Arg(of_real x) = 0 \<longleftrightarrow> 0 \<le> x" + by (simp add: Arg_eq_0) + +lemma Arg_eq_pi: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0" + apply (cases "z=0", simp) + using Arg_eq_0 [of "-z"] + apply (auto simp: complex_is_Real_iff Arg_minus) + apply (simp add: complex_Re_Im_cancel_iff) + apply (metis Arg_minus pi_gt_zero add.left_neutral minus_minus minus_zero) + done + +lemma Arg_eq_0_pi: "Arg z = 0 \<or> Arg z = pi \<longleftrightarrow> z \<in> \<real>" + using Arg_eq_0 Arg_eq_pi not_le by auto + +lemma Arg_inverse: "Arg(inverse z) = (if z \<in> \<real> \<and> 0 \<le> Re z then Arg z else 2*pi - Arg z)" + apply (cases "z=0", simp) + apply (rule Arg_unique [of "inverse (norm z)"]) + using Arg_ge_0 [of z] Arg_lt_2pi [of z] Arg_eq [of z] Arg_eq_0 [of z] exp_two_pi_i + apply (auto simp: of_real_numeral algebra_simps exp_diff divide_simps) + done + +lemma Arg_eq_iff: + assumes "w \<noteq> 0" "z \<noteq> 0" + shows "Arg w = Arg z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)" + using assms Arg_eq [of z] Arg_eq [of w] + apply auto + apply (rule_tac x="norm w / norm z" in exI) + apply (simp add: divide_simps) + by (metis mult.commute mult.left_commute) + +lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \<longleftrightarrow> Arg z = 0" + using complex_is_Real_iff + apply (simp add: Arg_eq_0) + apply (auto simp: divide_simps not_sum_power2_lt_zero) + done + +lemma Arg_divide: + assumes "w \<noteq> 0" "z \<noteq> 0" "Arg w \<le> Arg z" + shows "Arg(z / w) = Arg z - Arg w" + apply (rule Arg_unique [of "norm(z / w)"]) + using assms Arg_eq [of z] Arg_eq [of w] Arg_ge_0 [of w] Arg_lt_2pi [of z] + apply (auto simp: exp_diff norm_divide algebra_simps divide_simps) + done + +lemma Arg_le_div_sum: + assumes "w \<noteq> 0" "z \<noteq> 0" "Arg w \<le> Arg z" + shows "Arg z = Arg w + Arg(z / w)" + by (simp add: Arg_divide assms) + +lemma Arg_le_div_sum_eq: + assumes "w \<noteq> 0" "z \<noteq> 0" + shows "Arg w \<le> Arg z \<longleftrightarrow> Arg z = Arg w + Arg(z / w)" + using assms + by (auto simp: Arg_ge_0 intro: Arg_le_div_sum) + +lemma Arg_diff: + assumes "w \<noteq> 0" "z \<noteq> 0" + shows "Arg w - Arg z = (if Arg z \<le> Arg w then Arg(w / z) else Arg(w/z) - 2*pi)" + using assms + apply (auto simp: Arg_ge_0 Arg_divide not_le) + using Arg_divide [of w z] Arg_inverse [of "w/z"] + apply auto + by (metis Arg_eq_0 less_irrefl minus_diff_eq right_minus_eq) + +lemma Arg_add: + assumes "w \<noteq> 0" "z \<noteq> 0" + shows "Arg w + Arg z = (if Arg w + Arg z < 2*pi then Arg(w * z) else Arg(w * z) + 2*pi)" + using assms + using Arg_diff [of "w*z" z] Arg_le_div_sum_eq [of z "w*z"] + apply (auto simp: Arg_ge_0 Arg_divide not_le) + apply (metis Arg_lt_2pi add.commute) + apply (metis (no_types) Arg add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less) + done + +lemma Arg_times: + assumes "w \<noteq> 0" "z \<noteq> 0" + shows "Arg (w * z) = (if Arg w + Arg z < 2*pi then Arg w + Arg z + else (Arg w + Arg z) - 2*pi)" + using Arg_add [OF assms] + by auto + +lemma Arg_cnj: "Arg(cnj z) = (if z \<in> \<real> \<and> 0 \<le> Re z then Arg z else 2*pi - Arg z)" + apply (cases "z=0", simp) + apply (rule trans [of _ "Arg(inverse z)"]) + apply (simp add: Arg_eq_iff divide_simps complex_norm_square [symmetric] mult.commute) + apply (metis norm_eq_zero of_real_power zero_less_power2) + apply (auto simp: of_real_numeral Arg_inverse) + done + +lemma Arg_real: "z \<in> \<real> \<Longrightarrow> Arg z = (if 0 \<le> Re z then 0 else pi)" + using Arg_eq_0 Arg_eq_0_pi + by auto + +lemma Arg_exp: "0 \<le> Im z \<Longrightarrow> Im z < 2*pi \<Longrightarrow> Arg(exp z) = Im z" + by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) + +lemma complex_split_polar: + obtains r a::real where "z = complex_of_real r * (cos a + \<i> * sin a)" "0 \<le> r" "0 \<le> a" "a < 2*pi" + using Arg cis.ctr cis_conv_exp by fastforce + +lemma Re_Im_le_cmod: "Im w * sin \<phi> + Re w * cos \<phi> \<le> cmod w" +proof (cases w rule: complex_split_polar) + case (1 r a) with sin_cos_le1 [of a \<phi>] show ?thesis + apply (simp add: norm_mult cmod_unit_one) + by (metis (no_types, hide_lams) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le) +qed + +subsection\<open>Analytic properties of tangent function\<close> + +lemma cnj_tan: "cnj(tan z) = tan(cnj z)" + by (simp add: cnj_cos cnj_sin tan_def) + +lemma field_differentiable_at_tan: "~(cos z = 0) \<Longrightarrow> tan field_differentiable at z" + unfolding field_differentiable_def + using DERIV_tan by blast + +lemma field_differentiable_within_tan: "~(cos z = 0) + \<Longrightarrow> tan field_differentiable (at z within s)" + using field_differentiable_at_tan field_differentiable_at_within by blast + +lemma continuous_within_tan: "~(cos z = 0) \<Longrightarrow> continuous (at z within s) tan" + using continuous_at_imp_continuous_within isCont_tan by blast + +lemma continuous_on_tan [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 0)) \<Longrightarrow> continuous_on s tan" + by (simp add: continuous_at_imp_continuous_on) + +lemma holomorphic_on_tan: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 0)) \<Longrightarrow> tan holomorphic_on s" + by (simp add: field_differentiable_within_tan holomorphic_on_def) + + +subsection\<open>Complex logarithms (the conventional principal value)\<close> + +instantiation complex :: ln +begin + +definition ln_complex :: "complex \<Rightarrow> complex" + where "ln_complex \<equiv> \<lambda>z. THE w. exp w = z & -pi < Im(w) & Im(w) \<le> pi" + +lemma + assumes "z \<noteq> 0" + shows exp_Ln [simp]: "exp(ln z) = z" + and mpi_less_Im_Ln: "-pi < Im(ln z)" + and Im_Ln_le_pi: "Im(ln z) \<le> pi" +proof - + obtain \<psi> where z: "z / (cmod z) = Complex (cos \<psi>) (sin \<psi>)" + using complex_unimodular_polar [of "z / (norm z)"] assms + by (auto simp: norm_divide divide_simps) + obtain \<phi> where \<phi>: "- pi < \<phi>" "\<phi> \<le> pi" "sin \<phi> = sin \<psi>" "cos \<phi> = cos \<psi>" + using sincos_principal_value [of "\<psi>"] assms + by (auto simp: norm_divide divide_simps) + have "exp(ln z) = z & -pi < Im(ln z) & Im(ln z) \<le> pi" unfolding ln_complex_def + apply (rule theI [where a = "Complex (ln(norm z)) \<phi>"]) + using z assms \<phi> + apply (auto simp: field_simps exp_complex_eqI exp_eq_polar cis.code) + done + then show "exp(ln z) = z" "-pi < Im(ln z)" "Im(ln z) \<le> pi" + by auto +qed + +lemma Ln_exp [simp]: + assumes "-pi < Im(z)" "Im(z) \<le> pi" + shows "ln(exp z) = z" + apply (rule exp_complex_eqI) + using assms mpi_less_Im_Ln [of "exp z"] Im_Ln_le_pi [of "exp z"] + apply auto + done + +subsection\<open>Relation to Real Logarithm\<close> + +lemma Ln_of_real: + assumes "0 < z" + shows "ln(of_real z::complex) = of_real(ln z)" +proof - + have "ln(of_real (exp (ln z))::complex) = ln (exp (of_real (ln z)))" + by (simp add: exp_of_real) + also have "... = of_real(ln z)" + using assms + by (subst Ln_exp) auto + finally show ?thesis + using assms by simp +qed + +corollary Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<real>" + by (auto simp: Ln_of_real elim: Reals_cases) + +corollary Im_Ln_of_real [simp]: "r > 0 \<Longrightarrow> Im (ln (of_real r)) = 0" + by (simp add: Ln_of_real) + +lemma cmod_Ln_Reals [simp]: "z \<in> \<real> \<Longrightarrow> 0 < Re z \<Longrightarrow> cmod (ln z) = norm (ln (Re z))" + using Ln_of_real by force + +lemma Ln_1: "ln 1 = (0::complex)" +proof - + have "ln (exp 0) = (0::complex)" + by (metis (mono_tags, hide_lams) Ln_of_real exp_zero ln_one of_real_0 of_real_1 zero_less_one) + then show ?thesis + by simp +qed + +instance + by intro_classes (rule ln_complex_def Ln_1) + +end + +abbreviation Ln :: "complex \<Rightarrow> complex" + where "Ln \<equiv> ln" + +lemma Ln_eq_iff: "w \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> (Ln w = Ln z \<longleftrightarrow> w = z)" + by (metis exp_Ln) + +lemma Ln_unique: "exp(z) = w \<Longrightarrow> -pi < Im(z) \<Longrightarrow> Im(z) \<le> pi \<Longrightarrow> Ln w = z" + using Ln_exp by blast + +lemma Re_Ln [simp]: "z \<noteq> 0 \<Longrightarrow> Re(Ln z) = ln(norm z)" + by (metis exp_Ln ln_exp norm_exp_eq_Re) + +corollary ln_cmod_le: + assumes z: "z \<noteq> 0" + shows "ln (cmod z) \<le> cmod (Ln z)" + using norm_exp [of "Ln z", simplified exp_Ln [OF z]] + by (metis Re_Ln complex_Re_le_cmod z) + +proposition exists_complex_root: + fixes z :: complex + assumes "n \<noteq> 0" obtains w where "z = w ^ n" + apply (cases "z=0") + using assms apply (simp add: power_0_left) + apply (rule_tac w = "exp(Ln z / n)" in that) + apply (auto simp: assms exp_of_nat_mult [symmetric]) + done + +corollary exists_complex_root_nonzero: + fixes z::complex + assumes "z \<noteq> 0" "n \<noteq> 0" + obtains w where "w \<noteq> 0" "z = w ^ n" + by (metis exists_complex_root [of n z] assms power_0_left) + +subsection\<open>The Unwinding Number and the Ln-product Formula\<close> + +text\<open>Note that in this special case the unwinding number is -1, 0 or 1.\<close> + +definition unwinding :: "complex \<Rightarrow> complex" where + "unwinding(z) = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)" + +lemma unwinding_2pi: "(2*pi) * \<i> * unwinding(z) = z - Ln(exp z)" + by (simp add: unwinding_def) + +lemma Ln_times_unwinding: + "w \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \<i> * unwinding(Ln w + Ln z)" + using unwinding_2pi by (simp add: exp_add) + + +subsection\<open>Derivative of Ln away from the branch cut\<close> + +lemma + assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" + shows has_field_derivative_Ln: "(Ln has_field_derivative inverse(z)) (at z)" + and Im_Ln_less_pi: "Im (Ln z) < pi" +proof - + have znz: "z \<noteq> 0" + using assms by auto + then have "Im (Ln z) \<noteq> pi" + by (metis (no_types) Im_exp Ln_in_Reals assms complex_nonpos_Reals_iff complex_is_Real_iff exp_Ln mult_zero_right not_less pi_neq_zero sin_pi znz) + then show *: "Im (Ln z) < pi" using assms Im_Ln_le_pi + by (simp add: le_neq_trans znz) + have "(exp has_field_derivative z) (at (Ln z))" + by (metis znz DERIV_exp exp_Ln) + then show "(Ln has_field_derivative inverse(z)) (at z)" + apply (rule has_complex_derivative_inverse_strong_x + [where s = "{w. -pi < Im(w) \<and> Im(w) < pi}"]) + using znz * + apply (auto simp: Transcendental.continuous_on_exp [OF continuous_on_id] open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt mpi_less_Im_Ln) + done +qed + +declare has_field_derivative_Ln [derivative_intros] +declare has_field_derivative_Ln [THEN DERIV_chain2, derivative_intros] + +lemma field_differentiable_at_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Ln field_differentiable at z" + using field_differentiable_def has_field_derivative_Ln by blast + +lemma field_differentiable_within_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 + \<Longrightarrow> Ln field_differentiable (at z within s)" + using field_differentiable_at_Ln field_differentiable_within_subset by blast + +lemma continuous_at_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z) Ln" + by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Ln) + +lemma isCont_Ln' [simp]: + "\<lbrakk>isCont f z; f z \<notin> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. Ln (f x)) z" + by (blast intro: isCont_o2 [OF _ continuous_at_Ln]) + +lemma continuous_within_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within s) Ln" + using continuous_at_Ln continuous_at_imp_continuous_within by blast + +lemma continuous_on_Ln [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on s Ln" + by (simp add: continuous_at_imp_continuous_on continuous_within_Ln) + +lemma holomorphic_on_Ln: "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> Ln holomorphic_on s" + by (simp add: field_differentiable_within_Ln holomorphic_on_def) + + +subsection\<open>Quadrant-type results for Ln\<close> + +lemma cos_lt_zero_pi: "pi/2 < x \<Longrightarrow> x < 3*pi/2 \<Longrightarrow> cos x < 0" + using cos_minus_pi cos_gt_zero_pi [of "x-pi"] + by simp + +lemma Re_Ln_pos_lt: + assumes "z \<noteq> 0" + shows "\<bar>Im(Ln z)\<bar> < pi/2 \<longleftrightarrow> 0 < Re(z)" +proof - + { fix w + assume "w = Ln z" + then have w: "Im w \<le> pi" "- pi < Im w" + using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms + by auto + then have "\<bar>Im w\<bar> < pi/2 \<longleftrightarrow> 0 < Re(exp w)" + apply (auto simp: Re_exp zero_less_mult_iff cos_gt_zero_pi) + using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"] + apply (simp add: abs_if split: if_split_asm) + apply (metis (no_types) cos_minus cos_pi_half eq_divide_eq_numeral1(1) eq_numeral_simps(4) + less_numeral_extra(3) linorder_neqE_linordered_idom minus_mult_minus minus_mult_right + mult_numeral_1_right) + done + } + then show ?thesis using assms + by auto +qed + +lemma Re_Ln_pos_le: + assumes "z \<noteq> 0" + shows "\<bar>Im(Ln z)\<bar> \<le> pi/2 \<longleftrightarrow> 0 \<le> Re(z)" +proof - + { fix w + assume "w = Ln z" + then have w: "Im w \<le> pi" "- pi < Im w" + using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms + by auto + then have "\<bar>Im w\<bar> \<le> pi/2 \<longleftrightarrow> 0 \<le> Re(exp w)" + apply (auto simp: Re_exp zero_le_mult_iff cos_ge_zero) + using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le + apply (auto simp: abs_if split: if_split_asm) + done + } + then show ?thesis using assms + by auto +qed + +lemma Im_Ln_pos_lt: + assumes "z \<noteq> 0" + shows "0 < Im(Ln z) \<and> Im(Ln z) < pi \<longleftrightarrow> 0 < Im(z)" +proof - + { fix w + assume "w = Ln z" + then have w: "Im w \<le> pi" "- pi < Im w" + using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms + by auto + then have "0 < Im w \<and> Im w < pi \<longleftrightarrow> 0 < Im(exp w)" + using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"] + apply (auto simp: Im_exp zero_less_mult_iff) + using less_linear apply fastforce + using less_linear apply fastforce + done + } + then show ?thesis using assms + by auto +qed + +lemma Im_Ln_pos_le: + assumes "z \<noteq> 0" + shows "0 \<le> Im(Ln z) \<and> Im(Ln z) \<le> pi \<longleftrightarrow> 0 \<le> Im(z)" +proof - + { fix w + assume "w = Ln z" + then have w: "Im w \<le> pi" "- pi < Im w" + using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms + by auto + then have "0 \<le> Im w \<and> Im w \<le> pi \<longleftrightarrow> 0 \<le> Im(exp w)" + using sin_ge_zero [of "- (Im w)"] sin_ge_zero [of "(Im w)"] + apply (auto simp: Im_exp zero_le_mult_iff sin_ge_zero) + apply (metis not_le not_less_iff_gr_or_eq pi_not_less_zero sin_eq_0_pi) + done } + then show ?thesis using assms + by auto +qed + +lemma Re_Ln_pos_lt_imp: "0 < Re(z) \<Longrightarrow> \<bar>Im(Ln z)\<bar> < pi/2" + by (metis Re_Ln_pos_lt less_irrefl zero_complex.simps(1)) + +lemma Im_Ln_pos_lt_imp: "0 < Im(z) \<Longrightarrow> 0 < Im(Ln z) \<and> Im(Ln z) < pi" + by (metis Im_Ln_pos_lt not_le order_refl zero_complex.simps(2)) + +text\<open>A reference to the set of positive real numbers\<close> +lemma Im_Ln_eq_0: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = 0 \<longleftrightarrow> 0 < Re(z) \<and> Im(z) = 0)" +by (metis Im_complex_of_real Im_exp Ln_in_Reals Re_Ln_pos_lt Re_Ln_pos_lt_imp + Re_complex_of_real complex_is_Real_iff exp_Ln exp_of_real pi_gt_zero) + +lemma Im_Ln_eq_pi: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = pi \<longleftrightarrow> Re(z) < 0 \<and> Im(z) = 0)" +by (metis Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt add.left_neutral complex_eq less_eq_real_def + mult_zero_right not_less_iff_gr_or_eq pi_ge_zero pi_neq_zero rcis_zero_arg rcis_zero_mod) + + +subsection\<open>More Properties of Ln\<close> + +lemma cnj_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> cnj(Ln z) = Ln(cnj z)" + apply (cases "z=0", auto) + apply (rule exp_complex_eqI) + apply (auto simp: abs_if split: if_split_asm) + using Im_Ln_less_pi Im_Ln_le_pi apply force + apply (metis complex_cnj_zero_iff diff_minus_eq_add diff_strict_mono minus_less_iff + mpi_less_Im_Ln mult.commute mult_2_right) + by (metis exp_Ln exp_cnj) + +lemma Ln_inverse: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Ln(inverse z) = -(Ln z)" + apply (cases "z=0", auto) + apply (rule exp_complex_eqI) + using mpi_less_Im_Ln [of z] mpi_less_Im_Ln [of "inverse z"] + apply (auto simp: abs_if exp_minus split: if_split_asm) + apply (metis Im_Ln_less_pi Im_Ln_le_pi add.commute add_mono_thms_linordered_field(3) inverse_nonzero_iff_nonzero mult_2) + done + +lemma Ln_minus1 [simp]: "Ln(-1) = \<i> * pi" + apply (rule exp_complex_eqI) + using Im_Ln_le_pi [of "-1"] mpi_less_Im_Ln [of "-1"] cis_conv_exp cis_pi + apply (auto simp: abs_if) + done + +lemma Ln_ii [simp]: "Ln \<i> = \<i> * of_real pi/2" + using Ln_exp [of "\<i> * (of_real pi/2)"] + unfolding exp_Euler + by simp + +lemma Ln_minus_ii [simp]: "Ln(-\<i>) = - (\<i> * pi/2)" +proof - + have "Ln(-\<i>) = Ln(inverse \<i>)" by simp + also have "... = - (Ln \<i>)" using Ln_inverse by blast + also have "... = - (\<i> * pi/2)" by simp + finally show ?thesis . +qed + +lemma Ln_times: + assumes "w \<noteq> 0" "z \<noteq> 0" + shows "Ln(w * z) = + (if Im(Ln w + Ln z) \<le> -pi then + (Ln(w) + Ln(z)) + \<i> * of_real(2*pi) + else if Im(Ln w + Ln z) > pi then + (Ln(w) + Ln(z)) - \<i> * of_real(2*pi) + else Ln(w) + Ln(z))" + using pi_ge_zero Im_Ln_le_pi [of w] Im_Ln_le_pi [of z] + using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z] + by (auto simp: exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique) + +corollary Ln_times_simple: + "\<lbrakk>w \<noteq> 0; z \<noteq> 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \<le> pi\<rbrakk> + \<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z)" + by (simp add: Ln_times) + +corollary Ln_times_of_real: + "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(of_real r * z) = ln r + Ln(z)" + using mpi_less_Im_Ln Im_Ln_le_pi + by (force simp: Ln_times) + +corollary Ln_divide_of_real: + "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(z / of_real r) = Ln(z) - ln r" +using Ln_times_of_real [of "inverse r" z] +by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric] + del: of_real_inverse) + +lemma Ln_minus: + assumes "z \<noteq> 0" + shows "Ln(-z) = (if Im(z) \<le> 0 \<and> ~(Re(z) < 0 \<and> Im(z) = 0) + then Ln(z) + \<i> * pi + else Ln(z) - \<i> * pi)" (is "_ = ?rhs") + using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms + Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] + by (fastforce simp: exp_add exp_diff exp_Euler intro!: Ln_unique) + +lemma Ln_inverse_if: + assumes "z \<noteq> 0" + shows "Ln (inverse z) = (if z \<in> \<real>\<^sub>\<le>\<^sub>0 then -(Ln z) + \<i> * 2 * complex_of_real pi else -(Ln z))" +proof (cases "z \<in> \<real>\<^sub>\<le>\<^sub>0") + case False then show ?thesis + by (simp add: Ln_inverse) +next + case True + then have z: "Im z = 0" "Re z < 0" + using assms + apply (auto simp: complex_nonpos_Reals_iff) + by (metis complex_is_Real_iff le_imp_less_or_eq of_real_0 of_real_Re) + have "Ln(inverse z) = Ln(- (inverse (-z)))" + by simp + also have "... = Ln (inverse (-z)) + \<i> * complex_of_real pi" + using assms z + apply (simp add: Ln_minus) + apply (simp add: field_simps) + done + also have "... = - Ln (- z) + \<i> * complex_of_real pi" + apply (subst Ln_inverse) + using z by (auto simp add: complex_nonneg_Reals_iff) + also have "... = - (Ln z) + \<i> * 2 * complex_of_real pi" + apply (subst Ln_minus [OF assms]) + using assms z + apply simp + done + finally show ?thesis by (simp add: True) +qed + +lemma Ln_times_ii: + assumes "z \<noteq> 0" + shows "Ln(\<i> * z) = (if 0 \<le> Re(z) | Im(z) < 0 + then Ln(z) + \<i> * of_real pi/2 + else Ln(z) - \<i> * of_real(3 * pi/2))" + using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms + Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] Re_Ln_pos_le [of z] + by (auto simp: Ln_times) + +lemma Ln_of_nat: "0 < n \<Longrightarrow> Ln (of_nat n) = of_real (ln (of_nat n))" + by (subst of_real_of_nat_eq[symmetric], subst Ln_of_real[symmetric]) simp_all + +lemma Ln_of_nat_over_of_nat: + assumes "m > 0" "n > 0" + shows "Ln (of_nat m / of_nat n) = of_real (ln (of_nat m) - ln (of_nat n))" +proof - + have "of_nat m / of_nat n = (of_real (of_nat m / of_nat n) :: complex)" by simp + also from assms have "Ln ... = of_real (ln (of_nat m / of_nat n))" + by (simp add: Ln_of_real[symmetric]) + also from assms have "... = of_real (ln (of_nat m) - ln (of_nat n))" + by (simp add: ln_div) + finally show ?thesis . +qed + + +subsection\<open>Relation between Ln and Arg, and hence continuity of Arg\<close> + +lemma Arg_Ln: + assumes "0 < Arg z" shows "Arg z = Im(Ln(-z)) + pi" +proof (cases "z = 0") + case True + with assms show ?thesis + by simp +next + case False + then have "z / of_real(norm z) = exp(\<i> * of_real(Arg z))" + using Arg [of z] + by (metis abs_norm_cancel nonzero_mult_divide_cancel_left norm_of_real zero_less_norm_iff) + then have "- z / of_real(norm z) = exp (\<i> * (of_real (Arg z) - pi))" + using cis_conv_exp cis_pi + by (auto simp: exp_diff algebra_simps) + then have "ln (- z / of_real(norm z)) = ln (exp (\<i> * (of_real (Arg z) - pi)))" + by simp + also have "... = \<i> * (of_real(Arg z) - pi)" + using Arg [of z] assms pi_not_less_zero + by auto + finally have "Arg z = Im (Ln (- z / of_real (cmod z))) + pi" + by simp + also have "... = Im (Ln (-z) - ln (cmod z)) + pi" + by (metis diff_0_right minus_diff_eq zero_less_norm_iff Ln_divide_of_real False) + also have "... = Im (Ln (-z)) + pi" + by simp + finally show ?thesis . +qed + +lemma continuous_at_Arg: + assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0" + shows "continuous (at z) Arg" +proof - + have *: "isCont (\<lambda>z. Im (Ln (- z)) + pi) z" + by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+ + have [simp]: "\<And>x. \<lbrakk>Im x \<noteq> 0\<rbrakk> \<Longrightarrow> Im (Ln (- x)) + pi = Arg x" + using Arg_Ln Arg_gt_0 complex_is_Real_iff by auto + consider "Re z < 0" | "Im z \<noteq> 0" using assms + using complex_nonneg_Reals_iff not_le by blast + then have [simp]: "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg z" + using "*" by (simp add: isCont_def) (metis Arg_Ln Arg_gt_0 complex_is_Real_iff) + show ?thesis + apply (simp add: continuous_at) + apply (rule Lim_transform_within_open [where s= "-\<real>\<^sub>\<ge>\<^sub>0" and f = "\<lambda>z. Im(Ln(-z)) + pi"]) + apply (auto simp add: not_le Arg_Ln [OF Arg_gt_0] complex_nonneg_Reals_iff closed_def [symmetric]) + using assms apply (force simp add: complex_nonneg_Reals_iff) + done +qed + +lemma Ln_series: + fixes z :: complex + assumes "norm z < 1" + shows "(\<lambda>n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\<lambda>n. ?f n * z^n) sums _") +proof - + let ?F = "\<lambda>z. \<Sum>n. ?f n * z^n" and ?F' = "\<lambda>z. \<Sum>n. diffs ?f n * z^n" + have r: "conv_radius ?f = 1" + by (intro conv_radius_ratio_limit_nonzero[of _ 1]) + (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc) + + have "\<exists>c. \<forall>z\<in>ball 0 1. ln (1 + z) - ?F z = c" + proof (rule has_field_derivative_zero_constant) + fix z :: complex assume z': "z \<in> ball 0 1" + hence z: "norm z < 1" by (simp add: dist_0_norm) + define t :: complex where "t = of_real (1 + norm z) / 2" + from z have t: "norm z < norm t" "norm t < 1" unfolding t_def + by (simp_all add: field_simps norm_divide del: of_real_add) + + have "Re (-z) \<le> norm (-z)" by (rule complex_Re_le_cmod) + also from z have "... < 1" by simp + finally have "((\<lambda>z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)" + by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff) + moreover have "(?F has_field_derivative ?F' z) (at z)" using t r + by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all + ultimately have "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z)) + (at z within ball 0 1)" + by (intro derivative_intros) (simp_all add: at_within_open[OF z']) + also have "(\<lambda>n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r + by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all + from sums_split_initial_segment[OF this, of 1] + have "(\<lambda>i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc) + hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse) + also have "inverse (1 + z) - inverse (1 + z) = 0" by simp + finally show "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" . + qed simp_all + then obtain c where c: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> ln (1 + z) - ?F z = c" by blast + from c[of 0] have "c = 0" by (simp only: powser_zero) simp + with c[of z] assms have "ln (1 + z) = ?F z" by (simp add: dist_0_norm) + moreover have "summable (\<lambda>n. ?f n * z^n)" using assms r + by (intro summable_in_conv_radius) simp_all + ultimately show ?thesis by (simp add: sums_iff) +qed + +lemma Ln_approx_linear: + fixes z :: complex + assumes "norm z < 1" + shows "norm (ln (1 + z) - z) \<le> norm z^2 / (1 - norm z)" +proof - + let ?f = "\<lambda>n. (-1)^Suc n / of_nat n" + from assms have "(\<lambda>n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp + moreover have "(\<lambda>n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp + ultimately have "(\<lambda>n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)" + by (subst left_diff_distrib, intro sums_diff) simp_all + from sums_split_initial_segment[OF this, of "Suc 1"] + have "(\<lambda>i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)" + by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse) + hence "(Ln (1 + z) - z) = (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)" + by (simp add: sums_iff) + also have A: "summable (\<lambda>n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))" + by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]]) + (auto simp: assms field_simps intro!: always_eventually) + hence "norm (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \<le> + (\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))" + by (intro summable_norm) + (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc) + also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \<le> norm ((-z)^2 * (-z)^i) * 1" for i + by (intro mult_left_mono) (simp_all add: divide_simps) + hence "(\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i)) \<le> + (\<Sum>i. norm (-(z^2) * (-z)^i))" using A assms + apply (simp_all only: norm_power norm_inverse norm_divide norm_mult) + apply (intro suminf_le summable_mult summable_geometric) + apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc) + done + also have "... = norm z^2 * (\<Sum>i. norm z^i)" using assms + by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power) + also have "(\<Sum>i. norm z^i) = inverse (1 - norm z)" using assms + by (subst suminf_geometric) (simp_all add: divide_inverse) + also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse) + finally show ?thesis . +qed + + +text\<open>Relation between Arg and arctangent in upper halfplane\<close> +lemma Arg_arctan_upperhalf: + assumes "0 < Im z" + shows "Arg z = pi/2 - arctan(Re z / Im z)" +proof (cases "z = 0") + case True with assms show ?thesis + by simp +next + case False + show ?thesis + apply (rule Arg_unique [of "norm z"]) + using False assms arctan [of "Re z / Im z"] pi_ge_two pi_half_less_two + apply (auto simp: exp_Euler cos_diff sin_diff) + using norm_complex_def [of z, symmetric] + apply (simp add: sin_of_real cos_of_real sin_arctan cos_arctan field_simps real_sqrt_divide) + apply (metis complex_eq mult.assoc ring_class.ring_distribs(2)) + done +qed + +lemma Arg_eq_Im_Ln: + assumes "0 \<le> Im z" "0 < Re z" + shows "Arg z = Im (Ln z)" +proof (cases "z = 0 \<or> Im z = 0") + case True then show ?thesis + using assms Arg_eq_0 complex_is_Real_iff + apply auto + by (metis Arg_eq_0_pi Arg_eq_pi Im_Ln_eq_0 Im_Ln_eq_pi less_numeral_extra(3) zero_complex.simps(1)) +next + case False + then have "Arg z > 0" + using Arg_gt_0 complex_is_Real_iff by blast + then show ?thesis + using assms False + by (subst Arg_Ln) (auto simp: Ln_minus) +qed + +lemma continuous_within_upperhalf_Arg: + assumes "z \<noteq> 0" + shows "continuous (at z within {z. 0 \<le> Im z}) Arg" +proof (cases "z \<in> \<real>\<^sub>\<ge>\<^sub>0") + case False then show ?thesis + using continuous_at_Arg continuous_at_imp_continuous_within by auto +next + case True + then have z: "z \<in> \<real>" "0 < Re z" + using assms by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0) + then have [simp]: "Arg z = 0" "Im (Ln z) = 0" + by (auto simp: Arg_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff) + show ?thesis + proof (clarsimp simp add: continuous_within Lim_within dist_norm) + fix e::real + assume "0 < e" + moreover have "continuous (at z) (\<lambda>x. Im (Ln x))" + using z by (simp add: continuous_at_Ln complex_nonpos_Reals_iff) + ultimately + obtain d where d: "d>0" "\<And>x. x \<noteq> z \<Longrightarrow> cmod (x - z) < d \<Longrightarrow> \<bar>Im (Ln x)\<bar> < e" + by (auto simp: continuous_within Lim_within dist_norm) + { fix x + assume "cmod (x - z) < Re z / 2" + then have "\<bar>Re x - Re z\<bar> < Re z / 2" + by (metis le_less_trans abs_Re_le_cmod minus_complex.simps(1)) + then have "0 < Re x" + using z by linarith + } + then show "\<exists>d>0. \<forall>x. 0 \<le> Im x \<longrightarrow> x \<noteq> z \<and> cmod (x - z) < d \<longrightarrow> \<bar>Arg x\<bar> < e" + apply (rule_tac x="min d (Re z / 2)" in exI) + using z d + apply (auto simp: Arg_eq_Im_Ln) + done + qed +qed + +lemma continuous_on_upperhalf_Arg: "continuous_on ({z. 0 \<le> Im z} - {0}) Arg" + apply (auto simp: continuous_on_eq_continuous_within) + by (metis Diff_subset continuous_within_subset continuous_within_upperhalf_Arg) + +lemma open_Arg_less_Int: + assumes "0 \<le> s" "t \<le> 2*pi" + shows "open ({y. s < Arg y} \<inter> {y. Arg y < t})" +proof - + have 1: "continuous_on (UNIV - \<real>\<^sub>\<ge>\<^sub>0) Arg" + using continuous_at_Arg continuous_at_imp_continuous_within + by (auto simp: continuous_on_eq_continuous_within) + have 2: "open (UNIV - \<real>\<^sub>\<ge>\<^sub>0 :: complex set)" by (simp add: open_Diff) + have "open ({z. s < z} \<inter> {z. z < t})" + using open_lessThan [of t] open_greaterThan [of s] + by (metis greaterThan_def lessThan_def open_Int) + moreover have "{y. s < Arg y} \<inter> {y. Arg y < t} \<subseteq> - \<real>\<^sub>\<ge>\<^sub>0" + using assms by (auto simp: Arg_real complex_nonneg_Reals_iff complex_is_Real_iff) + ultimately show ?thesis + using continuous_imp_open_vimage [OF 1 2, of "{z. Re z > s} \<inter> {z. Re z < t}"] + by auto +qed + +lemma open_Arg_gt: "open {z. t < Arg z}" +proof (cases "t < 0") + case True then have "{z. t < Arg z} = UNIV" + using Arg_ge_0 less_le_trans by auto + then show ?thesis + by simp +next + case False then show ?thesis + using open_Arg_less_Int [of t "2*pi"] Arg_lt_2pi + by auto +qed + +lemma closed_Arg_le: "closed {z. Arg z \<le> t}" + using open_Arg_gt [of t] + by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le) + +subsection\<open>Complex Powers\<close> + +lemma powr_to_1 [simp]: "z powr 1 = (z::complex)" + by (simp add: powr_def) + +lemma powr_nat: + fixes n::nat and z::complex shows "z powr n = (if z = 0 then 0 else z^n)" + by (simp add: exp_of_nat_mult powr_def) + +lemma powr_add_complex: + fixes w::complex shows "w powr (z1 + z2) = w powr z1 * w powr z2" + by (simp add: powr_def algebra_simps exp_add) + +lemma powr_minus_complex: + fixes w::complex shows "w powr (-z) = inverse(w powr z)" + by (simp add: powr_def exp_minus) + +lemma powr_diff_complex: + fixes w::complex shows "w powr (z1 - z2) = w powr z1 / w powr z2" + by (simp add: powr_def algebra_simps exp_diff) + +lemma norm_powr_real: "w \<in> \<real> \<Longrightarrow> 0 < Re w \<Longrightarrow> norm(w powr z) = exp(Re z * ln(Re w))" + apply (simp add: powr_def) + using Im_Ln_eq_0 complex_is_Real_iff norm_complex_def + by auto + +lemma cnj_powr: + assumes "Im a = 0 \<Longrightarrow> Re a \<ge> 0" + shows "cnj (a powr b) = cnj a powr cnj b" +proof (cases "a = 0") + case False + with assms have "a \<notin> \<real>\<^sub>\<le>\<^sub>0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) + with False show ?thesis by (simp add: powr_def exp_cnj cnj_Ln) +qed simp + +lemma powr_real_real: + "\<lbrakk>w \<in> \<real>; z \<in> \<real>; 0 < Re w\<rbrakk> \<Longrightarrow> w powr z = exp(Re z * ln(Re w))" + apply (simp add: powr_def) + by (metis complex_eq complex_is_Real_iff diff_0 diff_0_right diff_minus_eq_add exp_ln exp_not_eq_zero + exp_of_real Ln_of_real mult_zero_right of_real_0 of_real_mult) + +lemma powr_of_real: + fixes x::real and y::real + shows "0 \<le> x \<Longrightarrow> of_real x powr (of_real y::complex) = of_real (x powr y)" + by (simp_all add: powr_def exp_eq_polar) + +lemma norm_powr_real_mono: + "\<lbrakk>w \<in> \<real>; 1 < Re w\<rbrakk> + \<Longrightarrow> cmod(w powr z1) \<le> cmod(w powr z2) \<longleftrightarrow> Re z1 \<le> Re z2" + by (auto simp: powr_def algebra_simps Reals_def Ln_of_real) + +lemma powr_times_real: + "\<lbrakk>x \<in> \<real>; y \<in> \<real>; 0 \<le> Re x; 0 \<le> Re y\<rbrakk> + \<Longrightarrow> (x * y) powr z = x powr z * y powr z" + by (auto simp: Reals_def powr_def Ln_times exp_add algebra_simps less_eq_real_def Ln_of_real) + +lemma powr_neg_real_complex: + shows "(- of_real x) powr a = (-1) powr (of_real (sgn x) * a) * of_real x powr (a :: complex)" +proof (cases "x = 0") + assume x: "x \<noteq> 0" + hence "(-x) powr a = exp (a * ln (-of_real x))" by (simp add: powr_def) + also from x have "ln (-of_real x) = Ln (of_real x) + of_real (sgn x) * pi * \<i>" + by (simp add: Ln_minus Ln_of_real) + also from x have "exp (a * ...) = cis pi powr (of_real (sgn x) * a) * of_real x powr a" + by (simp add: powr_def exp_add algebra_simps Ln_of_real cis_conv_exp) + also note cis_pi + finally show ?thesis by simp +qed simp_all + +lemma has_field_derivative_powr: + fixes z :: complex + shows "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ((\<lambda>z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)" + apply (cases "z=0", auto) + apply (simp add: powr_def) + apply (rule DERIV_transform_at [where d = "norm z" and f = "\<lambda>z. exp (s * Ln z)"]) + apply (auto simp: dist_complex_def) + apply (intro derivative_eq_intros | simp)+ + apply (simp add: field_simps exp_diff) + done + +declare has_field_derivative_powr[THEN DERIV_chain2, derivative_intros] + + +lemma has_field_derivative_powr_right: + "w \<noteq> 0 \<Longrightarrow> ((\<lambda>z. w powr z) has_field_derivative Ln w * w powr z) (at z)" + apply (simp add: powr_def) + apply (intro derivative_eq_intros | simp)+ + done + +lemma field_differentiable_powr_right: + fixes w::complex + shows + "w \<noteq> 0 \<Longrightarrow> (\<lambda>z. w powr z) field_differentiable (at z)" +using field_differentiable_def has_field_derivative_powr_right by blast + +lemma holomorphic_on_powr_right: + "f holomorphic_on s \<Longrightarrow> w \<noteq> 0 \<Longrightarrow> (\<lambda>z. w powr (f z)) holomorphic_on s" + unfolding holomorphic_on_def field_differentiable_def +by (metis (full_types) DERIV_chain' has_field_derivative_powr_right) + +lemma norm_powr_real_powr: + "w \<in> \<real> \<Longrightarrow> 0 \<le> Re w \<Longrightarrow> cmod (w powr z) = Re w powr Re z" + by (cases "w = 0") (auto simp add: norm_powr_real powr_def Im_Ln_eq_0 + complex_is_Real_iff in_Reals_norm complex_eq_iff) + +lemma tendsto_ln_complex [tendsto_intros]: + assumes "(f \<longlongrightarrow> a) F" "a \<notin> \<real>\<^sub>\<le>\<^sub>0" + shows "((\<lambda>z. ln (f z :: complex)) \<longlongrightarrow> ln a) F" + using tendsto_compose[OF continuous_at_Ln[of a, unfolded isCont_def] assms(1)] assms(2) by simp + +lemma tendsto_powr_complex: + fixes f g :: "_ \<Rightarrow> complex" + assumes a: "a \<notin> \<real>\<^sub>\<le>\<^sub>0" + assumes f: "(f \<longlongrightarrow> a) F" and g: "(g \<longlongrightarrow> b) F" + shows "((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F" +proof - + from a have [simp]: "a \<noteq> 0" by auto + from f g a have "((\<lambda>z. exp (g z * ln (f z))) \<longlongrightarrow> a powr b) F" (is ?P) + by (auto intro!: tendsto_intros simp: powr_def) + also { + have "eventually (\<lambda>z. z \<noteq> 0) (nhds a)" + by (intro t1_space_nhds) simp_all + with f have "eventually (\<lambda>z. f z \<noteq> 0) F" using filterlim_iff by blast + } + hence "?P \<longleftrightarrow> ((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F" + by (intro tendsto_cong refl) (simp_all add: powr_def mult_ac) + finally show ?thesis . +qed + +lemma tendsto_powr_complex_0: + fixes f g :: "'a \<Rightarrow> complex" + assumes f: "(f \<longlongrightarrow> 0) F" and g: "(g \<longlongrightarrow> b) F" and b: "Re b > 0" + shows "((\<lambda>z. f z powr g z) \<longlongrightarrow> 0) F" +proof (rule tendsto_norm_zero_cancel) + define h where + "h = (\<lambda>z. if f z = 0 then 0 else exp (Re (g z) * ln (cmod (f z)) + abs (Im (g z)) * pi))" + { + fix z :: 'a assume z: "f z \<noteq> 0" + define c where "c = abs (Im (g z)) * pi" + from mpi_less_Im_Ln[OF z] Im_Ln_le_pi[OF z] + have "abs (Im (Ln (f z))) \<le> pi" by simp + from mult_left_mono[OF this, of "abs (Im (g z))"] + have "abs (Im (g z) * Im (ln (f z))) \<le> c" by (simp add: abs_mult c_def) + hence "-Im (g z) * Im (ln (f z)) \<le> c" by simp + hence "norm (f z powr g z) \<le> h z" by (simp add: powr_def field_simps h_def c_def) + } + hence le: "norm (f z powr g z) \<le> h z" for z by (cases "f x = 0") (simp_all add: h_def) + + have g': "(g \<longlongrightarrow> b) (inf F (principal {z. f z \<noteq> 0}))" + by (rule tendsto_mono[OF _ g]) simp_all + have "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) (inf F (principal {z. f z \<noteq> 0}))" + by (subst tendsto_norm_zero_iff, rule tendsto_mono[OF _ f]) simp_all + moreover { + have "filterlim (\<lambda>x. norm (f x)) (principal {0<..}) (principal {z. f z \<noteq> 0})" + by (auto simp: filterlim_def) + hence "filterlim (\<lambda>x. norm (f x)) (principal {0<..}) + (inf F (principal {z. f z \<noteq> 0}))" + by (rule filterlim_mono) simp_all + } + ultimately have norm: "filterlim (\<lambda>x. norm (f x)) (at_right 0) (inf F (principal {z. f z \<noteq> 0}))" + by (simp add: filterlim_inf at_within_def) + + have A: "LIM x inf F (principal {z. f z \<noteq> 0}). Re (g x) * -ln (cmod (f x)) :> at_top" + by (rule filterlim_tendsto_pos_mult_at_top tendsto_intros g' b + filterlim_compose[OF filterlim_uminus_at_top_at_bot] filterlim_compose[OF ln_at_0] norm)+ + have B: "LIM x inf F (principal {z. f z \<noteq> 0}). + -\<bar>Im (g x)\<bar> * pi + -(Re (g x) * ln (cmod (f x))) :> at_top" + by (rule filterlim_tendsto_add_at_top tendsto_intros g')+ (insert A, simp_all) + have C: "(h \<longlongrightarrow> 0) F" unfolding h_def + by (intro filterlim_If tendsto_const filterlim_compose[OF exp_at_bot]) + (insert B, auto simp: filterlim_uminus_at_bot algebra_simps) + show "((\<lambda>x. norm (f x powr g x)) \<longlongrightarrow> 0) F" + by (rule Lim_null_comparison[OF always_eventually C]) (insert le, auto) +qed + +lemma tendsto_powr_complex' [tendsto_intros]: + fixes f g :: "_ \<Rightarrow> complex" + assumes fz: "a \<notin> \<real>\<^sub>\<le>\<^sub>0 \<or> (a = 0 \<and> Re b > 0)" + assumes fg: "(f \<longlongrightarrow> a) F" "(g \<longlongrightarrow> b) F" + shows "((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F" +proof (cases "a = 0") + case True + with assms show ?thesis by (auto intro!: tendsto_powr_complex_0) +next + case False + with assms show ?thesis by (auto intro!: tendsto_powr_complex elim!: nonpos_Reals_cases) +qed + +lemma continuous_powr_complex: + assumes "f (netlimit F) \<notin> \<real>\<^sub>\<le>\<^sub>0" "continuous F f" "continuous F g" + shows "continuous F (\<lambda>z. f z powr g z :: complex)" + using assms unfolding continuous_def by (intro tendsto_powr_complex) simp_all + +lemma isCont_powr_complex [continuous_intros]: + assumes "f z \<notin> \<real>\<^sub>\<le>\<^sub>0" "isCont f z" "isCont g z" + shows "isCont (\<lambda>z. f z powr g z :: complex) z" + using assms unfolding isCont_def by (intro tendsto_powr_complex) simp_all + +lemma continuous_on_powr_complex [continuous_intros]: + assumes "A \<subseteq> {z. Re (f z) \<ge> 0 \<or> Im (f z) \<noteq> 0}" + assumes "\<And>z. z \<in> A \<Longrightarrow> f z = 0 \<Longrightarrow> Re (g z) > 0" + assumes "continuous_on A f" "continuous_on A g" + shows "continuous_on A (\<lambda>z. f z powr g z)" + unfolding continuous_on_def +proof + fix z assume z: "z \<in> A" + show "((\<lambda>z. f z powr g z) \<longlongrightarrow> f z powr g z) (at z within A)" + proof (cases "f z = 0") + case False + from assms(1,2) z have "Re (f z) \<ge> 0 \<or> Im (f z) \<noteq> 0" "f z = 0 \<longrightarrow> Re (g z) > 0" by auto + with assms(3,4) z show ?thesis + by (intro tendsto_powr_complex') + (auto elim!: nonpos_Reals_cases simp: complex_eq_iff continuous_on_def) + next + case True + with assms z show ?thesis + by (auto intro!: tendsto_powr_complex_0 simp: continuous_on_def) + qed +qed + + +subsection\<open>Some Limits involving Logarithms\<close> + +lemma lim_Ln_over_power: + fixes s::complex + assumes "0 < Re s" + shows "((\<lambda>n. Ln n / (n powr s)) \<longlongrightarrow> 0) sequentially" +proof (simp add: lim_sequentially dist_norm, clarify) + fix e::real + assume e: "0 < e" + have "\<exists>xo>0. \<forall>x\<ge>xo. 0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" + proof (rule_tac x="2/(e * (Re s)\<^sup>2)" in exI, safe) + show "0 < 2 / (e * (Re s)\<^sup>2)" + using e assms by (simp add: field_simps) + next + fix x::real + assume x: "2 / (e * (Re s)\<^sup>2) \<le> x" + then have "x>0" + using e assms + by (metis less_le_trans mult_eq_0_iff mult_pos_pos pos_less_divide_eq power2_eq_square + zero_less_numeral) + then show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" + using e assms x + apply (auto simp: field_simps) + apply (rule_tac y = "e * (x\<^sup>2 * (Re s)\<^sup>2)" in le_less_trans) + apply (auto simp: power2_eq_square field_simps add_pos_pos) + done + qed + then have "\<exists>xo>0. \<forall>x\<ge>xo. x / e < 1 + (Re s * x) + (1/2) * (Re s * x)^2" + using e by (simp add: field_simps) + then have "\<exists>xo>0. \<forall>x\<ge>xo. x / e < exp (Re s * x)" + using assms + by (force intro: less_le_trans [OF _ exp_lower_taylor_quadratic]) + then have "\<exists>xo>0. \<forall>x\<ge>xo. x < e * exp (Re s * x)" + using e by (auto simp: field_simps) + with e show "\<exists>no. \<forall>n\<ge>no. norm (Ln (of_nat n) / of_nat n powr s) < e" + apply (auto simp: norm_divide norm_powr_real divide_simps) + apply (rule_tac x="nat \<lceil>exp xo\<rceil>" in exI) + apply clarify + apply (drule_tac x="ln n" in spec) + apply (auto simp: exp_less_mono nat_ceiling_le_eq not_le) + apply (metis exp_less_mono exp_ln not_le of_nat_0_less_iff) + done +qed + +lemma lim_Ln_over_n: "((\<lambda>n. Ln(of_nat n) / of_nat n) \<longlongrightarrow> 0) sequentially" + using lim_Ln_over_power [of 1] + by simp + +lemma Ln_Reals_eq: "x \<in> \<real> \<Longrightarrow> Re x > 0 \<Longrightarrow> Ln x = of_real (ln (Re x))" + using Ln_of_real by force + +lemma powr_Reals_eq: "x \<in> \<real> \<Longrightarrow> Re x > 0 \<Longrightarrow> x powr complex_of_real y = of_real (x powr y)" + by (simp add: powr_of_real) + +lemma lim_ln_over_power: + fixes s :: real + assumes "0 < s" + shows "((\<lambda>n. ln n / (n powr s)) \<longlongrightarrow> 0) sequentially" + using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms + apply (subst filterlim_sequentially_Suc [symmetric]) + apply (simp add: lim_sequentially dist_norm + Ln_Reals_eq norm_powr_real_powr norm_divide) + done + +lemma lim_ln_over_n: "((\<lambda>n. ln(real_of_nat n) / of_nat n) \<longlongrightarrow> 0) sequentially" + using lim_ln_over_power [of 1, THEN filterlim_sequentially_Suc [THEN iffD2]] + apply (subst filterlim_sequentially_Suc [symmetric]) + apply (simp add: lim_sequentially dist_norm) + done + +lemma lim_1_over_complex_power: + assumes "0 < Re s" + shows "((\<lambda>n. 1 / (of_nat n powr s)) \<longlongrightarrow> 0) sequentially" +proof - + have "\<forall>n>0. 3 \<le> n \<longrightarrow> 1 \<le> ln (real_of_nat n)" + using ln3_gt_1 + by (force intro: order_trans [of _ "ln 3"] ln3_gt_1) + moreover have "(\<lambda>n. cmod (Ln (of_nat n) / of_nat n powr s)) \<longlonglongrightarrow> 0" + using lim_Ln_over_power [OF assms] + by (metis tendsto_norm_zero_iff) + ultimately show ?thesis + apply (auto intro!: Lim_null_comparison [where g = "\<lambda>n. norm (Ln(of_nat n) / of_nat n powr s)"]) + apply (auto simp: norm_divide divide_simps eventually_sequentially) + done +qed + +lemma lim_1_over_real_power: + fixes s :: real + assumes "0 < s" + shows "((\<lambda>n. 1 / (of_nat n powr s)) \<longlongrightarrow> 0) sequentially" + using lim_1_over_complex_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms + apply (subst filterlim_sequentially_Suc [symmetric]) + apply (simp add: lim_sequentially dist_norm) + apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) + done + +lemma lim_1_over_Ln: "((\<lambda>n. 1 / Ln(of_nat n)) \<longlongrightarrow> 0) sequentially" +proof (clarsimp simp add: lim_sequentially dist_norm norm_divide divide_simps) + fix r::real + assume "0 < r" + have ir: "inverse (exp (inverse r)) > 0" + by simp + obtain n where n: "1 < of_nat n * inverse (exp (inverse r))" + using ex_less_of_nat_mult [of _ 1, OF ir] + by auto + then have "exp (inverse r) < of_nat n" + by (simp add: divide_simps) + then have "ln (exp (inverse r)) < ln (of_nat n)" + by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff) + with \<open>0 < r\<close> have "1 < r * ln (real_of_nat n)" + by (simp add: field_simps) + moreover have "n > 0" using n + using neq0_conv by fastforce + ultimately show "\<exists>no. \<forall>n. Ln (of_nat n) \<noteq> 0 \<longrightarrow> no \<le> n \<longrightarrow> 1 < r * cmod (Ln (of_nat n))" + using n \<open>0 < r\<close> + apply (rule_tac x=n in exI) + apply (auto simp: divide_simps) + apply (erule less_le_trans, auto) + done +qed + +lemma lim_1_over_ln: "((\<lambda>n. 1 / ln(real_of_nat n)) \<longlongrightarrow> 0) sequentially" + using lim_1_over_Ln [THEN filterlim_sequentially_Suc [THEN iffD2]] + apply (subst filterlim_sequentially_Suc [symmetric]) + apply (simp add: lim_sequentially dist_norm) + apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) + done + + +subsection\<open>Relation between Square Root and exp/ln, hence its derivative\<close> + +lemma csqrt_exp_Ln: + assumes "z \<noteq> 0" + shows "csqrt z = exp(Ln(z) / 2)" +proof - + have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))" + by (metis exp_double nonzero_mult_divide_cancel_left times_divide_eq_right zero_neq_numeral) + also have "... = z" + using assms exp_Ln by blast + finally have "csqrt z = csqrt ((exp (Ln z / 2))\<^sup>2)" + by simp + also have "... = exp (Ln z / 2)" + apply (subst csqrt_square) + using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms + apply (auto simp: Re_exp Im_exp zero_less_mult_iff zero_le_mult_iff, fastforce+) + done + finally show ?thesis using assms csqrt_square + by simp +qed + +lemma csqrt_inverse: + assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" + shows "csqrt (inverse z) = inverse (csqrt z)" +proof (cases "z=0", simp) + assume "z \<noteq> 0" + then show ?thesis + using assms csqrt_exp_Ln Ln_inverse exp_minus + by (simp add: csqrt_exp_Ln Ln_inverse exp_minus) +qed + +lemma cnj_csqrt: + assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" + shows "cnj(csqrt z) = csqrt(cnj z)" +proof (cases "z=0", simp) + assume "z \<noteq> 0" + then show ?thesis + by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj) +qed + +lemma has_field_derivative_csqrt: + assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" + shows "(csqrt has_field_derivative inverse(2 * csqrt z)) (at z)" +proof - + have z: "z \<noteq> 0" + using assms by auto + then have *: "inverse z = inverse (2*z) * 2" + by (simp add: divide_simps) + have [simp]: "exp (Ln z / 2) * inverse z = inverse (csqrt z)" + by (simp add: z field_simps csqrt_exp_Ln [symmetric]) (metis power2_csqrt power2_eq_square) + have "Im z = 0 \<Longrightarrow> 0 < Re z" + using assms complex_nonpos_Reals_iff not_less by blast + with z have "((\<lambda>z. exp (Ln z / 2)) has_field_derivative inverse (2 * csqrt z)) (at z)" + by (force intro: derivative_eq_intros * simp add: assms) + then show ?thesis + apply (rule DERIV_transform_at[where d = "norm z"]) + apply (intro z derivative_eq_intros | simp add: assms)+ + using z + apply (metis csqrt_exp_Ln dist_0_norm less_irrefl) + done +qed + +lemma field_differentiable_at_csqrt: + "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> csqrt field_differentiable at z" + using field_differentiable_def has_field_derivative_csqrt by blast + +lemma field_differentiable_within_csqrt: + "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> csqrt field_differentiable (at z within s)" + using field_differentiable_at_csqrt field_differentiable_within_subset by blast + +lemma continuous_at_csqrt: + "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z) csqrt" + by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at) + +corollary isCont_csqrt' [simp]: + "\<lbrakk>isCont f z; f z \<notin> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. csqrt (f x)) z" + by (blast intro: isCont_o2 [OF _ continuous_at_csqrt]) + +lemma continuous_within_csqrt: + "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within s) csqrt" + by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_csqrt) + +lemma continuous_on_csqrt [continuous_intros]: + "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on s csqrt" + by (simp add: continuous_at_imp_continuous_on continuous_within_csqrt) + +lemma holomorphic_on_csqrt: + "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> csqrt holomorphic_on s" + by (simp add: field_differentiable_within_csqrt holomorphic_on_def) + +lemma continuous_within_closed_nontrivial: + "closed s \<Longrightarrow> a \<notin> s ==> continuous (at a within s) f" + using open_Compl + by (force simp add: continuous_def eventually_at_topological filterlim_iff open_Collect_neg) + +lemma continuous_within_csqrt_posreal: + "continuous (at z within (\<real> \<inter> {w. 0 \<le> Re(w)})) csqrt" +proof (cases "z \<in> \<real>\<^sub>\<le>\<^sub>0") + case True + then have "Im z = 0" "Re z < 0 \<or> z = 0" + using cnj.code complex_cnj_zero_iff by (auto simp: complex_nonpos_Reals_iff) fastforce + then show ?thesis + apply (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge]) + apply (auto simp: continuous_within_eps_delta norm_conv_dist [symmetric]) + apply (rule_tac x="e^2" in exI) + apply (auto simp: Reals_def) + by (metis linear not_less real_sqrt_less_iff real_sqrt_pow2_iff real_sqrt_power) +next + case False + then show ?thesis by (blast intro: continuous_within_csqrt) +qed + +subsection\<open>Complex arctangent\<close> + +text\<open>The branch cut gives standard bounds in the real case.\<close> + +definition Arctan :: "complex \<Rightarrow> complex" where + "Arctan \<equiv> \<lambda>z. (\<i>/2) * Ln((1 - \<i>*z) / (1 + \<i>*z))" + +lemma Arctan_def_moebius: "Arctan z = \<i>/2 * Ln (moebius (-\<i>) 1 \<i> 1 z)" + by (simp add: Arctan_def moebius_def add_ac) + +lemma Ln_conv_Arctan: + assumes "z \<noteq> -1" + shows "Ln z = -2*\<i> * Arctan (moebius 1 (- 1) (- \<i>) (- \<i>) z)" +proof - + have "Arctan (moebius 1 (- 1) (- \<i>) (- \<i>) z) = + \<i>/2 * Ln (moebius (- \<i>) 1 \<i> 1 (moebius 1 (- 1) (- \<i>) (- \<i>) z))" + by (simp add: Arctan_def_moebius) + also from assms have "\<i> * z \<noteq> \<i> * (-1)" by (subst mult_left_cancel) simp + hence "\<i> * z - -\<i> \<noteq> 0" by (simp add: eq_neg_iff_add_eq_0) + from moebius_inverse'[OF _ this, of 1 1] + have "moebius (- \<i>) 1 \<i> 1 (moebius 1 (- 1) (- \<i>) (- \<i>) z) = z" by simp + finally show ?thesis by (simp add: field_simps) +qed + +lemma Arctan_0 [simp]: "Arctan 0 = 0" + by (simp add: Arctan_def) + +lemma Im_complex_div_lemma: "Im((1 - \<i>*z) / (1 + \<i>*z)) = 0 \<longleftrightarrow> Re z = 0" + by (auto simp: Im_complex_div_eq_0 algebra_simps) + +lemma Re_complex_div_lemma: "0 < Re((1 - \<i>*z) / (1 + \<i>*z)) \<longleftrightarrow> norm z < 1" + by (simp add: Re_complex_div_gt_0 algebra_simps cmod_def power2_eq_square) + +lemma tan_Arctan: + assumes "z\<^sup>2 \<noteq> -1" + shows [simp]:"tan(Arctan z) = z" +proof - + have "1 + \<i>*z \<noteq> 0" + by (metis assms complex_i_mult_minus i_squared minus_unique power2_eq_square power2_minus) + moreover + have "1 - \<i>*z \<noteq> 0" + by (metis assms complex_i_mult_minus i_squared power2_eq_square power2_minus right_minus_eq) + ultimately + show ?thesis + by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus csqrt_exp_Ln [symmetric] + divide_simps power2_eq_square [symmetric]) +qed + +lemma Arctan_tan [simp]: + assumes "\<bar>Re z\<bar> < pi/2" + shows "Arctan(tan z) = z" +proof - + have ge_pi2: "\<And>n::int. \<bar>of_int (2*n + 1) * pi/2\<bar> \<ge> pi/2" + by (case_tac n rule: int_cases) (auto simp: abs_mult) + have "exp (\<i>*z)*exp (\<i>*z) = -1 \<longleftrightarrow> exp (2*\<i>*z) = -1" + by (metis distrib_right exp_add mult_2) + also have "... \<longleftrightarrow> exp (2*\<i>*z) = exp (\<i>*pi)" + using cis_conv_exp cis_pi by auto + also have "... \<longleftrightarrow> exp (2*\<i>*z - \<i>*pi) = 1" + by (metis (no_types) diff_add_cancel diff_minus_eq_add exp_add exp_minus_inverse mult.commute) + also have "... \<longleftrightarrow> Re(\<i>*2*z - \<i>*pi) = 0 \<and> (\<exists>n::int. Im(\<i>*2*z - \<i>*pi) = of_int (2 * n) * pi)" + by (simp add: exp_eq_1) + also have "... \<longleftrightarrow> Im z = 0 \<and> (\<exists>n::int. 2 * Re z = of_int (2*n + 1) * pi)" + by (simp add: algebra_simps) + also have "... \<longleftrightarrow> False" + using assms ge_pi2 + apply (auto simp: algebra_simps) + by (metis abs_mult_pos not_less of_nat_less_0_iff of_nat_numeral) + finally have *: "exp (\<i>*z)*exp (\<i>*z) + 1 \<noteq> 0" + by (auto simp: add.commute minus_unique) + show ?thesis + using assms * + apply (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps + ii_times_eq_iff power2_eq_square [symmetric]) + apply (rule Ln_unique) + apply (auto simp: divide_simps exp_minus) + apply (simp add: algebra_simps exp_double [symmetric]) + done +qed + +lemma + assumes "Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1" + shows Re_Arctan_bounds: "\<bar>Re(Arctan z)\<bar> < pi/2" + and has_field_derivative_Arctan: "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" +proof - + have nz0: "1 + \<i>*z \<noteq> 0" + using assms + by (metis abs_one complex_i_mult_minus diff_0_right diff_minus_eq_add ii.simps(1) ii.simps(2) + less_irrefl minus_diff_eq mult.right_neutral right_minus_eq) + have "z \<noteq> -\<i>" using assms + by auto + then have zz: "1 + z * z \<noteq> 0" + by (metis abs_one assms i_squared ii.simps less_irrefl minus_unique square_eq_iff) + have nz1: "1 - \<i>*z \<noteq> 0" + using assms by (force simp add: ii_times_eq_iff) + have nz2: "inverse (1 + \<i>*z) \<noteq> 0" + using assms + by (metis Im_complex_div_lemma Re_complex_div_lemma cmod_eq_Im divide_complex_def + less_irrefl mult_zero_right zero_complex.simps(1) zero_complex.simps(2)) + have nzi: "((1 - \<i>*z) * inverse (1 + \<i>*z)) \<noteq> 0" + using nz1 nz2 by auto + have "Im ((1 - \<i>*z) / (1 + \<i>*z)) = 0 \<Longrightarrow> 0 < Re ((1 - \<i>*z) / (1 + \<i>*z))" + apply (simp add: divide_complex_def) + apply (simp add: divide_simps split: if_split_asm) + using assms + apply (auto simp: algebra_simps abs_square_less_1 [unfolded power2_eq_square]) + done + then have *: "((1 - \<i>*z) / (1 + \<i>*z)) \<notin> \<real>\<^sub>\<le>\<^sub>0" + by (auto simp add: complex_nonpos_Reals_iff) + show "\<bar>Re(Arctan z)\<bar> < pi/2" + unfolding Arctan_def divide_complex_def + using mpi_less_Im_Ln [OF nzi] + apply (auto simp: abs_if intro!: Im_Ln_less_pi * [unfolded divide_complex_def]) + done + show "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" + unfolding Arctan_def scaleR_conv_of_real + apply (rule DERIV_cong) + apply (intro derivative_eq_intros | simp add: nz0 *)+ + using nz0 nz1 zz + apply (simp add: divide_simps power2_eq_square) + apply (auto simp: algebra_simps) + done +qed + +lemma field_differentiable_at_Arctan: "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan field_differentiable at z" + using has_field_derivative_Arctan + by (auto simp: field_differentiable_def) + +lemma field_differentiable_within_Arctan: + "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan field_differentiable (at z within s)" + using field_differentiable_at_Arctan field_differentiable_at_within by blast + +declare has_field_derivative_Arctan [derivative_intros] +declare has_field_derivative_Arctan [THEN DERIV_chain2, derivative_intros] + +lemma continuous_at_Arctan: + "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> continuous (at z) Arctan" + by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Arctan) + +lemma continuous_within_Arctan: + "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> continuous (at z within s) Arctan" + using continuous_at_Arctan continuous_at_imp_continuous_within by blast + +lemma continuous_on_Arctan [continuous_intros]: + "(\<And>z. z \<in> s \<Longrightarrow> Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> continuous_on s Arctan" + by (auto simp: continuous_at_imp_continuous_on continuous_within_Arctan) + +lemma holomorphic_on_Arctan: + "(\<And>z. z \<in> s \<Longrightarrow> Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan holomorphic_on s" + by (simp add: field_differentiable_within_Arctan holomorphic_on_def) + +lemma Arctan_series: + assumes z: "norm (z :: complex) < 1" + defines "g \<equiv> \<lambda>n. if odd n then -\<i>*\<i>^n / n else 0" + defines "h \<equiv> \<lambda>z n. (-1)^n / of_nat (2*n+1) * (z::complex)^(2*n+1)" + shows "(\<lambda>n. g n * z^n) sums Arctan z" + and "h z sums Arctan z" +proof - + define G where [abs_def]: "G z = (\<Sum>n. g n * z^n)" for z + have summable: "summable (\<lambda>n. g n * u^n)" if "norm u < 1" for u + proof (cases "u = 0") + assume u: "u \<noteq> 0" + have "(\<lambda>n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\<lambda>n. ereal (inverse (norm u)^2) * + ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))" + proof + fix n + have "ereal (norm (h u n) / norm (h u (Suc n))) = + ereal (inverse (norm u)^2) * ereal ((of_nat (2*Suc n+1) / of_nat (Suc n)) / + (of_nat (2*Suc n-1) / of_nat (Suc n)))" + by (simp add: h_def norm_mult norm_power norm_divide divide_simps + power2_eq_square eval_nat_numeral del: of_nat_add of_nat_Suc) + also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))" + by (auto simp: divide_simps simp del: of_nat_Suc) simp_all? + also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc n))" + by (auto simp: divide_simps simp del: of_nat_Suc) simp_all? + finally show "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * + ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n))))" . + qed + also have "\<dots> \<longlonglongrightarrow> ereal (inverse (norm u)^2) * ereal ((2 + 0) / (2 - 0))" + by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all + finally have "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) = inverse (norm u)^2" + by (intro lim_imp_Liminf) simp_all + moreover from power_strict_mono[OF that, of 2] u have "inverse (norm u)^2 > 1" + by (simp add: divide_simps) + ultimately have A: "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp + from u have "summable (h u)" + by (intro summable_norm_cancel[OF ratio_test_convergence[OF _ A]]) + (auto simp: h_def norm_divide norm_mult norm_power simp del: of_nat_Suc + intro!: mult_pos_pos divide_pos_pos always_eventually) + thus "summable (\<lambda>n. g n * u^n)" + by (subst summable_mono_reindex[of "\<lambda>n. 2*n+1", symmetric]) + (auto simp: power_mult subseq_def g_def h_def elim!: oddE) + qed (simp add: h_def) + + have "\<exists>c. \<forall>u\<in>ball 0 1. Arctan u - G u = c" + proof (rule has_field_derivative_zero_constant) + fix u :: complex assume "u \<in> ball 0 1" + hence u: "norm u < 1" by (simp add: dist_0_norm) + define K where "K = (norm u + 1) / 2" + from u and abs_Im_le_cmod[of u] have Im_u: "\<bar>Im u\<bar> < 1" by linarith + from u have K: "0 \<le> K" "norm u < K" "K < 1" by (simp_all add: K_def) + hence "(G has_field_derivative (\<Sum>n. diffs g n * u ^ n)) (at u)" unfolding G_def + by (intro termdiffs_strong[of _ "of_real K"] summable) simp_all + also have "(\<lambda>n. diffs g n * u^n) = (\<lambda>n. if even n then (\<i>*u)^n else 0)" + by (intro ext) (simp_all del: of_nat_Suc add: g_def diffs_def power_mult_distrib) + also have "suminf \<dots> = (\<Sum>n. (-(u^2))^n)" + by (subst suminf_mono_reindex[of "\<lambda>n. 2*n", symmetric]) + (auto elim!: evenE simp: subseq_def power_mult power_mult_distrib) + also from u have "norm u^2 < 1^2" by (intro power_strict_mono) simp_all + hence "(\<Sum>n. (-(u^2))^n) = inverse (1 + u^2)" + by (subst suminf_geometric) (simp_all add: norm_power inverse_eq_divide) + finally have "(G has_field_derivative inverse (1 + u\<^sup>2)) (at u)" . + from DERIV_diff[OF has_field_derivative_Arctan this] Im_u u + show "((\<lambda>u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)" + by (simp_all add: dist_0_norm at_within_open[OF _ open_ball]) + qed simp_all + then obtain c where c: "\<And>u. norm u < 1 \<Longrightarrow> Arctan u - G u = c" by (auto simp: dist_0_norm) + from this[of 0] have "c = 0" by (simp add: G_def g_def powser_zero) + with c z have "Arctan z = G z" by simp + with summable[OF z] show "(\<lambda>n. g n * z^n) sums Arctan z" unfolding G_def by (simp add: sums_iff) + thus "h z sums Arctan z" by (subst (asm) sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric]) + (auto elim!: oddE simp: subseq_def power_mult g_def h_def) +qed + +text \<open>A quickly-converging series for the logarithm, based on the arctangent.\<close> +lemma ln_series_quadratic: + assumes x: "x > (0::real)" + shows "(\<lambda>n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x" +proof - + define y :: complex where "y = of_real ((x-1)/(x+1))" + from x have x': "complex_of_real x \<noteq> of_real (-1)" by (subst of_real_eq_iff) auto + from x have "\<bar>x - 1\<bar> < \<bar>x + 1\<bar>" by linarith + hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1" + by (simp add: norm_divide del: of_real_add of_real_diff) + hence "norm (\<i> * y) < 1" unfolding y_def by (subst norm_mult) simp + hence "(\<lambda>n. (-2*\<i>) * ((-1)^n / of_nat (2*n+1) * (\<i>*y)^(2*n+1))) sums ((-2*\<i>) * Arctan (\<i>*y))" + by (intro Arctan_series sums_mult) simp_all + also have "(\<lambda>n. (-2*\<i>) * ((-1)^n / of_nat (2*n+1) * (\<i>*y)^(2*n+1))) = + (\<lambda>n. (-2*\<i>) * ((-1)^n * (\<i>*y*(-y\<^sup>2)^n)/of_nat (2*n+1)))" + by (intro ext) (simp_all add: power_mult power_mult_distrib) + also have "\<dots> = (\<lambda>n. 2*y* ((-1) * (-y\<^sup>2))^n/of_nat (2*n+1))" + by (intro ext, subst power_mult_distrib) (simp add: algebra_simps power_mult) + also have "\<dots> = (\<lambda>n. 2*y^(2*n+1) / of_nat (2*n+1))" + by (subst power_add, subst power_mult) (simp add: mult_ac) + also have "\<dots> = (\<lambda>n. of_real (2*((x-1)/(x+1))^(2*n+1) / of_nat (2*n+1)))" + by (intro ext) (simp add: y_def) + also have "\<i> * y = (of_real x - 1) / (-\<i> * (of_real x + 1))" + by (subst divide_divide_eq_left [symmetric]) (simp add: y_def) + also have "\<dots> = moebius 1 (-1) (-\<i>) (-\<i>) (of_real x)" by (simp add: moebius_def algebra_simps) + also from x' have "-2*\<i>*Arctan \<dots> = Ln (of_real x)" by (intro Ln_conv_Arctan [symmetric]) simp_all + also from x have "\<dots> = ln x" by (rule Ln_of_real) + finally show ?thesis by (subst (asm) sums_of_real_iff) +qed + +subsection \<open>Real arctangent\<close> + +lemma norm_exp_ii_times [simp]: "norm (exp(\<i> * of_real y)) = 1" + by simp + +lemma norm_exp_imaginary: "norm(exp z) = 1 \<Longrightarrow> Re z = 0" + by (simp add: complex_norm_eq_1_exp) + +lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0" + unfolding Arctan_def divide_complex_def + apply (simp add: complex_eq_iff) + apply (rule norm_exp_imaginary) + apply (subst exp_Ln, auto) + apply (simp_all add: cmod_def complex_eq_iff) + apply (auto simp: divide_simps) + apply (metis power_one sum_power2_eq_zero_iff zero_neq_one, algebra) + done + +lemma arctan_eq_Re_Arctan: "arctan x = Re (Arctan (of_real x))" +proof (rule arctan_unique) + show "- (pi / 2) < Re (Arctan (complex_of_real x))" + apply (simp add: Arctan_def) + apply (rule Im_Ln_less_pi) + apply (auto simp: Im_complex_div_lemma complex_nonpos_Reals_iff) + done +next + have *: " (1 - \<i>*x) / (1 + \<i>*x) \<noteq> 0" + by (simp add: divide_simps) ( simp add: complex_eq_iff) + show "Re (Arctan (complex_of_real x)) < pi / 2" + using mpi_less_Im_Ln [OF *] + by (simp add: Arctan_def) +next + have "tan (Re (Arctan (of_real x))) = Re (tan (Arctan (of_real x)))" + apply (auto simp: tan_def Complex.Re_divide Re_sin Re_cos Im_sin Im_cos) + apply (simp add: field_simps) + by (simp add: power2_eq_square) + also have "... = x" + apply (subst tan_Arctan, auto) + by (metis diff_0_right minus_diff_eq mult_zero_left not_le of_real_1 of_real_eq_iff of_real_minus of_real_power power2_eq_square real_minus_mult_self_le zero_less_one) + finally show "tan (Re (Arctan (complex_of_real x))) = x" . +qed + +lemma Arctan_of_real: "Arctan (of_real x) = of_real (arctan x)" + unfolding arctan_eq_Re_Arctan divide_complex_def + by (simp add: complex_eq_iff) + +lemma Arctan_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Arctan z \<in> \<real>" + by (metis Reals_cases Reals_of_real Arctan_of_real) + +declare arctan_one [simp] + +lemma arctan_less_pi4_pos: "x < 1 \<Longrightarrow> arctan x < pi/4" + by (metis arctan_less_iff arctan_one) + +lemma arctan_less_pi4_neg: "-1 < x \<Longrightarrow> -(pi/4) < arctan x" + by (metis arctan_less_iff arctan_minus arctan_one) + +lemma arctan_less_pi4: "\<bar>x\<bar> < 1 \<Longrightarrow> \<bar>arctan x\<bar> < pi/4" + by (metis abs_less_iff arctan_less_pi4_pos arctan_minus) + +lemma arctan_le_pi4: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> \<bar>arctan x\<bar> \<le> pi/4" + by (metis abs_le_iff arctan_le_iff arctan_minus arctan_one) + +lemma abs_arctan: "\<bar>arctan x\<bar> = arctan \<bar>x\<bar>" + by (simp add: abs_if arctan_minus) + +lemma arctan_add_raw: + assumes "\<bar>arctan x + arctan y\<bar> < pi/2" + shows "arctan x + arctan y = arctan((x + y) / (1 - x * y))" +proof (rule arctan_unique [symmetric]) + show 12: "- (pi / 2) < arctan x + arctan y" "arctan x + arctan y < pi / 2" + using assms by linarith+ + show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" + using cos_gt_zero_pi [OF 12] + by (simp add: arctan tan_add) +qed + +lemma arctan_inverse: + assumes "0 < x" + shows "arctan(inverse x) = pi/2 - arctan x" +proof - + have "arctan(inverse x) = arctan(inverse(tan(arctan x)))" + by (simp add: arctan) + also have "... = arctan (tan (pi / 2 - arctan x))" + by (simp add: tan_cot) + also have "... = pi/2 - arctan x" + proof - + have "0 < pi - arctan x" + using arctan_ubound [of x] pi_gt_zero by linarith + with assms show ?thesis + by (simp add: Transcendental.arctan_tan) + qed + finally show ?thesis . +qed + +lemma arctan_add_small: + assumes "\<bar>x * y\<bar> < 1" + shows "(arctan x + arctan y = arctan((x + y) / (1 - x * y)))" +proof (cases "x = 0 \<or> y = 0") + case True then show ?thesis + by auto +next + case False + then have *: "\<bar>arctan x\<bar> < pi / 2 - \<bar>arctan y\<bar>" using assms + apply (auto simp add: abs_arctan arctan_inverse [symmetric] arctan_less_iff) + apply (simp add: divide_simps abs_mult) + done + show ?thesis + apply (rule arctan_add_raw) + using * by linarith +qed + +lemma abs_arctan_le: + fixes x::real shows "\<bar>arctan x\<bar> \<le> \<bar>x\<bar>" +proof - + { fix w::complex and z::complex + assume *: "w \<in> \<real>" "z \<in> \<real>" + have "cmod (Arctan w - Arctan z) \<le> 1 * cmod (w-z)" + apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1]) + apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan]) + apply (force simp add: Reals_def) + apply (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square) + using * by auto + } + then have "cmod (Arctan (of_real x) - Arctan 0) \<le> 1 * cmod (of_real x -0)" + using Reals_0 Reals_of_real by blast + then show ?thesis + by (simp add: Arctan_of_real) +qed + +lemma arctan_le_self: "0 \<le> x \<Longrightarrow> arctan x \<le> x" + by (metis abs_arctan_le abs_of_nonneg zero_le_arctan_iff) + +lemma abs_tan_ge: "\<bar>x\<bar> < pi/2 \<Longrightarrow> \<bar>x\<bar> \<le> \<bar>tan x\<bar>" + by (metis abs_arctan_le abs_less_iff arctan_tan minus_less_iff) + +lemma arctan_bounds: + assumes "0 \<le> x" "x < 1" + shows arctan_lower_bound: + "(\<Sum>k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \<le> arctan x" + (is "(\<Sum>k<_. (- 1)^ k * ?a k) \<le> _") + and arctan_upper_bound: + "arctan x \<le> (\<Sum>k<2 * n + 1. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" +proof - + have tendsto_zero: "?a \<longlonglongrightarrow> 0" + using assms + apply - + apply (rule tendsto_eq_rhs[where x="0 * 0"]) + subgoal by (intro tendsto_mult real_tendsto_divide_at_top) + (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real + intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially + tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top) + subgoal by simp + done + have nonneg: "0 \<le> ?a n" for n + by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms) + have le: "?a (Suc n) \<le> ?a n" for n + by (rule mult_mono[OF _ power_decreasing]) (auto simp: divide_simps assms less_imp_le) + from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n] + summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n] + assms + show "(\<Sum>k<2*n. (- 1)^ k * ?a k) \<le> arctan x" "arctan x \<le> (\<Sum>k<2 * n + 1. (- 1)^ k * ?a k)" + by (auto simp: arctan_series) +qed + +subsection \<open>Bounds on pi using real arctangent\<close> + +lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)" + using machin + by simp + +lemma pi_approx: "3.141592653588 \<le> pi" "pi \<le> 3.1415926535899" + unfolding pi_machin + using arctan_bounds[of "1/5" 4] + arctan_bounds[of "1/239" 4] + by (simp_all add: eval_nat_numeral) + + +subsection\<open>Inverse Sine\<close> + +definition Arcsin :: "complex \<Rightarrow> complex" where + "Arcsin \<equiv> \<lambda>z. -\<i> * Ln(\<i> * z + csqrt(1 - z\<^sup>2))" + +lemma Arcsin_body_lemma: "\<i> * z + csqrt(1 - z\<^sup>2) \<noteq> 0" + using power2_csqrt [of "1 - z\<^sup>2"] + apply auto + by (metis complex_i_mult_minus diff_add_cancel diff_minus_eq_add diff_self mult.assoc mult.left_commute numeral_One power2_csqrt power2_eq_square zero_neq_numeral) + +lemma Arcsin_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Re(\<i> * z + csqrt(1 - z\<^sup>2))" + using Complex.cmod_power2 [of z, symmetric] + by (simp add: real_less_rsqrt algebra_simps Re_power2 cmod_square_less_1_plus) + +lemma Re_Arcsin: "Re(Arcsin z) = Im (Ln (\<i> * z + csqrt(1 - z\<^sup>2)))" + by (simp add: Arcsin_def) + +lemma Im_Arcsin: "Im(Arcsin z) = - ln (cmod (\<i> * z + csqrt (1 - z\<^sup>2)))" + by (simp add: Arcsin_def Arcsin_body_lemma) + +lemma one_minus_z2_notin_nonpos_Reals: + assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)" + shows "1 - z\<^sup>2 \<notin> \<real>\<^sub>\<le>\<^sub>0" + using assms + apply (auto simp: complex_nonpos_Reals_iff Re_power2 Im_power2) + using power2_less_0 [of "Im z"] apply force + using abs_square_less_1 not_le by blast + +lemma isCont_Arcsin_lemma: + assumes le0: "Re (\<i> * z + csqrt (1 - z\<^sup>2)) \<le> 0" and "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)" + shows False +proof (cases "Im z = 0") + case True + then show ?thesis + using assms by (fastforce simp: cmod_def abs_square_less_1 [symmetric]) +next + case False + have neq: "(cmod z)\<^sup>2 \<noteq> 1 + cmod (1 - z\<^sup>2)" + proof (clarsimp simp add: cmod_def) + assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 = 1 + sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" + then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" + by simp + then show False using False + by (simp add: power2_eq_square algebra_simps) + qed + moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2" + using le0 + apply simp + apply (drule sqrt_le_D) + using cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1] + apply (simp add: norm_power Re_power2 norm_minus_commute [of 1]) + done + ultimately show False + by (simp add: Re_power2 Im_power2 cmod_power2) +qed + +lemma isCont_Arcsin: + assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)" + shows "isCont Arcsin z" +proof - + have *: "\<i> * z + csqrt (1 - z\<^sup>2) \<notin> \<real>\<^sub>\<le>\<^sub>0" + by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff) + show ?thesis + using assms + apply (simp add: Arcsin_def) + apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+ + apply (simp add: one_minus_z2_notin_nonpos_Reals assms) + apply (rule *) + done +qed + +lemma isCont_Arcsin' [simp]: + shows "isCont f z \<Longrightarrow> (Im (f z) = 0 \<Longrightarrow> \<bar>Re (f z)\<bar> < 1) \<Longrightarrow> isCont (\<lambda>x. Arcsin (f x)) z" + by (blast intro: isCont_o2 [OF _ isCont_Arcsin]) + +lemma sin_Arcsin [simp]: "sin(Arcsin z) = z" +proof - + have "\<i>*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \<longleftrightarrow> (\<i>*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0" + by (simp add: algebra_simps) \<comment>\<open>Cancelling a factor of 2\<close> + moreover have "... \<longleftrightarrow> (\<i>*z) + csqrt (1 - z\<^sup>2) = 0" + by (metis Arcsin_body_lemma distrib_right no_zero_divisors zero_neq_numeral) + ultimately show ?thesis + apply (simp add: sin_exp_eq Arcsin_def Arcsin_body_lemma exp_minus divide_simps) + apply (simp add: algebra_simps) + apply (simp add: power2_eq_square [symmetric] algebra_simps) + done +qed + +lemma Re_eq_pihalf_lemma: + "\<bar>Re z\<bar> = pi/2 \<Longrightarrow> Im z = 0 \<Longrightarrow> + Re ((exp (\<i>*z) + inverse (exp (\<i>*z))) / 2) = 0 \<and> 0 \<le> Im ((exp (\<i>*z) + inverse (exp (\<i>*z))) / 2)" + apply (simp add: cos_ii_times [symmetric] Re_cos Im_cos abs_if del: eq_divide_eq_numeral1) + by (metis cos_minus cos_pi_half) + +lemma Re_less_pihalf_lemma: + assumes "\<bar>Re z\<bar> < pi / 2" + shows "0 < Re ((exp (\<i>*z) + inverse (exp (\<i>*z))) / 2)" +proof - + have "0 < cos (Re z)" using assms + using cos_gt_zero_pi by auto + then show ?thesis + by (simp add: cos_ii_times [symmetric] Re_cos Im_cos add_pos_pos) +qed + +lemma Arcsin_sin: + assumes "\<bar>Re z\<bar> < pi/2 \<or> (\<bar>Re z\<bar> = pi/2 \<and> Im z = 0)" + shows "Arcsin(sin z) = z" +proof - + have "Arcsin(sin z) = - (\<i> * Ln (csqrt (1 - (\<i> * (exp (\<i>*z) - inverse (exp (\<i>*z))))\<^sup>2 / 4) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))" + by (simp add: sin_exp_eq Arcsin_def exp_minus power_divide) + also have "... = - (\<i> * Ln (csqrt (((exp (\<i>*z) + inverse (exp (\<i>*z)))/2)\<^sup>2) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))" + by (simp add: field_simps power2_eq_square) + also have "... = - (\<i> * Ln (((exp (\<i>*z) + inverse (exp (\<i>*z)))/2) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))" + apply (subst csqrt_square) + using assms Re_eq_pihalf_lemma Re_less_pihalf_lemma + apply auto + done + also have "... = - (\<i> * Ln (exp (\<i>*z)))" + by (simp add: field_simps power2_eq_square) + also have "... = z" + apply (subst Complex_Transcendental.Ln_exp) + using assms + apply (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm) + done + finally show ?thesis . +qed + +lemma Arcsin_unique: + "\<lbrakk>sin z = w; \<bar>Re z\<bar> < pi/2 \<or> (\<bar>Re z\<bar> = pi/2 \<and> Im z = 0)\<rbrakk> \<Longrightarrow> Arcsin w = z" + by (metis Arcsin_sin) + +lemma Arcsin_0 [simp]: "Arcsin 0 = 0" + by (metis Arcsin_sin norm_zero pi_half_gt_zero real_norm_def sin_zero zero_complex.simps(1)) + +lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2" + by (metis Arcsin_sin Im_complex_of_real Re_complex_of_real numeral_One of_real_numeral pi_half_ge_zero real_sqrt_abs real_sqrt_pow2 real_sqrt_power sin_of_real sin_pi_half) + +lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)" + by (metis Arcsin_1 Arcsin_sin Im_complex_of_real Re_complex_of_real abs_of_nonneg of_real_minus pi_half_ge_zero power2_minus real_sqrt_abs sin_Arcsin sin_minus) + +lemma has_field_derivative_Arcsin: + assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)" + shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)" +proof - + have "(sin (Arcsin z))\<^sup>2 \<noteq> 1" + using assms + apply atomize + apply (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1) + apply (metis abs_minus_cancel abs_one abs_power2 numeral_One numeral_neq_neg_one) + by (metis abs_minus_cancel abs_one abs_power2 one_neq_neg_one) + then have "cos (Arcsin z) \<noteq> 0" + by (metis diff_0_right power_zero_numeral sin_squared_eq) + then show ?thesis + apply (rule has_complex_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) + apply (auto intro: isCont_Arcsin assms) + done +qed + +declare has_field_derivative_Arcsin [derivative_intros] +declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] + +lemma field_differentiable_at_Arcsin: + "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin field_differentiable at z" + using field_differentiable_def has_field_derivative_Arcsin by blast + +lemma field_differentiable_within_Arcsin: + "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin field_differentiable (at z within s)" + using field_differentiable_at_Arcsin field_differentiable_within_subset by blast + +lemma continuous_within_Arcsin: + "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous (at z within s) Arcsin" + using continuous_at_imp_continuous_within isCont_Arcsin by blast + +lemma continuous_on_Arcsin [continuous_intros]: + "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous_on s Arcsin" + by (simp add: continuous_at_imp_continuous_on) + +lemma holomorphic_on_Arcsin: "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin holomorphic_on s" + by (simp add: field_differentiable_within_Arcsin holomorphic_on_def) + + +subsection\<open>Inverse Cosine\<close> + +definition Arccos :: "complex \<Rightarrow> complex" where + "Arccos \<equiv> \<lambda>z. -\<i> * Ln(z + \<i> * csqrt(1 - z\<^sup>2))" + +lemma Arccos_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Im(z + \<i> * csqrt(1 - z\<^sup>2))" + using Arcsin_range_lemma [of "-z"] + by simp + +lemma Arccos_body_lemma: "z + \<i> * csqrt(1 - z\<^sup>2) \<noteq> 0" + using Arcsin_body_lemma [of z] + by (metis complex_i_mult_minus diff_add_cancel minus_diff_eq minus_unique mult.assoc mult.left_commute + power2_csqrt power2_eq_square zero_neq_one) + +lemma Re_Arccos: "Re(Arccos z) = Im (Ln (z + \<i> * csqrt(1 - z\<^sup>2)))" + by (simp add: Arccos_def) + +lemma Im_Arccos: "Im(Arccos z) = - ln (cmod (z + \<i> * csqrt (1 - z\<^sup>2)))" + by (simp add: Arccos_def Arccos_body_lemma) + +text\<open>A very tricky argument to find!\<close> +lemma isCont_Arccos_lemma: + assumes eq0: "Im (z + \<i> * csqrt (1 - z\<^sup>2)) = 0" and "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)" + shows False +proof (cases "Im z = 0") + case True + then show ?thesis + using assms by (fastforce simp add: cmod_def abs_square_less_1 [symmetric]) +next + case False + have Imz: "Im z = - sqrt ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" + using eq0 abs_Re_le_cmod [of "1-z\<^sup>2"] + by (simp add: Re_power2 algebra_simps) + have "(cmod z)\<^sup>2 - 1 \<noteq> cmod (1 - z\<^sup>2)" + proof (clarsimp simp add: cmod_def) + assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 - 1 = sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" + then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" + by simp + then show False using False + by (simp add: power2_eq_square algebra_simps) + qed + moreover have "(Im z)\<^sup>2 = ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" + apply (subst Imz) + using abs_Re_le_cmod [of "1-z\<^sup>2"] + apply (simp add: Re_power2) + done + ultimately show False + by (simp add: cmod_power2) +qed + +lemma isCont_Arccos: + assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)" + shows "isCont Arccos z" +proof - + have "z + \<i> * csqrt (1 - z\<^sup>2) \<notin> \<real>\<^sub>\<le>\<^sub>0" + by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms) + with assms show ?thesis + apply (simp add: Arccos_def) + apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+ + apply (simp_all add: one_minus_z2_notin_nonpos_Reals assms) + done +qed + +lemma isCont_Arccos' [simp]: + shows "isCont f z \<Longrightarrow> (Im (f z) = 0 \<Longrightarrow> \<bar>Re (f z)\<bar> < 1) \<Longrightarrow> isCont (\<lambda>x. Arccos (f x)) z" + by (blast intro: isCont_o2 [OF _ isCont_Arccos]) + +lemma cos_Arccos [simp]: "cos(Arccos z) = z" +proof - + have "z*2 + \<i> * (2 * csqrt (1 - z\<^sup>2)) = 0 \<longleftrightarrow> z*2 + \<i> * csqrt (1 - z\<^sup>2)*2 = 0" + by (simp add: algebra_simps) \<comment>\<open>Cancelling a factor of 2\<close> + moreover have "... \<longleftrightarrow> z + \<i> * csqrt (1 - z\<^sup>2) = 0" + by (metis distrib_right mult_eq_0_iff zero_neq_numeral) + ultimately show ?thesis + apply (simp add: cos_exp_eq Arccos_def Arccos_body_lemma exp_minus field_simps) + apply (simp add: power2_eq_square [symmetric]) + done +qed + +lemma Arccos_cos: + assumes "0 < Re z & Re z < pi \<or> + Re z = 0 & 0 \<le> Im z \<or> + Re z = pi & Im z \<le> 0" + shows "Arccos(cos z) = z" +proof - + have *: "((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z))) = sin z" + by (simp add: sin_exp_eq exp_minus field_simps power2_eq_square) + have "1 - (exp (\<i> * z) + inverse (exp (\<i> * z)))\<^sup>2 / 4 = ((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z)))\<^sup>2" + by (simp add: field_simps power2_eq_square) + then have "Arccos(cos z) = - (\<i> * Ln ((exp (\<i> * z) + inverse (exp (\<i> * z))) / 2 + + \<i> * csqrt (((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z)))\<^sup>2)))" + by (simp add: cos_exp_eq Arccos_def exp_minus power_divide) + also have "... = - (\<i> * Ln ((exp (\<i> * z) + inverse (exp (\<i> * z))) / 2 + + \<i> * ((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z)))))" + apply (subst csqrt_square) + using assms Re_sin_pos [of z] Im_sin_nonneg [of z] Im_sin_nonneg2 [of z] + apply (auto simp: * Re_sin Im_sin) + done + also have "... = - (\<i> * Ln (exp (\<i>*z)))" + by (simp add: field_simps power2_eq_square) + also have "... = z" + using assms + apply (subst Complex_Transcendental.Ln_exp, auto) + done + finally show ?thesis . +qed + +lemma Arccos_unique: + "\<lbrakk>cos z = w; + 0 < Re z \<and> Re z < pi \<or> + Re z = 0 \<and> 0 \<le> Im z \<or> + Re z = pi \<and> Im z \<le> 0\<rbrakk> \<Longrightarrow> Arccos w = z" + using Arccos_cos by blast + +lemma Arccos_0 [simp]: "Arccos 0 = pi/2" + by (rule Arccos_unique) (auto simp: of_real_numeral) + +lemma Arccos_1 [simp]: "Arccos 1 = 0" + by (rule Arccos_unique) auto + +lemma Arccos_minus1: "Arccos(-1) = pi" + by (rule Arccos_unique) auto + +lemma has_field_derivative_Arccos: + assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)" + shows "(Arccos has_field_derivative - inverse(sin(Arccos z))) (at z)" +proof - + have "(cos (Arccos z))\<^sup>2 \<noteq> 1" + using assms + apply atomize + apply (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1) + apply (metis abs_minus_cancel abs_one abs_power2 numeral_One numeral_neq_neg_one) + apply (metis left_minus less_irrefl power_one sum_power2_gt_zero_iff zero_neq_neg_one) + done + then have "- sin (Arccos z) \<noteq> 0" + by (metis cos_squared_eq diff_0_right mult_zero_left neg_0_equal_iff_equal power2_eq_square) + then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)" + apply (rule has_complex_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]]) + apply (auto intro: isCont_Arccos assms) + done + then show ?thesis + by simp +qed + +declare has_field_derivative_Arcsin [derivative_intros] +declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] + +lemma field_differentiable_at_Arccos: + "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos field_differentiable at z" + using field_differentiable_def has_field_derivative_Arccos by blast + +lemma field_differentiable_within_Arccos: + "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos field_differentiable (at z within s)" + using field_differentiable_at_Arccos field_differentiable_within_subset by blast + +lemma continuous_within_Arccos: + "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous (at z within s) Arccos" + using continuous_at_imp_continuous_within isCont_Arccos by blast + +lemma continuous_on_Arccos [continuous_intros]: + "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous_on s Arccos" + by (simp add: continuous_at_imp_continuous_on) + +lemma holomorphic_on_Arccos: "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos holomorphic_on s" + by (simp add: field_differentiable_within_Arccos holomorphic_on_def) + + +subsection\<open>Upper and Lower Bounds for Inverse Sine and Cosine\<close> + +lemma Arcsin_bounds: "\<bar>Re z\<bar> < 1 \<Longrightarrow> \<bar>Re(Arcsin z)\<bar> < pi/2" + unfolding Re_Arcsin + by (blast intro: Re_Ln_pos_lt_imp Arcsin_range_lemma) + +lemma Arccos_bounds: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Re(Arccos z) \<and> Re(Arccos z) < pi" + unfolding Re_Arccos + by (blast intro!: Im_Ln_pos_lt_imp Arccos_range_lemma) + +lemma Re_Arccos_bounds: "-pi < Re(Arccos z) \<and> Re(Arccos z) \<le> pi" + unfolding Re_Arccos + by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arccos_body_lemma) + +lemma Re_Arccos_bound: "\<bar>Re(Arccos z)\<bar> \<le> pi" + by (meson Re_Arccos_bounds abs_le_iff less_eq_real_def minus_less_iff) + +lemma Re_Arcsin_bounds: "-pi < Re(Arcsin z) & Re(Arcsin z) \<le> pi" + unfolding Re_Arcsin + by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arcsin_body_lemma) + +lemma Re_Arcsin_bound: "\<bar>Re(Arcsin z)\<bar> \<le> pi" + by (meson Re_Arcsin_bounds abs_le_iff less_eq_real_def minus_less_iff) + + +subsection\<open>Interrelations between Arcsin and Arccos\<close> + +lemma cos_Arcsin_nonzero: + assumes "z\<^sup>2 \<noteq> 1" shows "cos(Arcsin z) \<noteq> 0" +proof - + have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)" + by (simp add: power_mult_distrib algebra_simps) + have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> z\<^sup>2 - 1" + proof + assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1" + then have "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (z\<^sup>2 - 1)\<^sup>2" + by simp + then have "z\<^sup>2 * (z\<^sup>2 - 1) = (z\<^sup>2 - 1)*(z\<^sup>2 - 1)" + using eq power2_eq_square by auto + then show False + using assms by simp + qed + then have "1 + \<i> * z * (csqrt (1 - z * z)) \<noteq> z\<^sup>2" + by (metis add_minus_cancel power2_eq_square uminus_add_conv_diff) + then have "2*(1 + \<i> * z * (csqrt (1 - z * z))) \<noteq> 2*z\<^sup>2" (*FIXME cancel_numeral_factor*) + by (metis mult_cancel_left zero_neq_numeral) + then have "(\<i> * z + csqrt (1 - z\<^sup>2))\<^sup>2 \<noteq> -1" + using assms + apply (auto simp: power2_sum) + apply (simp add: power2_eq_square algebra_simps) + done + then show ?thesis + apply (simp add: cos_exp_eq Arcsin_def exp_minus) + apply (simp add: divide_simps Arcsin_body_lemma) + apply (metis add.commute minus_unique power2_eq_square) + done +qed + +lemma sin_Arccos_nonzero: + assumes "z\<^sup>2 \<noteq> 1" shows "sin(Arccos z) \<noteq> 0" +proof - + have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)" + by (simp add: power_mult_distrib algebra_simps) + have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> 1 - z\<^sup>2" + proof + assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2" + then have "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (1 - z\<^sup>2)\<^sup>2" + by simp + then have "-(z\<^sup>2) * (1 - z\<^sup>2) = (1 - z\<^sup>2)*(1 - z\<^sup>2)" + using eq power2_eq_square by auto + then have "-(z\<^sup>2) = (1 - z\<^sup>2)" + using assms + by (metis add.commute add.right_neutral diff_add_cancel mult_right_cancel) + then show False + using assms by simp + qed + then have "z\<^sup>2 + \<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> 1" + by (simp add: algebra_simps) + then have "2*(z\<^sup>2 + \<i> * z * (csqrt (1 - z\<^sup>2))) \<noteq> 2*1" + by (metis mult_cancel_left2 zero_neq_numeral) (*FIXME cancel_numeral_factor*) + then have "(z + \<i> * csqrt (1 - z\<^sup>2))\<^sup>2 \<noteq> 1" + using assms + apply (auto simp: Power.comm_semiring_1_class.power2_sum power_mult_distrib) + apply (simp add: power2_eq_square algebra_simps) + done + then show ?thesis + apply (simp add: sin_exp_eq Arccos_def exp_minus) + apply (simp add: divide_simps Arccos_body_lemma) + apply (simp add: power2_eq_square) + done +qed + +lemma cos_sin_csqrt: + assumes "0 < cos(Re z) \<or> cos(Re z) = 0 \<and> Im z * sin(Re z) \<le> 0" + shows "cos z = csqrt(1 - (sin z)\<^sup>2)" + apply (rule csqrt_unique [THEN sym]) + apply (simp add: cos_squared_eq) + using assms + apply (auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff) + done + +lemma sin_cos_csqrt: + assumes "0 < sin(Re z) \<or> sin(Re z) = 0 \<and> 0 \<le> Im z * cos(Re z)" + shows "sin z = csqrt(1 - (cos z)\<^sup>2)" + apply (rule csqrt_unique [THEN sym]) + apply (simp add: sin_squared_eq) + using assms + apply (auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff) + done + +lemma Arcsin_Arccos_csqrt_pos: + "(0 < Re z | Re z = 0 & 0 \<le> Im z) \<Longrightarrow> Arcsin z = Arccos(csqrt(1 - z\<^sup>2))" + by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) + +lemma Arccos_Arcsin_csqrt_pos: + "(0 < Re z | Re z = 0 & 0 \<le> Im z) \<Longrightarrow> Arccos z = Arcsin(csqrt(1 - z\<^sup>2))" + by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) + +lemma sin_Arccos: + "0 < Re z | Re z = 0 & 0 \<le> Im z \<Longrightarrow> sin(Arccos z) = csqrt(1 - z\<^sup>2)" + by (simp add: Arccos_Arcsin_csqrt_pos) + +lemma cos_Arcsin: + "0 < Re z | Re z = 0 & 0 \<le> Im z \<Longrightarrow> cos(Arcsin z) = csqrt(1 - z\<^sup>2)" + by (simp add: Arcsin_Arccos_csqrt_pos) + + +subsection\<open>Relationship with Arcsin on the Real Numbers\<close> + +lemma Im_Arcsin_of_real: + assumes "\<bar>x\<bar> \<le> 1" + shows "Im (Arcsin (of_real x)) = 0" +proof - + have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))" + by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) + then have "cmod (\<i> * of_real x + csqrt (1 - (of_real x)\<^sup>2))^2 = 1" + using assms abs_square_le_1 + by (force simp add: Complex.cmod_power2) + then have "cmod (\<i> * of_real x + csqrt (1 - (of_real x)\<^sup>2)) = 1" + by (simp add: norm_complex_def) + then show ?thesis + by (simp add: Im_Arcsin exp_minus) +qed + +corollary Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>" + by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) + +lemma arcsin_eq_Re_Arcsin: + assumes "\<bar>x\<bar> \<le> 1" + shows "arcsin x = Re (Arcsin (of_real x))" +unfolding arcsin_def +proof (rule the_equality, safe) + show "- (pi / 2) \<le> Re (Arcsin (complex_of_real x))" + using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] + by (auto simp: Complex.in_Reals_norm Re_Arcsin) +next + show "Re (Arcsin (complex_of_real x)) \<le> pi / 2" + using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] + by (auto simp: Complex.in_Reals_norm Re_Arcsin) +next + show "sin (Re (Arcsin (complex_of_real x))) = x" + using Re_sin [of "Arcsin (of_real x)"] Arcsin_body_lemma [of "of_real x"] + by (simp add: Im_Arcsin_of_real assms) +next + fix x' + assume "- (pi / 2) \<le> x'" "x' \<le> pi / 2" "x = sin x'" + then show "x' = Re (Arcsin (complex_of_real (sin x')))" + apply (simp add: sin_of_real [symmetric]) + apply (subst Arcsin_sin) + apply (auto simp: ) + done +qed + +lemma of_real_arcsin: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arcsin x) = Arcsin(of_real x)" + by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0) + + +subsection\<open>Relationship with Arccos on the Real Numbers\<close> + +lemma Im_Arccos_of_real: + assumes "\<bar>x\<bar> \<le> 1" + shows "Im (Arccos (of_real x)) = 0" +proof - + have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))" + by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) + then have "cmod (of_real x + \<i> * csqrt (1 - (of_real x)\<^sup>2))^2 = 1" + using assms abs_square_le_1 + by (force simp add: Complex.cmod_power2) + then have "cmod (of_real x + \<i> * csqrt (1 - (of_real x)\<^sup>2)) = 1" + by (simp add: norm_complex_def) + then show ?thesis + by (simp add: Im_Arccos exp_minus) +qed + +corollary Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>" + by (metis Im_Arccos_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) + +lemma arccos_eq_Re_Arccos: + assumes "\<bar>x\<bar> \<le> 1" + shows "arccos x = Re (Arccos (of_real x))" +unfolding arccos_def +proof (rule the_equality, safe) + show "0 \<le> Re (Arccos (complex_of_real x))" + using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] + by (auto simp: Complex.in_Reals_norm Re_Arccos) +next + show "Re (Arccos (complex_of_real x)) \<le> pi" + using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] + by (auto simp: Complex.in_Reals_norm Re_Arccos) +next + show "cos (Re (Arccos (complex_of_real x))) = x" + using Re_cos [of "Arccos (of_real x)"] Arccos_body_lemma [of "of_real x"] + by (simp add: Im_Arccos_of_real assms) +next + fix x' + assume "0 \<le> x'" "x' \<le> pi" "x = cos x'" + then show "x' = Re (Arccos (complex_of_real (cos x')))" + apply (simp add: cos_of_real [symmetric]) + apply (subst Arccos_cos) + apply (auto simp: ) + done +qed + +lemma of_real_arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arccos x) = Arccos(of_real x)" + by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0) + +subsection\<open>Some interrelationships among the real inverse trig functions.\<close> + +lemma arccos_arctan: + assumes "-1 < x" "x < 1" + shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))" +proof - + have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0" + proof (rule sin_eq_0_pi) + show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)" + using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms + by (simp add: algebra_simps) + next + show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi" + using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms + by (simp add: algebra_simps) + next + show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0" + using assms + by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan + power2_eq_square square_eq_1_iff) + qed + then show ?thesis + by simp +qed + +lemma arcsin_plus_arccos: + assumes "-1 \<le> x" "x \<le> 1" + shows "arcsin x + arccos x = pi/2" +proof - + have "arcsin x = pi/2 - arccos x" + apply (rule sin_inj_pi) + using assms arcsin [OF assms] arccos [OF assms] + apply (auto simp: algebra_simps sin_diff) + done + then show ?thesis + by (simp add: algebra_simps) +qed + +lemma arcsin_arccos_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = pi/2 - arccos x" + using arcsin_plus_arccos by force + +lemma arccos_arcsin_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = pi/2 - arcsin x" + using arcsin_plus_arccos by force + +lemma arcsin_arctan: "-1 < x \<Longrightarrow> x < 1 \<Longrightarrow> arcsin x = arctan(x / sqrt(1 - x\<^sup>2))" + by (simp add: arccos_arctan arcsin_arccos_eq) + +lemma csqrt_1_diff_eq: "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))" + by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg) + +lemma arcsin_arccos_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = arccos(sqrt(1 - x\<^sup>2))" + apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) + apply (subst Arcsin_Arccos_csqrt_pos) + apply (auto simp: power_le_one csqrt_1_diff_eq) + done + +lemma arcsin_arccos_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arcsin x = -arccos(sqrt(1 - x\<^sup>2))" + using arcsin_arccos_sqrt_pos [of "-x"] + by (simp add: arcsin_minus) + +lemma arccos_arcsin_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = arcsin(sqrt(1 - x\<^sup>2))" + apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) + apply (subst Arccos_Arcsin_csqrt_pos) + apply (auto simp: power_le_one csqrt_1_diff_eq) + done + +lemma arccos_arcsin_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))" + using arccos_arcsin_sqrt_pos [of "-x"] + by (simp add: arccos_minus) + +subsection\<open>continuity results for arcsin and arccos.\<close> + +lemma continuous_on_Arcsin_real [continuous_intros]: + "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} Arcsin" +proof - + have "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (arcsin (Re x))) = + continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (Re (Arcsin (of_real (Re x)))))" + by (rule continuous_on_cong [OF refl]) (simp add: arcsin_eq_Re_Arcsin) + also have "... = ?thesis" + by (rule continuous_on_cong [OF refl]) simp + finally show ?thesis + using continuous_on_arcsin [OF continuous_on_Re [OF continuous_on_id], of "{w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}"] + continuous_on_of_real + by fastforce +qed + +lemma continuous_within_Arcsin_real: + "continuous (at z within {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}) Arcsin" +proof (cases "z \<in> {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}") + case True then show ?thesis + using continuous_on_Arcsin_real continuous_on_eq_continuous_within + by blast +next + case False + with closed_real_abs_le [of 1] show ?thesis + by (rule continuous_within_closed_nontrivial) +qed + +lemma continuous_on_Arccos_real: + "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} Arccos" +proof - + have "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (arccos (Re x))) = + continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (Re (Arccos (of_real (Re x)))))" + by (rule continuous_on_cong [OF refl]) (simp add: arccos_eq_Re_Arccos) + also have "... = ?thesis" + by (rule continuous_on_cong [OF refl]) simp + finally show ?thesis + using continuous_on_arccos [OF continuous_on_Re [OF continuous_on_id], of "{w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}"] + continuous_on_of_real + by fastforce +qed + +lemma continuous_within_Arccos_real: + "continuous (at z within {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}) Arccos" +proof (cases "z \<in> {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}") + case True then show ?thesis + using continuous_on_Arccos_real continuous_on_eq_continuous_within + by blast +next + case False + with closed_real_abs_le [of 1] show ?thesis + by (rule continuous_within_closed_nontrivial) +qed + + +subsection\<open>Roots of unity\<close> + +lemma complex_root_unity: + fixes j::nat + assumes "n \<noteq> 0" + shows "exp(2 * of_real pi * \<i> * of_nat j / of_nat n)^n = 1" +proof - + have *: "of_nat j * (complex_of_real pi * 2) = complex_of_real (2 * real j * pi)" + by (simp add: of_real_numeral) + then show ?thesis + apply (simp add: exp_of_nat_mult [symmetric] mult_ac exp_Euler) + apply (simp only: * cos_of_real sin_of_real) + apply (simp add: ) + done +qed + +lemma complex_root_unity_eq: + fixes j::nat and k::nat + assumes "1 \<le> n" + shows "(exp(2 * of_real pi * \<i> * of_nat j / of_nat n) = exp(2 * of_real pi * \<i> * of_nat k / of_nat n) + \<longleftrightarrow> j mod n = k mod n)" +proof - + have "(\<exists>z::int. \<i> * (of_nat j * (of_real pi * 2)) = + \<i> * (of_nat k * (of_real pi * 2)) + \<i> * (of_int z * (of_nat n * (of_real pi * 2)))) \<longleftrightarrow> + (\<exists>z::int. of_nat j * (\<i> * (of_real pi * 2)) = + (of_nat k + of_nat n * of_int z) * (\<i> * (of_real pi * 2)))" + by (simp add: algebra_simps) + also have "... \<longleftrightarrow> (\<exists>z::int. of_nat j = of_nat k + of_nat n * (of_int z :: complex))" + by simp + also have "... \<longleftrightarrow> (\<exists>z::int. of_nat j = of_nat k + of_nat n * z)" + apply (rule HOL.iff_exI) + apply (auto simp: ) + using of_int_eq_iff apply fastforce + by (metis of_int_add of_int_mult of_int_of_nat_eq) + also have "... \<longleftrightarrow> int j mod int n = int k mod int n" + by (auto simp: zmod_eq_dvd_iff dvd_def algebra_simps) + also have "... \<longleftrightarrow> j mod n = k mod n" + by (metis of_nat_eq_iff zmod_int) + finally have "(\<exists>z. \<i> * (of_nat j * (of_real pi * 2)) = + \<i> * (of_nat k * (of_real pi * 2)) + \<i> * (of_int z * (of_nat n * (of_real pi * 2)))) \<longleftrightarrow> j mod n = k mod n" . + note * = this + show ?thesis + using assms + by (simp add: exp_eq divide_simps mult_ac of_real_numeral *) +qed + +corollary bij_betw_roots_unity: + "bij_betw (\<lambda>j. exp(2 * of_real pi * \<i> * of_nat j / of_nat n)) + {..<n} {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j. j < n}" + by (auto simp: bij_betw_def inj_on_def complex_root_unity_eq) + +lemma complex_root_unity_eq_1: + fixes j::nat and k::nat + assumes "1 \<le> n" + shows "exp(2 * of_real pi * \<i> * of_nat j / of_nat n) = 1 \<longleftrightarrow> n dvd j" +proof - + have "1 = exp(2 * of_real pi * \<i> * (of_nat n / of_nat n))" + using assms by simp + then have "exp(2 * of_real pi * \<i> * (of_nat j / of_nat n)) = 1 \<longleftrightarrow> j mod n = n mod n" + using complex_root_unity_eq [of n j n] assms + by simp + then show ?thesis + by auto +qed + +lemma finite_complex_roots_unity_explicit: + "finite {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j::nat. j < n}" +by simp + +lemma card_complex_roots_unity_explicit: + "card {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j::nat. j < n} = n" + by (simp add: Finite_Set.bij_betw_same_card [OF bij_betw_roots_unity, symmetric]) + +lemma complex_roots_unity: + assumes "1 \<le> n" + shows "{z::complex. z^n = 1} = {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j::nat. j < n}" + apply (rule Finite_Set.card_seteq [symmetric]) + using assms + apply (auto simp: card_complex_roots_unity_explicit finite_roots_unity complex_root_unity card_roots_unity) + done + +lemma card_complex_roots_unity: "1 \<le> n \<Longrightarrow> card {z::complex. z^n = 1} = n" + by (simp add: card_complex_roots_unity_explicit complex_roots_unity) + +lemma complex_not_root_unity: + "1 \<le> n \<Longrightarrow> \<exists>u::complex. norm u = 1 \<and> u^n \<noteq> 1" + apply (rule_tac x="exp (of_real pi * \<i> * of_real (1 / n))" in exI) + apply (auto simp: Re_complex_div_eq_0 exp_of_nat_mult [symmetric] mult_ac exp_Euler) + done + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Conformal_Mappings.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,3681 @@ +section \<open>Conformal Mappings. Consequences of Cauchy's integral theorem.\<close> + +text\<open>By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\<close> + +text\<open>Also Cauchy's residue theorem by Wenda Li (2016)\<close> + +theory Conformal_Mappings +imports "~~/src/HOL/Analysis/Cauchy_Integral_Theorem" + +begin + +subsection\<open>Cauchy's inequality and more versions of Liouville\<close> + +lemma Cauchy_higher_deriv_bound: + assumes holf: "f holomorphic_on (ball z r)" + and contf: "continuous_on (cball z r) f" + and "0 < r" and "0 < n" + and fin : "\<And>w. w \<in> ball z r \<Longrightarrow> f w \<in> ball y B0" + shows "norm ((deriv ^^ n) f z) \<le> (fact n) * B0 / r^n" +proof - + have "0 < B0" using \<open>0 < r\<close> fin [of z] + by (metis ball_eq_empty ex_in_conv fin not_less) + have le_B0: "\<And>w. cmod (w - z) \<le> r \<Longrightarrow> cmod (f w - y) \<le> B0" + apply (rule continuous_on_closure_norm_le [of "ball z r" "\<lambda>w. f w - y"]) + apply (auto simp: \<open>0 < r\<close> dist_norm norm_minus_commute) + apply (rule continuous_intros contf)+ + using fin apply (simp add: dist_commute dist_norm less_eq_real_def) + done + have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w) z - (deriv ^^ n) (\<lambda>w. y) z" + using \<open>0 < n\<close> by simp + also have "... = (deriv ^^ n) (\<lambda>w. f w - y) z" + by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \<open>0 < r\<close>) + finally have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w - y) z" . + have contf': "continuous_on (cball z r) (\<lambda>u. f u - y)" + by (rule contf continuous_intros)+ + have holf': "(\<lambda>u. (f u - y)) holomorphic_on (ball z r)" + by (simp add: holf holomorphic_on_diff) + define a where "a = (2 * pi)/(fact n)" + have "0 < a" by (simp add: a_def) + have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" + using \<open>0 < r\<close> by (simp add: a_def divide_simps) + have der_dif: "(deriv ^^ n) (\<lambda>w. f w - y) z = (deriv ^^ n) f z" + using \<open>0 < r\<close> \<open>0 < n\<close> + by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) + have "norm ((2 * of_real pi * \<i>)/(fact n) * (deriv ^^ n) (\<lambda>w. f w - y) z) + \<le> (B0/r^(Suc n)) * (2 * pi * r)" + apply (rule has_contour_integral_bound_circlepath [of "(\<lambda>u. (f u - y)/(u - z)^(Suc n))" _ z]) + using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] + using \<open>0 < B0\<close> \<open>0 < r\<close> + apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) + done + then show ?thesis + using \<open>0 < r\<close> + by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) +qed + +proposition Cauchy_inequality: + assumes holf: "f holomorphic_on (ball \<xi> r)" + and contf: "continuous_on (cball \<xi> r) f" + and "0 < r" + and nof: "\<And>x. norm(\<xi>-x) = r \<Longrightarrow> norm(f x) \<le> B" + shows "norm ((deriv ^^ n) f \<xi>) \<le> (fact n) * B / r^n" +proof - + obtain x where "norm (\<xi>-x) = r" + by (metis abs_of_nonneg add_diff_cancel_left' \<open>0 < r\<close> diff_add_cancel + dual_order.strict_implies_order norm_of_real) + then have "0 \<le> B" + by (metis nof norm_not_less_zero not_le order_trans) + have "((\<lambda>u. f u / (u - \<xi>) ^ Suc n) has_contour_integral (2 * pi) * \<i> / fact n * (deriv ^^ n) f \<xi>) + (circlepath \<xi> r)" + apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) + using \<open>0 < r\<close> by simp + then have "norm ((2 * pi * \<i>)/(fact n) * (deriv ^^ n) f \<xi>) \<le> (B / r^(Suc n)) * (2 * pi * r)" + apply (rule has_contour_integral_bound_circlepath) + using \<open>0 \<le> B\<close> \<open>0 < r\<close> + apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) + done + then show ?thesis using \<open>0 < r\<close> + by (simp add: norm_divide norm_mult field_simps) +qed + +proposition Liouville_polynomial: + assumes holf: "f holomorphic_on UNIV" + and nof: "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) \<le> B * norm z ^ n" + shows "f \<xi> = (\<Sum>k\<le>n. (deriv^^k) f 0 / fact k * \<xi> ^ k)" +proof (cases rule: le_less_linear [THEN disjE]) + assume "B \<le> 0" + then have "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) = 0" + by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) + then have f0: "(f \<longlongrightarrow> 0) at_infinity" + using Lim_at_infinity by force + then have [simp]: "f = (\<lambda>w. 0)" + using Liouville_weak [OF holf, of 0] + by (simp add: eventually_at_infinity f0) meson + show ?thesis by simp +next + assume "0 < B" + have "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * (\<xi> - 0)^k) sums f \<xi>)" + apply (rule holomorphic_power_series [where r = "norm \<xi> + 1"]) + using holf holomorphic_on_subset apply auto + done + then have sumsf: "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * \<xi>^k) sums f \<xi>)" by simp + have "(deriv ^^ k) f 0 / fact k * \<xi> ^ k = 0" if "k>n" for k + proof (cases "(deriv ^^ k) f 0 = 0") + case True then show ?thesis by simp + next + case False + define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))" + have "1 \<le> abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))" + using \<open>0 < B\<close> by simp + then have wge1: "1 \<le> norm w" + by (metis norm_of_real w_def) + then have "w \<noteq> 0" by auto + have kB: "0 < fact k * B" + using \<open>0 < B\<close> by simp + then have "0 \<le> fact k * B / cmod ((deriv ^^ k) f 0)" + by simp + then have wgeA: "A \<le> cmod w" + by (simp only: w_def norm_of_real) + have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))" + using \<open>0 < B\<close> by simp + then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" + by (metis norm_of_real w_def) + then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" + using False by (simp add: divide_simps mult.commute split: if_split_asm) + also have "... \<le> fact k * (B * norm w ^ n) / norm w ^ k" + apply (rule Cauchy_inequality) + using holf holomorphic_on_subset apply force + using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast + using \<open>w \<noteq> 0\<close> apply (simp add:) + by (metis nof wgeA dist_0_norm dist_norm) + also have "... = fact k * (B * 1 / cmod w ^ (k-n))" + apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) + using \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> apply (simp add: divide_simps semiring_normalization_rules) + done + also have "... = fact k * B / cmod w ^ (k-n)" + by simp + finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . + then have "1 / cmod w < 1 / cmod w ^ (k - n)" + by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) + then have "cmod w ^ (k - n) < cmod w" + by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) + with self_le_power [OF wge1] have False + by (meson diff_is_0_eq not_gr0 not_le that) + then show ?thesis by blast + qed + then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \<xi> ^ (k + Suc n) = 0" for k + using not_less_eq by blast + then have "(\<lambda>i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \<xi> ^ (i + Suc n)) sums 0" + by (rule sums_0) + with sums_split_initial_segment [OF sumsf, where n = "Suc n"] + show ?thesis + using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce +qed + +text\<open>Every bounded entire function is a constant function.\<close> +theorem Liouville_theorem: + assumes holf: "f holomorphic_on UNIV" + and bf: "bounded (range f)" + obtains c where "\<And>z. f z = c" +proof - + obtain B where "\<And>z. cmod (f z) \<le> B" + by (meson bf bounded_pos rangeI) + then show ?thesis + using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast +qed + + + +text\<open>A holomorphic function f has only isolated zeros unless f is 0.\<close> + +proposition powser_0_nonzero: + fixes a :: "nat \<Rightarrow> 'a::{real_normed_field,banach}" + assumes r: "0 < r" + and sm: "\<And>x. norm (x - \<xi>) < r \<Longrightarrow> (\<lambda>n. a n * (x - \<xi>) ^ n) sums (f x)" + and [simp]: "f \<xi> = 0" + and m0: "a m \<noteq> 0" and "m>0" + obtains s where "0 < s" and "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0" +proof - + have "r \<le> conv_radius a" + using sm sums_summable by (auto simp: le_conv_radius_iff [where \<xi>=\<xi>]) + obtain m where am: "a m \<noteq> 0" and az [simp]: "(\<And>n. n<m \<Longrightarrow> a n = 0)" + apply (rule_tac m = "LEAST n. a n \<noteq> 0" in that) + using m0 + apply (rule LeastI2) + apply (fastforce intro: dest!: not_less_Least)+ + done + define b where "b i = a (i+m) / a m" for i + define g where "g x = suminf (\<lambda>i. b i * (x - \<xi>) ^ i)" for x + have [simp]: "b 0 = 1" + by (simp add: am b_def) + { fix x::'a + assume "norm (x - \<xi>) < r" + then have "(\<lambda>n. (a m * (x - \<xi>)^m) * (b n * (x - \<xi>)^n)) sums (f x)" + using am az sm sums_zero_iff_shift [of m "(\<lambda>n. a n * (x - \<xi>) ^ n)" "f x"] + by (simp add: b_def monoid_mult_class.power_add algebra_simps) + then have "x \<noteq> \<xi> \<Longrightarrow> (\<lambda>n. b n * (x - \<xi>)^n) sums (f x / (a m * (x - \<xi>)^m))" + using am by (simp add: sums_mult_D) + } note bsums = this + then have "norm (x - \<xi>) < r \<Longrightarrow> summable (\<lambda>n. b n * (x - \<xi>)^n)" for x + using sums_summable by (cases "x=\<xi>") auto + then have "r \<le> conv_radius b" + by (simp add: le_conv_radius_iff [where \<xi>=\<xi>]) + then have "r/2 < conv_radius b" + using not_le order_trans r by fastforce + then have "continuous_on (cball \<xi> (r/2)) g" + using powser_continuous_suminf [of "r/2" b \<xi>] by (simp add: g_def) + then obtain s where "s>0" "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> dist (g x) (g \<xi>) < 1/2" + apply (rule continuous_onE [where x=\<xi> and e = "1/2"]) + using r apply (auto simp: norm_minus_commute dist_norm) + done + moreover have "g \<xi> = 1" + by (simp add: g_def) + ultimately have gnz: "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> (g x) \<noteq> 0" + by fastforce + have "f x \<noteq> 0" if "x \<noteq> \<xi>" "norm (x - \<xi>) \<le> s" "norm (x - \<xi>) \<le> r/2" for x + using bsums [of x] that gnz [of x] + apply (auto simp: g_def) + using r sums_iff by fastforce + then show ?thesis + apply (rule_tac s="min s (r/2)" in that) + using \<open>0 < r\<close> \<open>0 < s\<close> by (auto simp: dist_commute dist_norm) +qed + +proposition isolated_zeros: + assumes holf: "f holomorphic_on S" + and "open S" "connected S" "\<xi> \<in> S" "f \<xi> = 0" "\<beta> \<in> S" "f \<beta> \<noteq> 0" + obtains r where "0 < r" "ball \<xi> r \<subseteq> S" "\<And>z. z \<in> ball \<xi> r - {\<xi>} \<Longrightarrow> f z \<noteq> 0" +proof - + obtain r where "0 < r" and r: "ball \<xi> r \<subseteq> S" + using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_ball_eq by blast + have powf: "((\<lambda>n. (deriv ^^ n) f \<xi> / (fact n) * (z - \<xi>)^n) sums f z)" if "z \<in> ball \<xi> r" for z + apply (rule holomorphic_power_series [OF _ that]) + apply (rule holomorphic_on_subset [OF holf r]) + done + obtain m where m: "(deriv ^^ m) f \<xi> / (fact m) \<noteq> 0" + using holomorphic_fun_eq_0_on_connected [OF holf \<open>open S\<close> \<open>connected S\<close> _ \<open>\<xi> \<in> S\<close> \<open>\<beta> \<in> S\<close>] \<open>f \<beta> \<noteq> 0\<close> + by auto + then have "m \<noteq> 0" using assms(5) funpow_0 by fastforce + obtain s where "0 < s" and s: "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0" + apply (rule powser_0_nonzero [OF \<open>0 < r\<close> powf \<open>f \<xi> = 0\<close> m]) + using \<open>m \<noteq> 0\<close> by (auto simp: dist_commute dist_norm) + have "0 < min r s" by (simp add: \<open>0 < r\<close> \<open>0 < s\<close>) + then show ?thesis + apply (rule that) + using r s by auto +qed + + +proposition analytic_continuation: + assumes holf: "f holomorphic_on S" + and S: "open S" "connected S" + and "U \<subseteq> S" "\<xi> \<in> S" + and "\<xi> islimpt U" + and fU0 [simp]: "\<And>z. z \<in> U \<Longrightarrow> f z = 0" + and "w \<in> S" + shows "f w = 0" +proof - + obtain e where "0 < e" and e: "cball \<xi> e \<subseteq> S" + using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_cball_eq by blast + define T where "T = cball \<xi> e \<inter> U" + have contf: "continuous_on (closure T) f" + by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on + holomorphic_on_subset inf.cobounded1) + have fT0 [simp]: "\<And>x. x \<in> T \<Longrightarrow> f x = 0" + by (simp add: T_def) + have "\<And>r. \<lbrakk>\<forall>e>0. \<exists>x'\<in>U. x' \<noteq> \<xi> \<and> dist x' \<xi> < e; 0 < r\<rbrakk> \<Longrightarrow> \<exists>x'\<in>cball \<xi> e \<inter> U. x' \<noteq> \<xi> \<and> dist x' \<xi> < r" + by (metis \<open>0 < e\<close> IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj) + then have "\<xi> islimpt T" using \<open>\<xi> islimpt U\<close> + by (auto simp: T_def islimpt_approachable) + then have "\<xi> \<in> closure T" + by (simp add: closure_def) + then have "f \<xi> = 0" + by (auto simp: continuous_constant_on_closure [OF contf]) + show ?thesis + apply (rule ccontr) + apply (rule isolated_zeros [OF holf \<open>open S\<close> \<open>connected S\<close> \<open>\<xi> \<in> S\<close> \<open>f \<xi> = 0\<close> \<open>w \<in> S\<close>], assumption) + by (metis open_ball \<open>\<xi> islimpt T\<close> centre_in_ball fT0 insertE insert_Diff islimptE) +qed + + +subsection\<open>Open mapping theorem\<close> + +lemma holomorphic_contract_to_zero: + assumes contf: "continuous_on (cball \<xi> r) f" + and holf: "f holomorphic_on ball \<xi> r" + and "0 < r" + and norm_less: "\<And>z. norm(\<xi> - z) = r \<Longrightarrow> norm(f \<xi>) < norm(f z)" + obtains z where "z \<in> ball \<xi> r" "f z = 0" +proof - + { assume fnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w \<noteq> 0" + then have "0 < norm (f \<xi>)" + by (simp add: \<open>0 < r\<close>) + have fnz': "\<And>w. w \<in> cball \<xi> r \<Longrightarrow> f w \<noteq> 0" + by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) + have "frontier(cball \<xi> r) \<noteq> {}" + using \<open>0 < r\<close> by simp + define g where [abs_def]: "g z = inverse (f z)" for z + have contg: "continuous_on (cball \<xi> r) g" + unfolding g_def using contf continuous_on_inverse fnz' by blast + have holg: "g holomorphic_on ball \<xi> r" + unfolding g_def using fnz holf holomorphic_on_inverse by blast + have "frontier (cball \<xi> r) \<subseteq> cball \<xi> r" + by (simp add: subset_iff) + then have contf': "continuous_on (frontier (cball \<xi> r)) f" + and contg': "continuous_on (frontier (cball \<xi> r)) g" + by (blast intro: contf contg continuous_on_subset)+ + have froc: "frontier(cball \<xi> r) \<noteq> {}" + using \<open>0 < r\<close> by simp + moreover have "continuous_on (frontier (cball \<xi> r)) (norm o f)" + using contf' continuous_on_compose continuous_on_norm_id by blast + ultimately obtain w where w: "w \<in> frontier(cball \<xi> r)" + and now: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (f w) \<le> norm (f x)" + apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]]) + apply (simp add:) + done + then have fw: "0 < norm (f w)" + by (simp add: fnz') + have "continuous_on (frontier (cball \<xi> r)) (norm o g)" + using contg' continuous_on_compose continuous_on_norm_id by blast + then obtain v where v: "v \<in> frontier(cball \<xi> r)" + and nov: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (g v) \<ge> norm (g x)" + apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]]) + apply (simp add:) + done + then have fv: "0 < norm (f v)" + by (simp add: fnz') + have "norm ((deriv ^^ 0) g \<xi>) \<le> fact 0 * norm (g v) / r ^ 0" + by (rule Cauchy_inequality [OF holg contg \<open>0 < r\<close>]) (simp add: dist_norm nov) + then have "cmod (g \<xi>) \<le> norm (g v)" + by simp + with w have wr: "norm (\<xi> - w) = r" and nfw: "norm (f w) \<le> norm (f \<xi>)" + apply (simp_all add: dist_norm) + by (metis \<open>0 < cmod (f \<xi>)\<close> g_def less_imp_inverse_less norm_inverse not_le now order_trans v) + with fw have False + using norm_less by force + } + with that show ?thesis by blast +qed + + +theorem open_mapping_thm: + assumes holf: "f holomorphic_on S" + and S: "open S" "connected S" + and "open U" "U \<subseteq> S" + and fne: "~ f constant_on S" + shows "open (f ` U)" +proof - + have *: "open (f ` U)" + if "U \<noteq> {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\<And>x. \<exists>y \<in> U. f y \<noteq> x" + for U + proof (clarsimp simp: open_contains_ball) + fix \<xi> assume \<xi>: "\<xi> \<in> U" + show "\<exists>e>0. ball (f \<xi>) e \<subseteq> f ` U" + proof - + have hol: "(\<lambda>z. f z - f \<xi>) holomorphic_on U" + by (rule holomorphic_intros that)+ + obtain s where "0 < s" and sbU: "ball \<xi> s \<subseteq> U" + and sne: "\<And>z. z \<in> ball \<xi> s - {\<xi>} \<Longrightarrow> (\<lambda>z. f z - f \<xi>) z \<noteq> 0" + using isolated_zeros [OF hol U \<xi>] by (metis fneU right_minus_eq) + obtain r where "0 < r" and r: "cball \<xi> r \<subseteq> ball \<xi> s" + apply (rule_tac r="s/2" in that) + using \<open>0 < s\<close> by auto + have "cball \<xi> r \<subseteq> U" + using sbU r by blast + then have frsbU: "frontier (cball \<xi> r) \<subseteq> U" + using Diff_subset frontier_def order_trans by fastforce + then have cof: "compact (frontier(cball \<xi> r))" + by blast + have frne: "frontier (cball \<xi> r) \<noteq> {}" + using \<open>0 < r\<close> by auto + have contfr: "continuous_on (frontier (cball \<xi> r)) (\<lambda>z. norm (f z - f \<xi>))" + apply (rule continuous_on_compose2 [OF Complex_Analysis_Basics.continuous_on_norm_id]) + using hol frsbU holomorphic_on_imp_continuous_on holomorphic_on_subset by blast+ + obtain w where "norm (\<xi> - w) = r" + and w: "(\<And>z. norm (\<xi> - z) = r \<Longrightarrow> norm (f w - f \<xi>) \<le> norm(f z - f \<xi>))" + apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) + apply (simp add: dist_norm) + done + moreover define \<epsilon> where "\<epsilon> \<equiv> norm (f w - f \<xi>) / 3" + ultimately have "0 < \<epsilon>" + using \<open>0 < r\<close> dist_complex_def r sne by auto + have "ball (f \<xi>) \<epsilon> \<subseteq> f ` U" + proof + fix \<gamma> + assume \<gamma>: "\<gamma> \<in> ball (f \<xi>) \<epsilon>" + have *: "cmod (\<gamma> - f \<xi>) < cmod (\<gamma> - f z)" if "cmod (\<xi> - z) = r" for z + proof - + have lt: "cmod (f w - f \<xi>) / 3 < cmod (\<gamma> - f z)" + using w [OF that] \<gamma> + using dist_triangle2 [of "f \<xi>" "\<gamma>" "f z"] dist_triangle2 [of "f \<xi>" "f z" \<gamma>] + by (simp add: \<epsilon>_def dist_norm norm_minus_commute) + show ?thesis + by (metis \<epsilon>_def dist_commute dist_norm less_trans lt mem_ball \<gamma>) + qed + have "continuous_on (cball \<xi> r) (\<lambda>z. \<gamma> - f z)" + apply (rule continuous_intros)+ + using \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close> + apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) + done + moreover have "(\<lambda>z. \<gamma> - f z) holomorphic_on ball \<xi> r" + apply (rule holomorphic_intros)+ + apply (metis \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close> holomorphic_on_subset interior_cball interior_subset) + done + ultimately obtain z where "z \<in> ball \<xi> r" "\<gamma> - f z = 0" + apply (rule holomorphic_contract_to_zero) + apply (blast intro!: \<open>0 < r\<close> *)+ + done + then show "\<gamma> \<in> f ` U" + using \<open>cball \<xi> r \<subseteq> U\<close> by fastforce + qed + then show ?thesis using \<open>0 < \<epsilon>\<close> by blast + qed + qed + have "open (f ` X)" if "X \<in> components U" for X + proof - + have holfU: "f holomorphic_on U" + using \<open>U \<subseteq> S\<close> holf holomorphic_on_subset by blast + have "X \<noteq> {}" + using that by (simp add: in_components_nonempty) + moreover have "open X" + using that \<open>open U\<close> open_components by auto + moreover have "connected X" + using that in_components_maximal by blast + moreover have "f holomorphic_on X" + by (meson that holfU holomorphic_on_subset in_components_maximal) + moreover have "\<exists>y\<in>X. f y \<noteq> x" for x + proof (rule ccontr) + assume not: "\<not> (\<exists>y\<in>X. f y \<noteq> x)" + have "X \<subseteq> S" + using \<open>U \<subseteq> S\<close> in_components_subset that by blast + obtain w where w: "w \<in> X" using \<open>X \<noteq> {}\<close> by blast + have wis: "w islimpt X" + using w \<open>open X\<close> interior_eq by auto + have hol: "(\<lambda>z. f z - x) holomorphic_on S" + by (simp add: holf holomorphic_on_diff) + with fne [unfolded constant_on_def] analytic_continuation [OF hol S \<open>X \<subseteq> S\<close> _ wis] + not \<open>X \<subseteq> S\<close> w + show False by auto + qed + ultimately show ?thesis + by (rule *) + qed + then have "open (f ` \<Union>components U)" + by (metis (no_types, lifting) imageE image_Union open_Union) + then show ?thesis + by force +qed + + +text\<open>No need for @{term S} to be connected. But the nonconstant condition is stronger.\<close> +corollary open_mapping_thm2: + assumes holf: "f holomorphic_on S" + and S: "open S" + and "open U" "U \<subseteq> S" + and fnc: "\<And>X. \<lbrakk>open X; X \<subseteq> S; X \<noteq> {}\<rbrakk> \<Longrightarrow> ~ f constant_on X" + shows "open (f ` U)" +proof - + have "S = \<Union>(components S)" by simp + with \<open>U \<subseteq> S\<close> have "U = (\<Union>C \<in> components S. C \<inter> U)" by auto + then have "f ` U = (\<Union>C \<in> components S. f ` (C \<inter> U))" + using image_UN by fastforce + moreover + { fix C assume "C \<in> components S" + with S \<open>C \<in> components S\<close> open_components in_components_connected + have C: "open C" "connected C" by auto + have "C \<subseteq> S" + by (metis \<open>C \<in> components S\<close> in_components_maximal) + have nf: "\<not> f constant_on C" + apply (rule fnc) + using C \<open>C \<subseteq> S\<close> \<open>C \<in> components S\<close> in_components_nonempty by auto + have "f holomorphic_on C" + by (metis holf holomorphic_on_subset \<open>C \<subseteq> S\<close>) + then have "open (f ` (C \<inter> U))" + apply (rule open_mapping_thm [OF _ C _ _ nf]) + apply (simp add: C \<open>open U\<close> open_Int, blast) + done + } ultimately show ?thesis + by force +qed + +corollary open_mapping_thm3: + assumes holf: "f holomorphic_on S" + and "open S" and injf: "inj_on f S" + shows "open (f ` S)" +apply (rule open_mapping_thm2 [OF holf]) +using assms +apply (simp_all add:) +using injective_not_constant subset_inj_on by blast + + + +subsection\<open>Maximum Modulus Principle\<close> + +text\<open>If @{term f} is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is + properly within the domain of @{term f}.\<close> + +proposition maximum_modulus_principle: + assumes holf: "f holomorphic_on S" + and S: "open S" "connected S" + and "open U" "U \<subseteq> S" "\<xi> \<in> U" + and no: "\<And>z. z \<in> U \<Longrightarrow> norm(f z) \<le> norm(f \<xi>)" + shows "f constant_on S" +proof (rule ccontr) + assume "\<not> f constant_on S" + then have "open (f ` U)" + using open_mapping_thm assms by blast + moreover have "~ open (f ` U)" + proof - + have "\<exists>t. cmod (f \<xi> - t) < e \<and> t \<notin> f ` U" if "0 < e" for e + apply (rule_tac x="if 0 < Re(f \<xi>) then f \<xi> + (e/2) else f \<xi> - (e/2)" in exI) + using that + apply (simp add: dist_norm) + apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym) + done + then show ?thesis + unfolding open_contains_ball by (metis \<open>\<xi> \<in> U\<close> contra_subsetD dist_norm imageI mem_ball) + qed + ultimately show False + by blast +qed + + +proposition maximum_modulus_frontier: + assumes holf: "f holomorphic_on (interior S)" + and contf: "continuous_on (closure S) f" + and bos: "bounded S" + and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> norm(f z) \<le> B" + and "\<xi> \<in> S" + shows "norm(f \<xi>) \<le> B" +proof - + have "compact (closure S)" using bos + by (simp add: bounded_closure compact_eq_bounded_closed) + moreover have "continuous_on (closure S) (cmod \<circ> f)" + using contf continuous_on_compose continuous_on_norm_id by blast + ultimately obtain z where zin: "z \<in> closure S" and z: "\<And>y. y \<in> closure S \<Longrightarrow> (cmod \<circ> f) y \<le> (cmod \<circ> f) z" + using continuous_attains_sup [of "closure S" "norm o f"] \<open>\<xi> \<in> S\<close> by auto + then consider "z \<in> frontier S" | "z \<in> interior S" using frontier_def by auto + then have "norm(f z) \<le> B" + proof cases + case 1 then show ?thesis using leB by blast + next + case 2 + have zin: "z \<in> connected_component_set (interior S) z" + by (simp add: 2) + have "f constant_on (connected_component_set (interior S) z)" + apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin]) + apply (metis connected_component_subset holf holomorphic_on_subset) + apply (simp_all add: open_connected_component) + by (metis closure_subset comp_eq_dest_lhs interior_subset subsetCE z connected_component_in) + then obtain c where c: "\<And>w. w \<in> connected_component_set (interior S) z \<Longrightarrow> f w = c" + by (auto simp: constant_on_def) + have "f ` closure(connected_component_set (interior S) z) \<subseteq> {c}" + apply (rule image_closure_subset) + apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) + using c + apply auto + done + then have cc: "\<And>w. w \<in> closure(connected_component_set (interior S) z) \<Longrightarrow> f w = c" by blast + have "frontier(connected_component_set (interior S) z) \<noteq> {}" + apply (simp add: frontier_eq_empty) + by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV) + then obtain w where w: "w \<in> frontier(connected_component_set (interior S) z)" + by auto + then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def) + also have "... \<le> B" + apply (rule leB) + using w +using frontier_interior_subset frontier_of_connected_component_subset by blast + finally show ?thesis . + qed + then show ?thesis + using z \<open>\<xi> \<in> S\<close> closure_subset by fastforce +qed + +corollary maximum_real_frontier: + assumes holf: "f holomorphic_on (interior S)" + and contf: "continuous_on (closure S) f" + and bos: "bounded S" + and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> Re(f z) \<le> B" + and "\<xi> \<in> S" + shows "Re(f \<xi>) \<le> B" +using maximum_modulus_frontier [of "exp o f" S "exp B"] + Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms +by auto + + +subsection\<open>Factoring out a zero according to its order\<close> + +lemma holomorphic_factor_order_of_zero: + assumes holf: "f holomorphic_on S" + and os: "open S" + and "\<xi> \<in> S" "0 < n" + and dnz: "(deriv ^^ n) f \<xi> \<noteq> 0" + and dfz: "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0" + obtains g r where "0 < r" + "g holomorphic_on ball \<xi> r" + "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w" + "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" +proof - + obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE) + then have holfb: "f holomorphic_on ball \<xi> r" + using holf holomorphic_on_subset by blast + define g where "g w = suminf (\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i)" for w + have sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w" + and feq: "f w - f \<xi> = (w - \<xi>)^n * g w" + if w: "w \<in> ball \<xi> r" for w + proof - + define powf where "powf = (\<lambda>i. (deriv ^^ i) f \<xi>/(fact i) * (w - \<xi>)^i)" + have sing: "{..<n} - {i. powf i = 0} = (if f \<xi> = 0 then {} else {0})" + unfolding powf_def using \<open>0 < n\<close> dfz by (auto simp: dfz; metis funpow_0 not_gr0) + have "powf sums f w" + unfolding powf_def by (rule holomorphic_power_series [OF holfb w]) + moreover have "(\<Sum>i<n. powf i) = f \<xi>" + apply (subst Groups_Big.comm_monoid_add_class.setsum.setdiff_irrelevant [symmetric]) + apply (simp add:) + apply (simp only: dfz sing) + apply (simp add: powf_def) + done + ultimately have fsums: "(\<lambda>i. powf (i+n)) sums (f w - f \<xi>)" + using w sums_iff_shift' by metis + then have *: "summable (\<lambda>i. (w - \<xi>) ^ n * ((deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n)))" + unfolding powf_def using sums_summable + by (auto simp: power_add mult_ac) + have "summable (\<lambda>i. (deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n))" + proof (cases "w=\<xi>") + case False then show ?thesis + using summable_mult [OF *, of "1 / (w - \<xi>) ^ n"] by (simp add:) + next + case True then show ?thesis + by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"] + split: if_split_asm) + qed + then show sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w" + by (simp add: summable_sums_iff g_def) + show "f w - f \<xi> = (w - \<xi>)^n * g w" + apply (rule sums_unique2) + apply (rule fsums [unfolded powf_def]) + using sums_mult [OF sumsg, of "(w - \<xi>) ^ n"] + by (auto simp: power_add mult_ac) + qed + then have holg: "g holomorphic_on ball \<xi> r" + by (meson sumsg power_series_holomorphic) + then have contg: "continuous_on (ball \<xi> r) g" + by (blast intro: holomorphic_on_imp_continuous_on) + have "g \<xi> \<noteq> 0" + using dnz unfolding g_def + by (subst suminf_finite [of "{0}"]) auto + obtain d where "0 < d" and d: "\<And>w. w \<in> ball \<xi> d \<Longrightarrow> g w \<noteq> 0" + apply (rule exE [OF continuous_on_avoid [OF contg _ \<open>g \<xi> \<noteq> 0\<close>]]) + using \<open>0 < r\<close> + apply force + by (metis \<open>0 < r\<close> less_trans mem_ball not_less_iff_gr_or_eq) + show ?thesis + apply (rule that [where g=g and r ="min r d"]) + using \<open>0 < r\<close> \<open>0 < d\<close> holg + apply (auto simp: feq holomorphic_on_subset subset_ball d) + done +qed + + +lemma holomorphic_factor_order_of_zero_strong: + assumes holf: "f holomorphic_on S" "open S" "\<xi> \<in> S" "0 < n" + and "(deriv ^^ n) f \<xi> \<noteq> 0" + and "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0" + obtains g r where "0 < r" + "g holomorphic_on ball \<xi> r" + "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n" + "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" +proof - + obtain g r where "0 < r" + and holg: "g holomorphic_on ball \<xi> r" + and feq: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w" + and gne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" + by (auto intro: holomorphic_factor_order_of_zero [OF assms]) + have con: "continuous_on (ball \<xi> r) (\<lambda>z. deriv g z / g z)" + by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on) + have cd: "\<And>x. dist \<xi> x < r \<Longrightarrow> (\<lambda>z. deriv g z / g z) field_differentiable at x" + apply (rule derivative_intros)+ + using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) + apply (metis Topology_Euclidean_Space.open_ball at_within_open holg holomorphic_on_def mem_ball) + using gne mem_ball by blast + obtain h where h: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> (h has_field_derivative deriv g x / g x) (at x)" + apply (rule exE [OF holomorphic_convex_primitive [of "ball \<xi> r" "{}" "\<lambda>z. deriv g z / g z"]]) + apply (auto simp: con cd) + apply (metis open_ball at_within_open mem_ball) + done + then have "continuous_on (ball \<xi> r) h" + by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open) + then have con: "continuous_on (ball \<xi> r) (\<lambda>x. exp (h x) / g x)" + by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne) + have 0: "dist \<xi> x < r \<Longrightarrow> ((\<lambda>x. exp (h x) / g x) has_field_derivative 0) (at x)" for x + apply (rule h derivative_eq_intros | simp)+ + apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2]) + using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h) + done + obtain c where c: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> exp (h x) / g x = c" + by (rule DERIV_zero_connected_constant [of "ball \<xi> r" "{}" "\<lambda>x. exp(h x) / g x"]) (auto simp: con 0) + have hol: "(\<lambda>z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \<xi> r" + apply (rule holomorphic_on_compose [unfolded o_def, where g = exp]) + apply (rule holomorphic_intros)+ + using h holomorphic_on_open apply blast + apply (rule holomorphic_intros)+ + using \<open>0 < n\<close> apply (simp add:) + apply (rule holomorphic_intros)+ + done + show ?thesis + apply (rule that [where g="\<lambda>z. exp((Ln(inverse c) + h z)/n)" and r =r]) + using \<open>0 < r\<close> \<open>0 < n\<close> + apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric]) + apply (rule hol) + apply (simp add: Transcendental.exp_add gne) + done +qed + + +lemma + fixes k :: "'a::wellorder" + assumes a_def: "a == LEAST x. P x" and P: "P k" + shows def_LeastI: "P a" and def_Least_le: "a \<le> k" +unfolding a_def +by (rule LeastI Least_le; rule P)+ + +lemma holomorphic_factor_zero_nonconstant: + assumes holf: "f holomorphic_on S" and S: "open S" "connected S" + and "\<xi> \<in> S" "f \<xi> = 0" + and nonconst: "\<And>c. \<exists>z \<in> S. f z \<noteq> c" + obtains g r n + where "0 < n" "0 < r" "ball \<xi> r \<subseteq> S" + "g holomorphic_on ball \<xi> r" + "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w = (w - \<xi>)^n * g w" + "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" +proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0") + case True then show ?thesis + using holomorphic_fun_eq_const_on_connected [OF holf S _ \<open>\<xi> \<in> S\<close>] nonconst by auto +next + case False + then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0" by blast + obtain r0 where "r0 > 0" "ball \<xi> r0 \<subseteq> S" using S openE \<open>\<xi> \<in> S\<close> by auto + define n where "n \<equiv> LEAST n. (deriv ^^ n) f \<xi> \<noteq> 0" + have n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0" + by (rule def_LeastI [OF n_def]) (rule n0) + then have "0 < n" using \<open>f \<xi> = 0\<close> + using funpow_0 by fastforce + have n_min: "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0" + using def_Least_le [OF n_def] not_le by blast + then obtain g r1 + where "0 < r1" "g holomorphic_on ball \<xi> r1" + "\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> f w = (w - \<xi>) ^ n * g w" + "\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> g w \<noteq> 0" + by (auto intro: holomorphic_factor_order_of_zero [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne] simp: \<open>f \<xi> = 0\<close>) + then show ?thesis + apply (rule_tac g=g and r="min r0 r1" and n=n in that) + using \<open>0 < n\<close> \<open>0 < r0\<close> \<open>0 < r1\<close> \<open>ball \<xi> r0 \<subseteq> S\<close> + apply (auto simp: subset_ball intro: holomorphic_on_subset) + done +qed + + +lemma holomorphic_lower_bound_difference: + assumes holf: "f holomorphic_on S" and S: "open S" "connected S" + and "\<xi> \<in> S" and "\<phi> \<in> S" + and fne: "f \<phi> \<noteq> f \<xi>" + obtains k n r + where "0 < k" "0 < r" + "ball \<xi> r \<subseteq> S" + "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> k * norm(w - \<xi>)^n \<le> norm(f w - f \<xi>)" +proof - + define n where "n = (LEAST n. 0 < n \<and> (deriv ^^ n) f \<xi> \<noteq> 0)" + obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0" + using fne holomorphic_fun_eq_const_on_connected [OF holf S] \<open>\<xi> \<in> S\<close> \<open>\<phi> \<in> S\<close> by blast + then have "0 < n" and n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0" + unfolding n_def by (metis (mono_tags, lifting) LeastI)+ + have n_min: "\<And>k. \<lbrakk>0 < k; k < n\<rbrakk> \<Longrightarrow> (deriv ^^ k) f \<xi> = 0" + unfolding n_def by (blast dest: not_less_Least) + then obtain g r + where "0 < r" and holg: "g holomorphic_on ball \<xi> r" + and fne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>) ^ n * g w" + and gnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" + by (auto intro: holomorphic_factor_order_of_zero [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne]) + obtain e where "e>0" and e: "ball \<xi> e \<subseteq> S" using assms by (blast elim!: openE) + then have holfb: "f holomorphic_on ball \<xi> e" + using holf holomorphic_on_subset by blast + define d where "d = (min e r) / 2" + have "0 < d" using \<open>0 < r\<close> \<open>0 < e\<close> by (simp add: d_def) + have "d < r" + using \<open>0 < r\<close> by (auto simp: d_def) + then have cbb: "cball \<xi> d \<subseteq> ball \<xi> r" + by (auto simp: cball_subset_ball_iff) + then have "g holomorphic_on cball \<xi> d" + by (rule holomorphic_on_subset [OF holg]) + then have "closed (g ` cball \<xi> d)" + by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on) + moreover have "g ` cball \<xi> d \<noteq> {}" + using \<open>0 < d\<close> by auto + ultimately obtain x where x: "x \<in> g ` cball \<xi> d" and "\<And>y. y \<in> g ` cball \<xi> d \<Longrightarrow> dist 0 x \<le> dist 0 y" + by (rule distance_attains_inf) blast + then have leg: "\<And>w. w \<in> cball \<xi> d \<Longrightarrow> norm x \<le> norm (g w)" + by auto + have "ball \<xi> d \<subseteq> cball \<xi> d" by auto + also have "... \<subseteq> ball \<xi> e" using \<open>0 < d\<close> d_def by auto + also have "... \<subseteq> S" by (rule e) + finally have dS: "ball \<xi> d \<subseteq> S" . + moreover have "x \<noteq> 0" using gnz x \<open>d < r\<close> by auto + ultimately show ?thesis + apply (rule_tac k="norm x" and n=n and r=d in that) + using \<open>d < r\<close> leg + apply (auto simp: \<open>0 < d\<close> fne norm_mult norm_power algebra_simps mult_right_mono) + done +qed + +lemma + assumes holf: "f holomorphic_on (S - {\<xi>})" and \<xi>: "\<xi> \<in> interior S" + shows holomorphic_on_extend_lim: + "(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow> + ((\<lambda>z. (z - \<xi>) * f z) \<longlongrightarrow> 0) (at \<xi>)" + (is "?P = ?Q") + and holomorphic_on_extend_bounded: + "(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow> + (\<exists>B. eventually (\<lambda>z. norm(f z) \<le> B) (at \<xi>))" + (is "?P = ?R") +proof - + obtain \<delta> where "0 < \<delta>" and \<delta>: "ball \<xi> \<delta> \<subseteq> S" + using \<xi> mem_interior by blast + have "?R" if holg: "g holomorphic_on S" and gf: "\<And>z. z \<in> S - {\<xi>} \<Longrightarrow> g z = f z" for g + proof - + have *: "\<forall>\<^sub>F z in at \<xi>. dist (g z) (g \<xi>) < 1 \<longrightarrow> cmod (f z) \<le> cmod (g \<xi>) + 1" + apply (simp add: eventually_at) + apply (rule_tac x="\<delta>" in exI) + using \<delta> \<open>0 < \<delta>\<close> + apply (clarsimp simp:) + apply (drule_tac c=x in subsetD) + apply (simp add: dist_commute) + by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD) + have "continuous_on (interior S) g" + by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset) + then have "\<And>x. x \<in> interior S \<Longrightarrow> (g \<longlongrightarrow> g x) (at x)" + using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast + then have "(g \<longlongrightarrow> g \<xi>) (at \<xi>)" + by (simp add: \<xi>) + then show ?thesis + apply (rule_tac x="norm(g \<xi>) + 1" in exI) + apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) + done + qed + moreover have "?Q" if "\<forall>\<^sub>F z in at \<xi>. cmod (f z) \<le> B" for B + by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) + moreover have "?P" if "(\<lambda>z. (z - \<xi>) * f z) \<midarrow>\<xi>\<rightarrow> 0" + proof - + define h where [abs_def]: "h z = (z - \<xi>)^2 * f z" for z + have h0: "(h has_field_derivative 0) (at \<xi>)" + apply (simp add: h_def Derivative.DERIV_within_iff) + apply (rule Lim_transform_within [OF that, of 1]) + apply (auto simp: divide_simps power2_eq_square) + done + have holh: "h holomorphic_on S" + proof (simp add: holomorphic_on_def, clarify) + fix z assume "z \<in> S" + show "h field_differentiable at z within S" + proof (cases "z = \<xi>") + case True then show ?thesis + using field_differentiable_at_within field_differentiable_def h0 by blast + next + case False + then have "f field_differentiable at z within S" + using holomorphic_onD [OF holf, of z] \<open>z \<in> S\<close> + unfolding field_differentiable_def DERIV_within_iff + by (force intro: exI [where x="dist \<xi> z"] elim: Lim_transform_within_set [unfolded eventually_at]) + then show ?thesis + by (simp add: h_def power2_eq_square derivative_intros) + qed + qed + define g where [abs_def]: "g z = (if z = \<xi> then deriv h \<xi> else (h z - h \<xi>) / (z - \<xi>))" for z + have holg: "g holomorphic_on S" + unfolding g_def by (rule pole_lemma [OF holh \<xi>]) + show ?thesis + apply (rule_tac x="\<lambda>z. if z = \<xi> then deriv g \<xi> else (g z - g \<xi>)/(z - \<xi>)" in exI) + apply (rule conjI) + apply (rule pole_lemma [OF holg \<xi>]) + apply (auto simp: g_def power2_eq_square divide_simps) + using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square) + done + qed + ultimately show "?P = ?Q" and "?P = ?R" + by meson+ +qed + + +proposition pole_at_infinity: + assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \<longlongrightarrow> l) at_infinity" + obtains a n where "\<And>z. f z = (\<Sum>i\<le>n. a i * z^i)" +proof (cases "l = 0") + case False + with tendsto_inverse [OF lim] show ?thesis + apply (rule_tac a="(\<lambda>n. inverse l)" and n=0 in that) + apply (simp add: Liouville_weak [OF holf, of "inverse l"]) + done +next + case True + then have [simp]: "l = 0" . + show ?thesis + proof (cases "\<exists>r. 0 < r \<and> (\<forall>z \<in> ball 0 r - {0}. f(inverse z) \<noteq> 0)") + case True + then obtain r where "0 < r" and r: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> f(inverse z) \<noteq> 0" + by auto + have 1: "inverse \<circ> f \<circ> inverse holomorphic_on ball 0 r - {0}" + by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+ + have 2: "0 \<in> interior (ball 0 r)" + using \<open>0 < r\<close> by simp + have "\<exists>B. 0<B \<and> eventually (\<lambda>z. cmod ((inverse \<circ> f \<circ> inverse) z) \<le> B) (at 0)" + apply (rule exI [where x=1]) + apply (simp add:) + using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] + apply (rule eventually_mono) + apply (simp add: dist_norm) + done + with holomorphic_on_extend_bounded [OF 1 2] + obtain g where holg: "g holomorphic_on ball 0 r" + and geq: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> g z = (inverse \<circ> f \<circ> inverse) z" + by meson + have ifi0: "(inverse \<circ> f \<circ> inverse) \<midarrow>0\<rightarrow> 0" + using \<open>l = 0\<close> lim lim_at_infinity_0 by blast + have g2g0: "g \<midarrow>0\<rightarrow> g 0" + using \<open>0 < r\<close> centre_in_ball continuous_at continuous_on_eq_continuous_at holg + by (blast intro: holomorphic_on_imp_continuous_on) + have g2g1: "g \<midarrow>0\<rightarrow> 0" + apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) + using \<open>0 < r\<close> by (auto simp: geq) + have [simp]: "g 0 = 0" + by (rule tendsto_unique [OF _ g2g0 g2g1]) simp + have "ball 0 r - {0::complex} \<noteq> {}" + using \<open>0 < r\<close> + apply (clarsimp simp: ball_def dist_norm) + apply (drule_tac c="of_real r/2" in subsetD, auto) + done + then obtain w::complex where "w \<noteq> 0" and w: "norm w < r" by force + then have "g w \<noteq> 0" by (simp add: geq r) + obtain B n e where "0 < B" "0 < e" "e \<le> r" + and leg: "\<And>w. norm w < e \<Longrightarrow> B * cmod w ^ n \<le> cmod (g w)" + apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w]) + using \<open>0 < r\<close> w \<open>g w \<noteq> 0\<close> by (auto simp: ball_subset_ball_iff) + have "cmod (f z) \<le> cmod z ^ n / B" if "2/e \<le> cmod z" for z + proof - + have ize: "inverse z \<in> ball 0 e - {0}" using that \<open>0 < e\<close> + by (auto simp: norm_divide divide_simps algebra_simps) + then have [simp]: "z \<noteq> 0" and izr: "inverse z \<in> ball 0 r - {0}" using \<open>e \<le> r\<close> + by auto + then have [simp]: "f z \<noteq> 0" + using r [of "inverse z"] by simp + have [simp]: "f z = inverse (g (inverse z))" + using izr geq [of "inverse z"] by simp + show ?thesis using ize leg [of "inverse z"] \<open>0 < B\<close> \<open>0 < e\<close> + by (simp add: divide_simps norm_divide algebra_simps) + qed + then show ?thesis + apply (rule_tac a = "\<lambda>k. (deriv ^^ k) f 0 / (fact k)" and n=n in that) + apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf]) + apply (simp add:) + done + next + case False + then have fi0: "\<And>r. r > 0 \<Longrightarrow> \<exists>z\<in>ball 0 r - {0}. f (inverse z) = 0" + by simp + have fz0: "f z = 0" if "0 < r" and lt1: "\<And>x. x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> inverse (cmod (f (inverse x))) < 1" + for z r + proof - + have f0: "(f \<longlongrightarrow> 0) at_infinity" + proof - + have DIM_complex[intro]: "2 \<le> DIM(complex)" \<comment>\<open>should not be necessary!\<close> + by simp + have "continuous_on (inverse ` (ball 0 r - {0})) f" + using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast + then have "connected ((f \<circ> inverse) ` (ball 0 r - {0}))" + apply (intro connected_continuous_image continuous_intros) + apply (force intro: connected_punctured_ball)+ + done + then have "\<lbrakk>w \<noteq> 0; cmod w < r\<rbrakk> \<Longrightarrow> f (inverse w) = 0" for w + apply (rule disjE [OF connected_closedD [where A = "{0}" and B = "- ball 0 1"]], auto) + apply (metis (mono_tags, hide_lams) not_less_iff_gr_or_eq one_less_inverse lt1 zero_less_norm_iff) + using False \<open>0 < r\<close> apply fastforce + by (metis (no_types, hide_lams) Compl_iff IntI comp_apply empty_iff image_eqI insert_Diff_single insert_iff mem_ball_0 not_less_iff_gr_or_eq one_less_inverse that(2) zero_less_norm_iff) + then show ?thesis + apply (simp add: lim_at_infinity_0) + apply (rule Lim_eventually) + apply (simp add: eventually_at) + apply (rule_tac x=r in exI) + apply (simp add: \<open>0 < r\<close> dist_norm) + done + qed + obtain w where "w \<in> ball 0 r - {0}" and "f (inverse w) = 0" + using False \<open>0 < r\<close> by blast + then show ?thesis + by (auto simp: f0 Liouville_weak [OF holf, of 0]) + qed + show ?thesis + apply (rule that [of "\<lambda>n. 0" 0]) + using lim [unfolded lim_at_infinity_0] + apply (simp add: Lim_at dist_norm norm_inverse) + apply (drule_tac x=1 in spec) + using fz0 apply auto + done + qed +qed + + +subsection\<open>Entire proper functions are precisely the non-trivial polynomials\<close> + +proposition proper_map_polyfun: + fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}" + assumes "closed S" and "compact K" and c: "c i \<noteq> 0" "1 \<le> i" "i \<le> n" + shows "compact (S \<inter> {z. (\<Sum>i\<le>n. c i * z^i) \<in> K})" +proof - + obtain B where "B > 0" and B: "\<And>x. x \<in> K \<Longrightarrow> norm x \<le> B" + by (metis compact_imp_bounded \<open>compact K\<close> bounded_pos) + have *: "norm x \<le> b" + if "\<And>x. b \<le> norm x \<Longrightarrow> B + 1 \<le> norm (\<Sum>i\<le>n. c i * x ^ i)" + "(\<Sum>i\<le>n. c i * x ^ i) \<in> K" for b x + proof - + have "norm (\<Sum>i\<le>n. c i * x ^ i) \<le> B" + using B that by blast + moreover have "\<not> B + 1 \<le> B" + by simp + ultimately show "norm x \<le> b" + using that by (metis (no_types) less_eq_real_def not_less order_trans) + qed + have "bounded {z. (\<Sum>i\<le>n. c i * z ^ i) \<in> K}" + using polyfun_extremal [where c=c and B="B+1", OF c] + by (auto simp: bounded_pos eventually_at_infinity_pos *) + moreover have "closed {z. (\<Sum>i\<le>n. c i * z ^ i) \<in> K}" + apply (rule allI continuous_closed_preimage_univ continuous_intros)+ + using \<open>compact K\<close> compact_eq_bounded_closed by blast + ultimately show ?thesis + using closed_Int_compact [OF \<open>closed S\<close>] compact_eq_bounded_closed by blast +qed + +corollary proper_map_polyfun_univ: + fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}" + assumes "compact K" "c i \<noteq> 0" "1 \<le> i" "i \<le> n" + shows "compact ({z. (\<Sum>i\<le>n. c i * z^i) \<in> K})" +using proper_map_polyfun [of UNIV K c i n] assms by simp + + +proposition proper_map_polyfun_eq: + assumes "f holomorphic_on UNIV" + shows "(\<forall>k. compact k \<longrightarrow> compact {z. f z \<in> k}) \<longleftrightarrow> + (\<exists>c n. 0 < n \<and> (c n \<noteq> 0) \<and> f = (\<lambda>z. \<Sum>i\<le>n. c i * z^i))" + (is "?lhs = ?rhs") +proof + assume compf [rule_format]: ?lhs + have 2: "\<exists>k. 0 < k \<and> a k \<noteq> 0 \<and> f = (\<lambda>z. \<Sum>i \<le> k. a i * z ^ i)" + if "\<And>z. f z = (\<Sum>i\<le>n. a i * z ^ i)" for a n + proof (cases "\<forall>i\<le>n. 0<i \<longrightarrow> a i = 0") + case True + then have [simp]: "\<And>z. f z = a 0" + by (simp add: that setsum_atMost_shift) + have False using compf [of "{a 0}"] by simp + then show ?thesis .. + next + case False + then obtain k where k: "0 < k" "k\<le>n" "a k \<noteq> 0" by force + define m where "m = (GREATEST k. k\<le>n \<and> a k \<noteq> 0)" + have m: "m\<le>n \<and> a m \<noteq> 0" + unfolding m_def + apply (rule GreatestI [where b = "Suc n"]) + using k apply auto + done + have [simp]: "a i = 0" if "m < i" "i \<le> n" for i + using Greatest_le [where b = "Suc n" and P = "\<lambda>k. k\<le>n \<and> a k \<noteq> 0"] + using m_def not_le that by auto + have "k \<le> m" + unfolding m_def + apply (rule Greatest_le [where b = "Suc n"]) + using k apply auto + done + with k m show ?thesis + by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.setsum.mono_neutral_right) + qed + have "((inverse \<circ> f) \<longlongrightarrow> 0) at_infinity" + proof (rule Lim_at_infinityI) + fix e::real assume "0 < e" + with compf [of "cball 0 (inverse e)"] + show "\<exists>B. \<forall>x. B \<le> cmod x \<longrightarrow> dist ((inverse \<circ> f) x) 0 \<le> e" + apply (simp add:) + apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) + apply (rule_tac x="b+1" in exI) + apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one) + done + qed + then show ?rhs + apply (rule pole_at_infinity [OF assms]) + using 2 apply blast + done +next + assume ?rhs + then obtain c n where "0 < n" "c n \<noteq> 0" "f = (\<lambda>z. \<Sum>i\<le>n. c i * z ^ i)" by blast + then have "compact {z. f z \<in> k}" if "compact k" for k + by (auto intro: proper_map_polyfun_univ [OF that]) + then show ?lhs by blast +qed + + +subsection\<open>Relating invertibility and nonvanishing of derivative\<close> + +proposition has_complex_derivative_locally_injective: + assumes holf: "f holomorphic_on S" + and S: "\<xi> \<in> S" "open S" + and dnz: "deriv f \<xi> \<noteq> 0" + obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)" +proof - + have *: "\<exists>d>0. \<forall>x. dist \<xi> x < d \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) < e" if "e > 0" for e + proof - + have contdf: "continuous_on S (deriv f)" + by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \<open>open S\<close>) + obtain \<delta> where "\<delta>>0" and \<delta>: "\<And>x. \<lbrakk>x \<in> S; dist x \<xi> \<le> \<delta>\<rbrakk> \<Longrightarrow> cmod (deriv f x - deriv f \<xi>) \<le> e/2" + using continuous_onE [OF contdf \<open>\<xi> \<in> S\<close>, of "e/2"] \<open>0 < e\<close> + by (metis dist_complex_def half_gt_zero less_imp_le) + obtain \<epsilon> where "\<epsilon>>0" "ball \<xi> \<epsilon> \<subseteq> S" + by (metis openE [OF \<open>open S\<close> \<open>\<xi> \<in> S\<close>]) + with \<open>\<delta>>0\<close> have "\<exists>\<delta>>0. \<forall>x. dist \<xi> x < \<delta> \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) \<le> e/2" + apply (rule_tac x="min \<delta> \<epsilon>" in exI) + apply (intro conjI allI impI Operator_Norm.onorm_le) + apply (simp add:) + apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult) + apply (rule mult_right_mono [OF \<delta>]) + apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \<delta>) + done + with \<open>e>0\<close> show ?thesis by force + qed + have "inj (op * (deriv f \<xi>))" + using dnz by simp + then obtain g' where g': "linear g'" "g' \<circ> op * (deriv f \<xi>) = id" + using linear_injective_left_inverse [of "op * (deriv f \<xi>)"] + by (auto simp: linear_times) + show ?thesis + apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\<lambda>z h. deriv f z * h" and g' = g']) + using g' * + apply (simp_all add: linear_conv_bounded_linear that) + using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf + holomorphic_on_imp_differentiable_at \<open>open S\<close> apply blast + done +qed + + +proposition has_complex_derivative_locally_invertible: + assumes holf: "f holomorphic_on S" + and S: "\<xi> \<in> S" "open S" + and dnz: "deriv f \<xi> \<noteq> 0" + obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "open (f ` (ball \<xi> r))" "inj_on f (ball \<xi> r)" +proof - + obtain r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)" + by (blast intro: that has_complex_derivative_locally_injective [OF assms]) + then have \<xi>: "\<xi> \<in> ball \<xi> r" by simp + then have nc: "~ f constant_on ball \<xi> r" + using \<open>inj_on f (ball \<xi> r)\<close> injective_not_constant by fastforce + have holf': "f holomorphic_on ball \<xi> r" + using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast + have "open (f ` ball \<xi> r)" + apply (rule open_mapping_thm [OF holf']) + using nc apply auto + done + then show ?thesis + using \<open>0 < r\<close> \<open>ball \<xi> r \<subseteq> S\<close> \<open>inj_on f (ball \<xi> r)\<close> that by blast +qed + + +proposition holomorphic_injective_imp_regular: + assumes holf: "f holomorphic_on S" + and "open S" and injf: "inj_on f S" + and "\<xi> \<in> S" + shows "deriv f \<xi> \<noteq> 0" +proof - + obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE) + have holf': "f holomorphic_on ball \<xi> r" + using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast + show ?thesis + proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0") + case True + have fcon: "f w = f \<xi>" if "w \<in> ball \<xi> r" for w + apply (rule holomorphic_fun_eq_const_on_connected [OF holf']) + using True \<open>0 < r\<close> that by auto + have False + using fcon [of "\<xi> + r/2"] \<open>0 < r\<close> r injf unfolding inj_on_def + by (metis \<open>\<xi> \<in> S\<close> contra_subsetD dist_commute fcon mem_ball perfect_choose_dist) + then show ?thesis .. + next + case False + then obtain n0 where n0: "n0 > 0 \<and> (deriv ^^ n0) f \<xi> \<noteq> 0" by blast + define n where [abs_def]: "n = (LEAST n. n > 0 \<and> (deriv ^^ n) f \<xi> \<noteq> 0)" + have n_ne: "n > 0" "(deriv ^^ n) f \<xi> \<noteq> 0" + using def_LeastI [OF n_def n0] by auto + have n_min: "\<And>k. 0 < k \<Longrightarrow> k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0" + using def_Least_le [OF n_def] not_le by auto + obtain g \<delta> where "0 < \<delta>" + and holg: "g holomorphic_on ball \<xi> \<delta>" + and fd: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n" + and gnz: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> g w \<noteq> 0" + apply (rule holomorphic_factor_order_of_zero_strong [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> n_ne]) + apply (blast intro: n_min)+ + done + show ?thesis + proof (cases "n=1") + case True + with n_ne show ?thesis by auto + next + case False + have holgw: "(\<lambda>w. (w - \<xi>) * g w) holomorphic_on ball \<xi> (min r \<delta>)" + apply (rule holomorphic_intros)+ + using holg by (simp add: holomorphic_on_subset subset_ball) + have gd: "\<And>w. dist \<xi> w < \<delta> \<Longrightarrow> (g has_field_derivative deriv g w) (at w)" + using holg + by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH) + have *: "\<And>w. w \<in> ball \<xi> (min r \<delta>) + \<Longrightarrow> ((\<lambda>w. (w - \<xi>) * g w) has_field_derivative ((w - \<xi>) * deriv g w + g w)) + (at w)" + by (rule gd derivative_eq_intros | simp)+ + have [simp]: "deriv (\<lambda>w. (w - \<xi>) * g w) \<xi> \<noteq> 0" + using * [of \<xi>] \<open>0 < \<delta>\<close> \<open>0 < r\<close> by (simp add: DERIV_imp_deriv gnz) + obtain T where "\<xi> \<in> T" "open T" and Tsb: "T \<subseteq> ball \<xi> (min r \<delta>)" and oimT: "open ((\<lambda>w. (w - \<xi>) * g w) ` T)" + apply (rule has_complex_derivative_locally_invertible [OF holgw, of \<xi>]) + using \<open>0 < r\<close> \<open>0 < \<delta>\<close> + apply (simp_all add:) + by (meson Topology_Euclidean_Space.open_ball centre_in_ball) + define U where "U = (\<lambda>w. (w - \<xi>) * g w) ` T" + have "open U" by (metis oimT U_def) + have "0 \<in> U" + apply (auto simp: U_def) + apply (rule image_eqI [where x = \<xi>]) + apply (auto simp: \<open>\<xi> \<in> T\<close>) + done + then obtain \<epsilon> where "\<epsilon>>0" and \<epsilon>: "cball 0 \<epsilon> \<subseteq> U" + using \<open>open U\<close> open_contains_cball by blast + then have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> cball 0 \<epsilon>" + "\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> cball 0 \<epsilon>" + by (auto simp: norm_mult) + with \<epsilon> have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> U" + "\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> U" by blast+ + then obtain y0 y1 where "y0 \<in> T" and y0: "(y0 - \<xi>) * g y0 = \<epsilon> * exp(2 * of_real pi * \<i> * (0/n))" + and "y1 \<in> T" and y1: "(y1 - \<xi>) * g y1 = \<epsilon> * exp(2 * of_real pi * \<i> * (1/n))" + by (auto simp: U_def) + then have "y0 \<in> ball \<xi> \<delta>" "y1 \<in> ball \<xi> \<delta>" using Tsb by auto + moreover have "y0 \<noteq> y1" + using y0 y1 \<open>\<epsilon> > 0\<close> complex_root_unity_eq_1 [of n 1] \<open>n > 0\<close> False by auto + moreover have "T \<subseteq> S" + by (meson Tsb min.cobounded1 order_trans r subset_ball) + ultimately have False + using inj_onD [OF injf, of y0 y1] \<open>y0 \<in> T\<close> \<open>y1 \<in> T\<close> + using fd [of y0] fd [of y1] complex_root_unity [of n 1] + apply (simp add: y0 y1 power_mult_distrib) + apply (force simp: algebra_simps) + done + then show ?thesis .. + qed + qed +qed + + +text\<open>Hence a nice clean inverse function theorem\<close> + +proposition holomorphic_has_inverse: + assumes holf: "f holomorphic_on S" + and "open S" and injf: "inj_on f S" + obtains g where "g holomorphic_on (f ` S)" + "\<And>z. z \<in> S \<Longrightarrow> deriv f z * deriv g (f z) = 1" + "\<And>z. z \<in> S \<Longrightarrow> g(f z) = z" +proof - + have ofs: "open (f ` S)" + by (rule open_mapping_thm3 [OF assms]) + have contf: "continuous_on S f" + by (simp add: holf holomorphic_on_imp_continuous_on) + have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \<in> S" for z + proof - + have 1: "(f has_field_derivative deriv f z) (at z)" + using DERIV_deriv_iff_field_differentiable \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_on_imp_differentiable_at + by blast + have 2: "deriv f z \<noteq> 0" + using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast + show ?thesis + apply (rule has_complex_derivative_inverse_strong [OF 1 2 \<open>open S\<close> \<open>z \<in> S\<close>]) + apply (simp add: holf holomorphic_on_imp_continuous_on) + by (simp add: injf the_inv_into_f_f) + qed + show ?thesis + proof + show "the_inv_into S f holomorphic_on f ` S" + by (simp add: holomorphic_on_open ofs) (blast intro: *) + next + fix z assume "z \<in> S" + have "deriv f z \<noteq> 0" + using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast + then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" + using * [OF \<open>z \<in> S\<close>] by (simp add: DERIV_imp_deriv) + next + fix z assume "z \<in> S" + show "the_inv_into S f (f z) = z" + by (simp add: \<open>z \<in> S\<close> injf the_inv_into_f_f) + qed +qed + + +subsection\<open>The Schwarz Lemma\<close> + +lemma Schwarz1: + assumes holf: "f holomorphic_on S" + and contf: "continuous_on (closure S) f" + and S: "open S" "connected S" + and boS: "bounded S" + and "S \<noteq> {}" + obtains w where "w \<in> frontier S" + "\<And>z. z \<in> closure S \<Longrightarrow> norm (f z) \<le> norm (f w)" +proof - + have connf: "continuous_on (closure S) (norm o f)" + using contf continuous_on_compose continuous_on_norm_id by blast + have coc: "compact (closure S)" + by (simp add: \<open>bounded S\<close> bounded_closure compact_eq_bounded_closed) + then obtain x where x: "x \<in> closure S" and xmax: "\<And>z. z \<in> closure S \<Longrightarrow> norm(f z) \<le> norm(f x)" + apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]]) + using \<open>S \<noteq> {}\<close> apply auto + done + then show ?thesis + proof (cases "x \<in> frontier S") + case True + then show ?thesis using that xmax by blast + next + case False + then have "x \<in> S" + using \<open>open S\<close> frontier_def interior_eq x by auto + then have "f constant_on S" + apply (rule maximum_modulus_principle [OF holf S \<open>open S\<close> order_refl]) + using closure_subset apply (blast intro: xmax) + done + then have "f constant_on (closure S)" + by (rule constant_on_closureI [OF _ contf]) + then obtain c where c: "\<And>x. x \<in> closure S \<Longrightarrow> f x = c" + by (meson constant_on_def) + obtain w where "w \<in> frontier S" + by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) + then show ?thesis + by (simp add: c frontier_def that) + qed +qed + +lemma Schwarz2: + "\<lbrakk>f holomorphic_on ball 0 r; + 0 < s; ball w s \<subseteq> ball 0 r; + \<And>z. norm (w-z) < s \<Longrightarrow> norm(f z) \<le> norm(f w)\<rbrakk> + \<Longrightarrow> f constant_on ball 0 r" +by (rule maximum_modulus_principle [where U = "ball w s" and \<xi> = w]) (simp_all add: dist_norm) + +lemma Schwarz3: + assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" + obtains h where "h holomorphic_on (ball 0 r)" and "\<And>z. norm z < r \<Longrightarrow> f z = z * (h z)" and "deriv f 0 = h 0" +proof - + define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z + have d0: "deriv f 0 = h 0" + by (simp add: h_def) + moreover have "h holomorphic_on (ball 0 r)" + by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) + moreover have "norm z < r \<Longrightarrow> f z = z * h z" for z + by (simp add: h_def) + ultimately show ?thesis + using that by blast +qed + + +proposition Schwarz_Lemma: + assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" + and no: "\<And>z. norm z < 1 \<Longrightarrow> norm (f z) < 1" + and \<xi>: "norm \<xi> < 1" + shows "norm (f \<xi>) \<le> norm \<xi>" and "norm(deriv f 0) \<le> 1" + and "((\<exists>z. norm z < 1 \<and> z \<noteq> 0 \<and> norm(f z) = norm z) \<or> norm(deriv f 0) = 1) + \<Longrightarrow> \<exists>\<alpha>. (\<forall>z. norm z < 1 \<longrightarrow> f z = \<alpha> * z) \<and> norm \<alpha> = 1" (is "?P \<Longrightarrow> ?Q") +proof - + obtain h where holh: "h holomorphic_on (ball 0 1)" + and fz_eq: "\<And>z. norm z < 1 \<Longrightarrow> f z = z * (h z)" and df0: "deriv f 0 = h 0" + by (rule Schwarz3 [OF holf]) auto + have noh_le: "norm (h z) \<le> 1" if z: "norm z < 1" for z + proof - + have "norm (h z) < a" if a: "1 < a" for a + proof - + have "max (inverse a) (norm z) < 1" + using z a by (simp_all add: inverse_less_1_iff) + then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" + using Rats_dense_in_real by blast + then have nzr: "norm z < r" and ira: "inverse r < a" + using z a less_imp_inverse_less by force+ + then have "0 < r" + by (meson norm_not_less_zero not_le order.strict_trans2) + have holh': "h holomorphic_on ball 0 r" + by (meson holh \<open>r < 1\<close> holomorphic_on_subset less_eq_real_def subset_ball) + have conth': "continuous_on (cball 0 r) h" + by (meson \<open>r < 1\<close> dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) + obtain w where w: "norm w = r" and lenw: "\<And>z. norm z < r \<Longrightarrow> norm(h z) \<le> norm(h w)" + apply (rule Schwarz1 [OF holh']) using conth' \<open>0 < r\<close> by auto + have "h w = f w / w" using fz_eq \<open>r < 1\<close> nzr w by auto + then have "cmod (h z) < inverse r" + by (metis \<open>0 < r\<close> \<open>r < 1\<close> divide_strict_right_mono inverse_eq_divide + le_less_trans lenw no norm_divide nzr w) + then show ?thesis using ira by linarith + qed + then show "norm (h z) \<le> 1" + using not_le by blast + qed + show "cmod (f \<xi>) \<le> cmod \<xi>" + proof (cases "\<xi> = 0") + case True then show ?thesis by auto + next + case False + then show ?thesis + by (simp add: noh_le fz_eq \<xi> mult_left_le norm_mult) + qed + show no_df0: "norm(deriv f 0) \<le> 1" + by (simp add: \<open>\<And>z. cmod z < 1 \<Longrightarrow> cmod (h z) \<le> 1\<close> df0) + show "?Q" if "?P" + using that + proof + assume "\<exists>z. cmod z < 1 \<and> z \<noteq> 0 \<and> cmod (f z) = cmod z" + then obtain \<gamma> where \<gamma>: "cmod \<gamma> < 1" "\<gamma> \<noteq> 0" "cmod (f \<gamma>) = cmod \<gamma>" by blast + then have [simp]: "norm (h \<gamma>) = 1" + by (simp add: fz_eq norm_mult) + have "ball \<gamma> (1 - cmod \<gamma>) \<subseteq> ball 0 1" + by (simp add: ball_subset_ball_iff) + moreover have "\<And>z. cmod (\<gamma> - z) < 1 - cmod \<gamma> \<Longrightarrow> cmod (h z) \<le> cmod (h \<gamma>)" + apply (simp add: algebra_simps) + by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) + ultimately obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c" + using Schwarz2 [OF holh, of "1 - norm \<gamma>" \<gamma>, unfolded constant_on_def] \<gamma> by auto + then have "norm c = 1" + using \<gamma> by force + with c show ?thesis + using fz_eq by auto + next + assume [simp]: "cmod (deriv f 0) = 1" + then obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c" + using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le + by auto + moreover have "norm c = 1" using df0 c by auto + ultimately show ?thesis + using fz_eq by auto + qed +qed + +subsection\<open>The Schwarz reflection principle\<close> + +lemma hol_pal_lem0: + assumes "d \<bullet> a \<le> k" "k \<le> d \<bullet> b" + obtains c where + "c \<in> closed_segment a b" "d \<bullet> c = k" + "\<And>z. z \<in> closed_segment a c \<Longrightarrow> d \<bullet> z \<le> k" + "\<And>z. z \<in> closed_segment c b \<Longrightarrow> k \<le> d \<bullet> z" +proof - + obtain c where cin: "c \<in> closed_segment a b" and keq: "k = d \<bullet> c" + using connected_ivt_hyperplane [of "closed_segment a b" a b d k] + by (auto simp: assms) + have "closed_segment a c \<subseteq> {z. d \<bullet> z \<le> k}" "closed_segment c b \<subseteq> {z. k \<le> d \<bullet> z}" + unfolding segment_convex_hull using assms keq + by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) + then show ?thesis using cin that by fastforce +qed + +lemma hol_pal_lem1: + assumes "convex S" "open S" + and abc: "a \<in> S" "b \<in> S" "c \<in> S" + "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k" "d \<bullet> c \<le> k" + and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof - + have "interior (convex hull {a, b, c}) \<subseteq> interior(S \<inter> {x. d \<bullet> x \<le> k})" + apply (rule interior_mono) + apply (rule hull_minimal) + apply (simp add: abc lek) + apply (rule convex_Int [OF \<open>convex S\<close> convex_halfspace_le]) + done + also have "... \<subseteq> {z \<in> S. d \<bullet> z < k}" + by (force simp: interior_open [OF \<open>open S\<close>] \<open>d \<noteq> 0\<close>) + finally have *: "interior (convex hull {a, b, c}) \<subseteq> {z \<in> S. d \<bullet> z < k}" . + have "continuous_on (convex hull {a,b,c}) f" + using \<open>convex S\<close> contf abc continuous_on_subset subset_hull + by fastforce + moreover have "f holomorphic_on interior (convex hull {a,b,c})" + by (rule holomorphic_on_subset [OF holf1 *]) + ultimately show ?thesis + using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 + by blast +qed + +lemma hol_pal_lem2: + assumes S: "convex S" "open S" + and abc: "a \<in> S" "b \<in> S" "c \<in> S" + and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k" + and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" + and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "d \<bullet> c \<le> k") + case True show ?thesis + by (rule hol_pal_lem1 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 contf]) +next + case False + then have "d \<bullet> c > k" by force + obtain a' where a': "a' \<in> closed_segment b c" and "d \<bullet> a' = k" + and ba': "\<And>z. z \<in> closed_segment b a' \<Longrightarrow> d \<bullet> z \<le> k" + and a'c: "\<And>z. z \<in> closed_segment a' c \<Longrightarrow> k \<le> d \<bullet> z" + apply (rule hol_pal_lem0 [of d b k c, OF \<open>d \<bullet> b \<le> k\<close>]) + using False by auto + obtain b' where b': "b' \<in> closed_segment a c" and "d \<bullet> b' = k" + and ab': "\<And>z. z \<in> closed_segment a b' \<Longrightarrow> d \<bullet> z \<le> k" + and b'c: "\<And>z. z \<in> closed_segment b' c \<Longrightarrow> k \<le> d \<bullet> z" + apply (rule hol_pal_lem0 [of d a k c, OF \<open>d \<bullet> a \<le> k\<close>]) + using False by auto + have a'b': "a' \<in> S \<and> b' \<in> S" + using a' abc b' convex_contains_segment \<open>convex S\<close> by auto + have "continuous_on (closed_segment c a) f" + by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>) + then have 1: "contour_integral (linepath c a) f = + contour_integral (linepath c b') f + contour_integral (linepath b' a) f" + apply (rule contour_integral_split_linepath) + using b' by (simp add: closed_segment_commute) + have "continuous_on (closed_segment b c) f" + by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>) + then have 2: "contour_integral (linepath b c) f = + contour_integral (linepath b a') f + contour_integral (linepath a' c) f" + by (rule contour_integral_split_linepath [OF _ a']) + have 3: "contour_integral (reversepath (linepath b' a')) f = + - contour_integral (linepath b' a') f" + by (rule contour_integral_reversepath [OF valid_path_linepath]) + have fcd_le: "f field_differentiable at x" + if "x \<in> interior S \<and> x \<in> interior {x. d \<bullet> x \<le> k}" for x + proof - + have "f holomorphic_on S \<inter> {c. d \<bullet> c < k}" + by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) + then have "\<exists>C D. x \<in> interior C \<inter> interior D \<and> f holomorphic_on interior C \<inter> interior D" + using that + by (metis Collect_mem_eq Int_Collect \<open>d \<noteq> 0\<close> interior_halfspace_le interior_open \<open>open S\<close>) + then show "f field_differentiable at x" + by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) + qed + have ab_le: "\<And>x. x \<in> closed_segment a b \<Longrightarrow> d \<bullet> x \<le> k" + proof - + fix x :: complex + assume "x \<in> closed_segment a b" + then have "\<And>C. x \<in> C \<or> b \<notin> C \<or> a \<notin> C \<or> \<not> convex C" + by (meson contra_subsetD convex_contains_segment) + then show "d \<bullet> x \<le> k" + by (metis lek convex_halfspace_le mem_Collect_eq) + qed + have "continuous_on (S \<inter> {x. d \<bullet> x \<le> k}) f" using contf + by (simp add: continuous_on_subset) + then have "(f has_contour_integral 0) + (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" + apply (rule Cauchy_theorem_convex [where k = "{}"]) + apply (simp_all add: path_image_join convex_Int convex_halfspace_le \<open>convex S\<close> fcd_le ab_le + closed_segment_subset abc a'b' ba') + by (metis \<open>d \<bullet> a' = k\<close> \<open>d \<bullet> b' = k\<close> convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl) + then have 4: "contour_integral (linepath a b) f + + contour_integral (linepath b a') f + + contour_integral (linepath a' b') f + + contour_integral (linepath b' a) f = 0" + by (rule has_chain_integral_chain_integral4) + have fcd_ge: "f field_differentiable at x" + if "x \<in> interior S \<and> x \<in> interior {x. k \<le> d \<bullet> x}" for x + proof - + have f2: "f holomorphic_on S \<inter> {c. k < d \<bullet> c}" + by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) + have f3: "interior S = S" + by (simp add: interior_open \<open>open S\<close>) + then have "x \<in> S \<inter> interior {c. k \<le> d \<bullet> c}" + using that by simp + then show "f field_differentiable at x" + using f3 f2 unfolding holomorphic_on_def + by (metis (no_types) \<open>d \<noteq> 0\<close> at_within_interior interior_Int interior_halfspace_ge interior_interior) + qed + have "continuous_on (S \<inter> {x. k \<le> d \<bullet> x}) f" using contf + by (simp add: continuous_on_subset) + then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" + apply (rule Cauchy_theorem_convex [where k = "{}"]) + apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \<open>convex S\<close> + fcd_ge closed_segment_subset abc a'b' a'c) + by (metis \<open>d \<bullet> a' = k\<close> b'c closed_segment_commute convex_contains_segment + convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl) + then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" + by (rule has_chain_integral_chain_integral3) + show ?thesis + using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) +qed + +lemma hol_pal_lem3: + assumes S: "convex S" "open S" + and abc: "a \<in> S" "b \<in> S" "c \<in> S" + and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" + and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" + and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "d \<bullet> b \<le> k") + case True show ?thesis + by (rule hol_pal_lem2 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 holf2 contf]) +next + case False + show ?thesis + proof (cases "d \<bullet> c \<le> k") + case True + have "contour_integral (linepath c a) f + + contour_integral (linepath a b) f + + contour_integral (linepath b c) f = 0" + by (rule hol_pal_lem2 [OF S \<open>c \<in> S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close> \<open>d \<noteq> 0\<close> \<open>d \<bullet> c \<le> k\<close> lek holf1 holf2 contf]) + then show ?thesis + by (simp add: algebra_simps) + next + case False + have "contour_integral (linepath b c) f + + contour_integral (linepath c a) f + + contour_integral (linepath a b) f = 0" + apply (rule hol_pal_lem2 [OF S \<open>b \<in> S\<close> \<open>c \<in> S\<close> \<open>a \<in> S\<close>, of "-d" "-k"]) + using \<open>d \<noteq> 0\<close> \<open>\<not> d \<bullet> b \<le> k\<close> False by (simp_all add: holf1 holf2 contf) + then show ?thesis + by (simp add: algebra_simps) + qed +qed + +lemma hol_pal_lem4: + assumes S: "convex S" "open S" + and abc: "a \<in> S" "b \<in> S" "c \<in> S" and "d \<noteq> 0" + and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" + and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "d \<bullet> a \<le> k") + case True show ?thesis + by (rule hol_pal_lem3 [OF S abc \<open>d \<noteq> 0\<close> True holf1 holf2 contf]) +next + case False + show ?thesis + apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"]) + using \<open>d \<noteq> 0\<close> False by (simp_all add: holf1 holf2 contf) +qed + +proposition holomorphic_on_paste_across_line: + assumes S: "open S" and "d \<noteq> 0" + and holf1: "f holomorphic_on (S \<inter> {z. d \<bullet> z < k})" + and holf2: "f holomorphic_on (S \<inter> {z. k < d \<bullet> z})" + and contf: "continuous_on S f" + shows "f holomorphic_on S" +proof - + have *: "\<exists>t. open t \<and> p \<in> t \<and> continuous_on t f \<and> + (\<forall>a b c. convex hull {a, b, c} \<subseteq> t \<longrightarrow> + contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + if "p \<in> S" for p + proof - + obtain e where "e>0" and e: "ball p e \<subseteq> S" + using \<open>p \<in> S\<close> openE S by blast + then have "continuous_on (ball p e) f" + using contf continuous_on_subset by blast + moreover have "f holomorphic_on {z. dist p z < e \<and> d \<bullet> z < k}" + apply (rule holomorphic_on_subset [OF holf1]) + using e by auto + moreover have "f holomorphic_on {z. dist p z < e \<and> k < d \<bullet> z}" + apply (rule holomorphic_on_subset [OF holf2]) + using e by auto + ultimately show ?thesis + apply (rule_tac x="ball p e" in exI) + using \<open>e > 0\<close> e \<open>d \<noteq> 0\<close> + apply (simp add:, clarify) + apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k]) + apply (auto simp: subset_hull) + done + qed + show ?thesis + by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) +qed + +proposition Schwarz_reflection: + assumes "open S" and cnjs: "cnj ` S \<subseteq> S" + and holf: "f holomorphic_on (S \<inter> {z. 0 < Im z})" + and contf: "continuous_on (S \<inter> {z. 0 \<le> Im z}) f" + and f: "\<And>z. \<lbrakk>z \<in> S; z \<in> \<real>\<rbrakk> \<Longrightarrow> (f z) \<in> \<real>" + shows "(\<lambda>z. if 0 \<le> Im z then f z else cnj(f(cnj z))) holomorphic_on S" +proof - + have 1: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. 0 < Im z})" + by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) + have cont_cfc: "continuous_on (S \<inter> {z. Im z \<le> 0}) (cnj o f o cnj)" + apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) + using cnjs apply auto + done + have "cnj \<circ> f \<circ> cnj field_differentiable at x within S \<inter> {z. Im z < 0}" + if "x \<in> S" "Im x < 0" "f field_differentiable at (cnj x) within S \<inter> {z. 0 < Im z}" for x + using that + apply (simp add: field_differentiable_def Derivative.DERIV_within_iff Lim_within dist_norm, clarify) + apply (rule_tac x="cnj f'" in exI) + apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) + apply (drule_tac x="cnj xa" in bspec) + using cnjs apply force + apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) + done + then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \<inter> {z. Im z < 0})" + using holf cnjs + by (force simp: holomorphic_on_def) + have 2: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. Im z < 0})" + apply (rule iffD1 [OF holomorphic_cong [OF refl]]) + using hol_cfc by auto + have [simp]: "(S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0}) = S" + by force + have "continuous_on ((S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0})) + (\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))" + apply (rule continuous_on_cases_local) + using cont_cfc contf + apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) + using f Reals_cnj_iff complex_is_Real_iff apply auto + done + then have 3: "continuous_on S (\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))" + by force + show ?thesis + apply (rule holomorphic_on_paste_across_line [OF \<open>open S\<close>, of "- \<i>" _ 0]) + using 1 2 3 + apply auto + done +qed + +subsection\<open>Bloch's theorem\<close> + +lemma Bloch_lemma_0: + assumes holf: "f holomorphic_on cball 0 r" and "0 < r" + and [simp]: "f 0 = 0" + and le: "\<And>z. norm z < r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f 0)" + shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \<subseteq> f ` ball 0 r" +proof - + have "sqrt 2 < 3/2" + by (rule real_less_lsqrt) (auto simp: power2_eq_square) + then have sq3: "0 < 3 - 2 * sqrt 2" by simp + show ?thesis + proof (cases "deriv f 0 = 0") + case True then show ?thesis by simp + next + case False + define C where "C = 2 * norm(deriv f 0)" + have "0 < C" using False by (simp add: C_def) + have holf': "f holomorphic_on ball 0 r" using holf + using ball_subset_cball holomorphic_on_subset by blast + then have holdf': "deriv f holomorphic_on ball 0 r" + by (rule holomorphic_deriv [OF _ open_ball]) + have "Le1": "norm(deriv f z - deriv f 0) \<le> norm z / (r - norm z) * C" + if "norm z < r" for z + proof - + have T1: "norm(deriv f z - deriv f 0) \<le> norm z / (R - norm z) * C" + if R: "norm z < R" "R < r" for R + proof - + have "0 < R" using R + by (metis less_trans norm_zero zero_less_norm_iff) + have df_le: "\<And>x. norm x < r \<Longrightarrow> norm (deriv f x) \<le> C" + using le by (simp add: C_def) + have hol_df: "deriv f holomorphic_on cball 0 R" + apply (rule holomorphic_on_subset) using R holdf' by auto + have *: "((\<lambda>w. deriv f w / (w - z)) has_contour_integral 2 * pi * \<i> * deriv f z) (circlepath 0 R)" + if "norm z < R" for z + using \<open>0 < R\<close> that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] + by (force simp: winding_number_circlepath) + have **: "((\<lambda>x. deriv f x / (x - z) - deriv f x / x) has_contour_integral + of_real (2 * pi) * \<i> * (deriv f z - deriv f 0)) + (circlepath 0 R)" + using has_contour_integral_diff [OF * [of z] * [of 0]] \<open>0 < R\<close> that + by (simp add: algebra_simps) + have [simp]: "\<And>x. norm x = R \<Longrightarrow> x \<noteq> z" using that(1) by blast + have "norm (deriv f x / (x - z) - deriv f x / x) + \<le> C * norm z / (R * (R - norm z))" + if "norm x = R" for x + proof - + have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = + norm (deriv f x) * norm z" + by (simp add: norm_mult right_diff_distrib') + show ?thesis + using \<open>0 < R\<close> \<open>0 < C\<close> R that + apply (simp add: norm_mult norm_divide divide_simps) + using df_le norm_triangle_ineq2 \<open>0 < C\<close> apply (auto intro!: mult_mono) + done + qed + then show ?thesis + using has_contour_integral_bound_circlepath + [OF **, of "C * norm z/(R*(R - norm z))"] + \<open>0 < R\<close> \<open>0 < C\<close> R + apply (simp add: norm_mult norm_divide) + apply (simp add: divide_simps mult.commute) + done + qed + obtain r' where r': "norm z < r'" "r' < r" + using Rats_dense_in_real [of "norm z" r] \<open>norm z < r\<close> by blast + then have [simp]: "closure {r'<..<r} = {r'..r}" by simp + show ?thesis + apply (rule continuous_ge_on_closure + [where f = "\<lambda>r. norm z / (r - norm z) * C" and s = "{r'<..<r}", + OF _ _ T1]) + apply (intro continuous_intros) + using that r' + apply (auto simp: not_le) + done + qed + have "*": "(norm z - norm z^2/(r - norm z)) * norm(deriv f 0) \<le> norm(f z)" + if r: "norm z < r" for z + proof - + have 1: "\<And>x. x \<in> ball 0 r \<Longrightarrow> + ((\<lambda>z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) + (at x within ball 0 r)" + by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ + have 2: "closed_segment 0 z \<subseteq> ball 0 r" + by (metis \<open>0 < r\<close> convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) + have 3: "(\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" + apply (rule integrable_on_cmult_right [where 'b=real, simplified]) + apply (rule integrable_on_cdivide [where 'b=real, simplified]) + apply (rule integrable_on_cmult_left [where 'b=real, simplified]) + apply (rule ident_integrable_on) + done + have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm z * norm z * x * C / (r - norm z)" + if x: "0 \<le> x" "x \<le> 1" for x + proof - + have [simp]: "x * norm z < r" + using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) + have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \<le> norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" + apply (rule Le1) using r x \<open>0 < r\<close> by simp + also have "... \<le> norm (x *\<^sub>R z) / (r - norm z) * C" + using r x \<open>0 < r\<close> + apply (simp add: divide_simps) + by (simp add: \<open>0 < C\<close> mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) + finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm (x *\<^sub>R z) / (r - norm z) * C * norm z" + by (rule mult_right_mono) simp + with x show ?thesis by (simp add: algebra_simps) + qed + have le_norm: "abc \<le> norm d - e \<Longrightarrow> norm(f - d) \<le> e \<Longrightarrow> abc \<le> norm f" for abc d e and f::complex + by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) + have "norm (integral {0..1} (\<lambda>x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) + \<le> integral {0..1} (\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C)" + apply (rule integral_norm_bound_integral) + using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 + apply (simp add: has_contour_integral_linepath has_integral_integrable_integral) + apply (rule 3) + apply (simp add: norm_mult power2_eq_square 4) + done + then have int_le: "norm (f z - deriv f 0 * z) \<le> (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" + using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 + apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) + done + show ?thesis + apply (rule le_norm [OF _ int_le]) + using \<open>norm z < r\<close> + apply (simp add: power2_eq_square divide_simps C_def norm_mult) + proof - + have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \<le> norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" + by (simp add: linordered_field_class.sign_simps(38)) + then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \<le> norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" + by (simp add: linordered_field_class.sign_simps(38) mult.commute mult.left_commute) + qed + qed + have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" + by (auto simp: sqrt2_less_2) + have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" + apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) + apply (subst closure_ball) + using \<open>0 < r\<close> mult_pos_pos sq201 + apply (auto simp: cball_subset_cball_iff) + done + have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" + apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force) + using \<open>0 < r\<close> mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff) + using False \<open>0 < r\<close> centre_in_ball holf' holomorphic_nonconstant by blast + have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = + ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" + by simp + also have "... \<subseteq> f ` ball 0 ((1 - sqrt 2 / 2) * r)" + proof - + have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \<le> norm (f z)" + if "norm z = (1 - sqrt 2 / 2) * r" for z + apply (rule order_trans [OF _ *]) + using \<open>0 < r\<close> + apply (simp_all add: field_simps power2_eq_square that) + apply (simp add: mult.assoc [symmetric]) + done + show ?thesis + apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) + using \<open>0 < r\<close> sq201 3 apply simp_all + using C_def \<open>0 < C\<close> sq3 apply force + done + qed + also have "... \<subseteq> f ` ball 0 r" + apply (rule image_subsetI [OF imageI], simp) + apply (erule less_le_trans) + using \<open>0 < r\<close> apply (auto simp: field_simps) + done + finally show ?thesis . + qed +qed + +lemma Bloch_lemma: + assumes holf: "f holomorphic_on cball a r" and "0 < r" + and le: "\<And>z. z \<in> ball a r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f a)" + shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \<subseteq> f ` ball a r" +proof - + have fz: "(\<lambda>z. f (a + z)) = f o (\<lambda>z. (a + z))" + by (simp add: o_def) + have hol0: "(\<lambda>z. f (a + z)) holomorphic_on cball 0 r" + unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ + then have [simp]: "\<And>x. norm x < r \<Longrightarrow> (\<lambda>z. f (a + z)) field_differentiable at x" + by (metis Topology_Euclidean_Space.open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) + have [simp]: "\<And>z. norm z < r \<Longrightarrow> f field_differentiable at (a + z)" + by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute) + then have [simp]: "f field_differentiable at a" + by (metis add.comm_neutral \<open>0 < r\<close> norm_eq_zero) + have hol1: "(\<lambda>z. f (a + z) - f a) holomorphic_on cball 0 r" + by (intro holomorphic_intros hol0) + then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\<lambda>z. f (a + z) - f a) 0)) + \<subseteq> (\<lambda>z. f (a + z) - f a) ` ball 0 r" + apply (rule Bloch_lemma_0) + apply (simp_all add: \<open>0 < r\<close>) + apply (simp add: fz complex_derivative_chain) + apply (simp add: dist_norm le) + done + then show ?thesis + apply clarify + apply (drule_tac c="x - f a" in subsetD) + apply (force simp: fz \<open>0 < r\<close> dist_norm complex_derivative_chain field_differentiable_compose)+ + done +qed + +proposition Bloch_unit: + assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" + obtains b r where "1/12 < r" "ball b r \<subseteq> f ` (ball a 1)" +proof - + define r :: real where "r = 249/256" + have "0 < r" "r < 1" by (auto simp: r_def) + define g where "g z = deriv f z * of_real(r - norm(z - a))" for z + have "deriv f holomorphic_on ball a 1" + by (rule holomorphic_deriv [OF holf open_ball]) + then have "continuous_on (ball a 1) (deriv f)" + using holomorphic_on_imp_continuous_on by blast + then have "continuous_on (cball a r) (deriv f)" + by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \<open>r < 1\<close>) + then have "continuous_on (cball a r) g" + by (simp add: g_def continuous_intros) + then have 1: "compact (g ` cball a r)" + by (rule compact_continuous_image [OF _ compact_cball]) + have 2: "g ` cball a r \<noteq> {}" + using \<open>r > 0\<close> by auto + obtain p where pr: "p \<in> cball a r" + and pge: "\<And>y. y \<in> cball a r \<Longrightarrow> norm (g y) \<le> norm (g p)" + using distance_attains_sup [OF 1 2, of 0] by force + define t where "t = (r - norm(p - a)) / 2" + have "norm (p - a) \<noteq> r" + using pge [of a] \<open>r > 0\<close> by (auto simp: g_def norm_mult) + then have "norm (p - a) < r" using pr + by (simp add: norm_minus_commute dist_norm) + then have "0 < t" + by (simp add: t_def) + have cpt: "cball p t \<subseteq> ball a r" + using \<open>0 < t\<close> by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) + have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \<le> norm (deriv f p)" + if "y \<in> cball a r" for y + proof - + have [simp]: "norm (y - a) \<le> r" + using that by (simp add: dist_norm norm_minus_commute) + have "norm (g y) \<le> norm (g p)" + using pge [OF that] by simp + then have "norm (deriv f y) * abs (r - norm (y - a)) \<le> norm (deriv f p) * abs (r - norm (p - a))" + by (simp only: dist_norm g_def norm_mult norm_of_real) + with that \<open>norm (p - a) < r\<close> show ?thesis + by (simp add: dist_norm divide_simps) + qed + have le_norm_dfp: "r / (r - norm (p - a)) \<le> norm (deriv f p)" + using gen_le_dfp [of a] \<open>r > 0\<close> by auto + have 1: "f holomorphic_on cball p t" + apply (rule holomorphic_on_subset [OF holf]) + using cpt \<open>r < 1\<close> order_subst1 subset_ball by auto + have 2: "norm (deriv f z) \<le> 2 * norm (deriv f p)" if "z \<in> ball p t" for z + proof - + have z: "z \<in> cball a r" + by (meson ball_subset_cball subsetD cpt that) + then have "norm(z - a) < r" + by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) + have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \<le> norm (deriv f p)" + using gen_le_dfp [OF z] by simp + with \<open>norm (z - a) < r\<close> \<open>norm (p - a) < r\<close> + have "norm (deriv f z) \<le> (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" + by (simp add: field_simps) + also have "... \<le> 2 * norm (deriv f p)" + apply (rule mult_right_mono) + using that \<open>norm (p - a) < r\<close> \<open>norm(z - a) < r\<close> + apply (simp_all add: field_simps t_def dist_norm [symmetric]) + using dist_triangle3 [of z a p] by linarith + finally show ?thesis . + qed + have sqrt2: "sqrt 2 < 2113/1494" + by (rule real_less_lsqrt) (auto simp: power2_eq_square) + then have sq3: "0 < 3 - 2 * sqrt 2" by simp + have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" + using sq3 sqrt2 by (auto simp: field_simps r_def) + also have "... \<le> cmod (deriv f p) * (r - cmod (p - a))" + using \<open>norm (p - a) < r\<close> le_norm_dfp by (simp add: pos_divide_le_eq) + finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" + using pos_divide_less_eq half_gt_zero_iff sq3 by blast + then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" + using sq3 by (simp add: mult.commute t_def) + have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball p t" + by (rule Bloch_lemma [OF 1 \<open>0 < t\<close> 2]) + also have "... \<subseteq> f ` ball a 1" + apply (rule image_mono) + apply (rule order_trans [OF ball_subset_cball]) + apply (rule order_trans [OF cpt]) + using \<open>0 < t\<close> \<open>r < 1\<close> apply (simp add: ball_subset_ball_iff dist_norm) + done + finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball a 1" . + with ** show ?thesis + by (rule that) +qed + + +theorem Bloch: + assumes holf: "f holomorphic_on ball a r" and "0 < r" + and r': "r' \<le> r * norm (deriv f a) / 12" + obtains b where "ball b r' \<subseteq> f ` (ball a r)" +proof (cases "deriv f a = 0") + case True with r' show ?thesis + using ball_eq_empty that by fastforce +next + case False + define C where "C = deriv f a" + have "0 < norm C" using False by (simp add: C_def) + have dfa: "f field_differentiable at a" + apply (rule holomorphic_on_imp_differentiable_at [OF holf]) + using \<open>0 < r\<close> by auto + have fo: "(\<lambda>z. f (a + of_real r * z)) = f o (\<lambda>z. (a + of_real r * z))" + by (simp add: o_def) + have holf': "f holomorphic_on (\<lambda>z. a + complex_of_real r * z) ` ball 0 1" + apply (rule holomorphic_on_subset [OF holf]) + using \<open>0 < r\<close> apply (force simp: dist_norm norm_mult) + done + have 1: "(\<lambda>z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" + apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ + using \<open>0 < r\<close> by (simp add: C_def False) + have "((\<lambda>z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative + (deriv f (a + of_real r * z) / C)) (at z)" + if "norm z < 1" for z + proof - + have *: "((\<lambda>x. f (a + of_real r * x)) has_field_derivative + (deriv f (a + of_real r * z) * of_real r)) (at z)" + apply (simp add: fo) + apply (rule DERIV_chain [OF field_differentiable_derivI]) + apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) + using \<open>0 < r\<close> apply (simp add: dist_norm norm_mult that) + apply (rule derivative_eq_intros | simp)+ + done + show ?thesis + apply (rule derivative_eq_intros * | simp)+ + using \<open>0 < r\<close> by (auto simp: C_def False) + qed + have 2: "deriv (\<lambda>z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" + apply (subst deriv_cdivide_right) + apply (simp add: field_differentiable_def fo) + apply (rule exI) + apply (rule DERIV_chain [OF field_differentiable_derivI]) + apply (simp add: dfa) + apply (rule derivative_eq_intros | simp add: C_def False fo)+ + using \<open>0 < r\<close> + apply (simp add: C_def False fo) + apply (simp add: derivative_intros dfa complex_derivative_chain) + done + have sb1: "op * (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1 + \<subseteq> f ` ball a r" + using \<open>0 < r\<close> by (auto simp: dist_norm norm_mult C_def False) + have sb2: "ball (C * r * b) r' \<subseteq> op * (C * r) ` ball b t" + if "1 / 12 < t" for b t + proof - + have *: "r * cmod (deriv f a) / 12 \<le> r * (t * cmod (deriv f a))" + using that \<open>0 < r\<close> less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right + by auto + show ?thesis + apply clarify + apply (rule_tac x="x / (C * r)" in image_eqI) + using \<open>0 < r\<close> + apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) + apply (erule less_le_trans) + apply (rule order_trans [OF r' *]) + done + qed + show ?thesis + apply (rule Bloch_unit [OF 1 2]) + apply (rename_tac t) + apply (rule_tac b="(C * of_real r) * b" in that) + apply (drule image_mono [where f = "\<lambda>z. (C * of_real r) * z"]) + using sb1 sb2 + apply force + done +qed + +corollary Bloch_general: + assumes holf: "f holomorphic_on s" and "a \<in> s" + and tle: "\<And>z. z \<in> frontier s \<Longrightarrow> t \<le> dist a z" + and rle: "r \<le> t * norm(deriv f a) / 12" + obtains b where "ball b r \<subseteq> f ` s" +proof - + consider "r \<le> 0" | "0 < t * norm(deriv f a) / 12" using rle by force + then show ?thesis + proof cases + case 1 then show ?thesis + by (simp add: Topology_Euclidean_Space.ball_empty that) + next + case 2 + show ?thesis + proof (cases "deriv f a = 0") + case True then show ?thesis + using rle by (simp add: Topology_Euclidean_Space.ball_empty that) + next + case False + then have "t > 0" + using 2 by (force simp: zero_less_mult_iff) + have "~ ball a t \<subseteq> s \<Longrightarrow> ball a t \<inter> frontier s \<noteq> {}" + apply (rule connected_Int_frontier [of "ball a t" s], simp_all) + using \<open>0 < t\<close> \<open>a \<in> s\<close> centre_in_ball apply blast + done + with tle have *: "ball a t \<subseteq> s" by fastforce + then have 1: "f holomorphic_on ball a t" + using holf using holomorphic_on_subset by blast + show ?thesis + apply (rule Bloch [OF 1 \<open>t > 0\<close> rle]) + apply (rule_tac b=b in that) + using * apply force + done + qed + qed +qed + +subsection \<open>Foundations of Cauchy's residue theorem\<close> + +text\<open>Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. + Interactive Theorem Proving\<close> + +definition residue :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where + "residue f z = (SOME int. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e + \<longrightarrow> (f has_contour_integral 2*pi* \<i> *int) (circlepath z \<epsilon>))" + +lemma contour_integral_circlepath_eq: + assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0<e1" "e1\<le>e2" + and e2_cball:"cball z e2 \<subseteq> s" + shows + "f contour_integrable_on circlepath z e1" + "f contour_integrable_on circlepath z e2" + "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" +proof - + define l where "l \<equiv> linepath (z+e2) (z+e1)" + have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto + have "e2>0" using \<open>e1>0\<close> \<open>e1\<le>e2\<close> by auto + have zl_img:"z\<notin>path_image l" + proof + assume "z \<in> path_image l" + then have "e2 \<le> cmod (e2 - e1)" + using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \<open>e1>0\<close> \<open>e2>0\<close> unfolding l_def + by (auto simp add:closed_segment_commute) + thus False using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close> + apply (subst (asm) norm_of_real) + by auto + qed + define g where "g \<equiv> circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" + show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" + proof - + show "f contour_integrable_on circlepath z e2" + apply (intro contour_integrable_continuous_circlepath[OF + continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) + using \<open>e2>0\<close> e2_cball by auto + show "f contour_integrable_on (circlepath z e1)" + apply (intro contour_integrable_continuous_circlepath[OF + continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) + using \<open>e1>0\<close> \<open>e1\<le>e2\<close> e2_cball by auto + qed + have [simp]:"f contour_integrable_on l" + proof - + have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close> + by (intro closed_segment_subset,auto simp add:dist_norm) + hence "closed_segment (z + e2) (z + e1) \<subseteq> s - {z}" using zl_img e2_cball unfolding l_def + by auto + then show "f contour_integrable_on l" unfolding l_def + apply (intro contour_integrable_continuous_linepath[OF + continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) + by auto + qed + let ?ig="\<lambda>g. contour_integral g f" + have "(f has_contour_integral 0) g" + proof (rule Cauchy_theorem_global[OF _ f_holo]) + show "open (s - {z})" using \<open>open s\<close> by auto + show "valid_path g" unfolding g_def l_def by auto + show "pathfinish g = pathstart g" unfolding g_def l_def by auto + next + have path_img:"path_image g \<subseteq> cball z e2" + proof - + have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close> + by (intro closed_segment_subset,auto simp add:dist_norm) + moreover have "sphere z \<bar>e1\<bar> \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1\<le>e2\<close> \<open>e1>0\<close> by auto + ultimately show ?thesis unfolding g_def l_def using \<open>e2>0\<close> + by (simp add: path_image_join closed_segment_commute) + qed + show "path_image g \<subseteq> s - {z}" + proof - + have "z\<notin>path_image g" using zl_img + unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) + moreover note \<open>cball z e2 \<subseteq> s\<close> and path_img + ultimately show ?thesis by auto + qed + show "winding_number g w = 0" when"w \<notin> s - {z}" for w + proof - + have "winding_number g w = 0" when "w\<notin>s" using that e2_cball + apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) + by (auto simp add:g_def l_def) + moreover have "winding_number g z=0" + proof - + let ?Wz="\<lambda>g. winding_number g z" + have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + + ?Wz (reversepath l)" + using \<open>e2>0\<close> \<open>e1>0\<close> zl_img unfolding g_def l_def + by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ + also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" + using zl_img + apply (subst (2) winding_number_reversepath) + by (auto simp add:l_def closed_segment_commute) + also have "... = 0" + proof - + have "?Wz (circlepath z e2) = 1" using \<open>e2>0\<close> + by (auto intro: winding_number_circlepath_centre) + moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \<open>e1>0\<close> + apply (subst winding_number_reversepath) + by (auto intro: winding_number_circlepath_centre) + ultimately show ?thesis by auto + qed + finally show ?thesis . + qed + ultimately show ?thesis using that by auto + qed + qed + then have "0 = ?ig g" using contour_integral_unique by simp + also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + + ?ig (reversepath l)" + unfolding g_def + by (auto simp add:contour_integrable_reversepath_eq) + also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" + by (auto simp add:contour_integral_reversepath) + finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" + by simp +qed + +lemma base_residue: + assumes "open s" "z\<in>s" "r>0" and f_holo:"f holomorphic_on (s - {z})" + and r_cball:"cball z r \<subseteq> s" + shows "(f has_contour_integral 2 * pi * \<i> * (residue f z)) (circlepath z r)" +proof - + obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" + using open_contains_cball[of s] \<open>open s\<close> \<open>z\<in>s\<close> by auto + define c where "c \<equiv> 2 * pi * \<i>" + define i where "i \<equiv> contour_integral (circlepath z e) f / c" + have "(f has_contour_integral c*i) (circlepath z \<epsilon>)" when "\<epsilon>>0" "\<epsilon><e" for \<epsilon> + proof - + have "contour_integral (circlepath z e) f = contour_integral (circlepath z \<epsilon>) f" + "f contour_integrable_on circlepath z \<epsilon>" + "f contour_integrable_on circlepath z e" + using \<open>\<epsilon><e\<close> + by (intro contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> _ e_cball],auto)+ + then show ?thesis unfolding i_def c_def + by (auto intro:has_contour_integral_integral) + qed + then have "\<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)" + unfolding residue_def c_def + apply (rule_tac someI[of _ i],intro exI[where x=e]) + by (auto simp add:\<open>e>0\<close> c_def) + then obtain e' where "e'>0" + and e'_def:"\<forall>\<epsilon>>0. \<epsilon><e' \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)" + by auto + let ?int="\<lambda>e. contour_integral (circlepath z e) f" + def \<epsilon>\<equiv>"Min {r,e'} / 2" + have "\<epsilon>>0" "\<epsilon>\<le>r" "\<epsilon><e'" using \<open>r>0\<close> \<open>e'>0\<close> unfolding \<epsilon>_def by auto + have "(f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)" + using e'_def[rule_format,OF \<open>\<epsilon>>0\<close> \<open>\<epsilon><e'\<close>] . + then show ?thesis unfolding c_def + using contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> \<open>\<epsilon>\<le>r\<close> r_cball] + by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \<epsilon>" "circlepath z r"]) +qed + + +lemma residue_holo: + assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s" + shows "residue f z = 0" +proof - + define c where "c \<equiv> 2 * pi * \<i>" + obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> + using open_contains_cball_eq by blast + have "(f has_contour_integral c*residue f z) (circlepath z e)" + using f_holo + by (auto intro: base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) + moreover have "(f has_contour_integral 0) (circlepath z e)" + using f_holo e_cball \<open>e>0\<close> + by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) + ultimately have "c*residue f z =0" + using has_contour_integral_unique by blast + thus ?thesis unfolding c_def by auto +qed + + +lemma residue_const:"residue (\<lambda>_. c) z = 0" + by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) + + +lemma residue_add: + assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" + and g_holo:"g holomorphic_on s - {z}" + shows "residue (\<lambda>z. f z + g z) z= residue f z + residue g z" +proof - + define c where "c \<equiv> 2 * pi * \<i>" + define fg where "fg \<equiv> (\<lambda>z. f z+g z)" + obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> + using open_contains_cball_eq by blast + have "(fg has_contour_integral c * residue fg z) (circlepath z e)" + unfolding fg_def using f_holo g_holo + apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) + by (auto intro:holomorphic_intros) + moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" + unfolding fg_def using f_holo g_holo + by (auto intro: has_contour_integral_add base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) + ultimately have "c*(residue f z + residue g z) = c * residue fg z" + using has_contour_integral_unique by (auto simp add:distrib_left) + thus ?thesis unfolding fg_def + by (auto simp add:c_def) +qed + + +lemma residue_lmul: + assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\<lambda>z. c * (f z)) z= c * residue f z" +proof (cases "c=0") + case True + thus ?thesis using residue_const by auto +next + case False + def c'\<equiv>"2 * pi * \<i>" + def f'\<equiv>"(\<lambda>z. c * (f z))" + obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> + using open_contains_cball_eq by blast + have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" + unfolding f'_def using f_holo + apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def]) + by (auto intro:holomorphic_intros) + moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" + unfolding f'_def using f_holo + by (auto intro: has_contour_integral_lmul + base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def]) + ultimately have "c' * residue f' z = c * (c' * residue f z)" + using has_contour_integral_unique by auto + thus ?thesis unfolding f'_def c'_def using False + by (auto simp add:field_simps) +qed + +lemma residue_rmul: + assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\<lambda>z. (f z) * c) z= residue f z * c" +using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) + +lemma residue_div: + assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\<lambda>z. (f z) / c) z= residue f z / c " +using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) + +lemma residue_neg: + assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\<lambda>z. - (f z)) z= - residue f z" +using residue_lmul[OF assms,of "-1"] by auto + +lemma residue_diff: + assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" + and g_holo:"g holomorphic_on s - {z}" + shows "residue (\<lambda>z. f z - g z) z= residue f z - residue g z" +using residue_add[OF assms(1,2,3),of "\<lambda>z. - g z"] residue_neg[OF assms(1,2,4)] +by (auto intro:holomorphic_intros g_holo) + +lemma residue_simple: + assumes "open s" "z\<in>s" and f_holo:"f holomorphic_on s" + shows "residue (\<lambda>w. f w / (w - z)) z = f z" +proof - + define c where "c \<equiv> 2 * pi * \<i>" + def f'\<equiv>"\<lambda>w. f w / (w - z)" + obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> + using open_contains_cball_eq by blast + have "(f' has_contour_integral c * f z) (circlepath z e)" + unfolding f'_def c_def using \<open>e>0\<close> f_holo e_cball + by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) + moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" + unfolding f'_def using f_holo + apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) + by (auto intro!:holomorphic_intros) + ultimately have "c * f z = c * residue f' z" + using has_contour_integral_unique by blast + thus ?thesis unfolding c_def f'_def by auto +qed + + + +subsubsection \<open>Cauchy's residue theorem\<close> + +lemma get_integrable_path: + assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\<in>s-pts" "b\<in>s-pts" + obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" + "path_image g \<subseteq> s-pts" "f contour_integrable_on g" using assms +proof (induct arbitrary:s thesis a rule:finite_induct[OF \<open>finite pts\<close>]) + case 1 + obtain g where "valid_path g" "path_image g \<subseteq> s" "pathstart g = a" "pathfinish g = b" + using connected_open_polynomial_connected[OF \<open>open s\<close>,of a b ] \<open>connected (s - {})\<close> + valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto + moreover have "f contour_integrable_on g" + using contour_integrable_holomorphic_simple[OF _ \<open>open s\<close> \<open>valid_path g\<close> \<open>path_image g \<subseteq> s\<close>,of f] + \<open>f holomorphic_on s - {}\<close> + by auto + ultimately show ?case using "1"(1)[of g] by auto +next + case idt:(2 p pts) + obtain e where "e>0" and e:"\<forall>w\<in>ball a e. w \<in> s \<and> (w \<noteq> a \<longrightarrow> w \<notin> insert p pts)" + using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>, of a] + \<open>a \<in> s - insert p pts\<close> + by auto + define a' where "a' \<equiv> a+e/2" + have "a'\<in>s-{p} -pts" using e[rule_format,of "a+e/2"] \<open>e>0\<close> + by (auto simp add:dist_complex_def a'_def) + then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" + "path_image g' \<subseteq> s - {p} - pts" "f contour_integrable_on g'" + using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) + by (metis Diff_insert2 open_delete) + define g where "g \<equiv> linepath a a' +++ g'" + have "valid_path g" unfolding g_def by (auto intro: valid_path_join) + moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto + moreover have "path_image g \<subseteq> s - insert p pts" unfolding g_def + proof (rule subset_path_image_join) + have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close> + by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) + then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9) + by auto + next + show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast + qed + moreover have "f contour_integrable_on g" + proof - + have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close> + by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) + then have "continuous_on (closed_segment a a') f" + using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] + apply (elim continuous_on_subset) + by auto + then have "f contour_integrable_on linepath a a'" + using contour_integrable_continuous_linepath by auto + then show ?thesis unfolding g_def + apply (rule contour_integrable_joinI) + by (auto simp add: \<open>e>0\<close>) + qed + ultimately show ?case using idt.prems(1)[of g] by auto +qed + +lemma Cauchy_theorem_aux: + assumes "open s" "connected (s-pts)" "finite pts" "pts \<subseteq> s" "f holomorphic_on s-pts" + "valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> s-pts" + "\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" + "\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))" + shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)" + using assms +proof (induct arbitrary:s g rule:finite_induct[OF \<open>finite pts\<close>]) + case 1 + then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) +next + case (2 p pts) + note fin[simp] = \<open>finite (insert p pts)\<close> + and connected = \<open>connected (s - insert p pts)\<close> + and valid[simp] = \<open>valid_path g\<close> + and g_loop[simp] = \<open>pathfinish g = pathstart g\<close> + and holo[simp]= \<open>f holomorphic_on s - insert p pts\<close> + and path_img = \<open>path_image g \<subseteq> s - insert p pts\<close> + and winding = \<open>\<forall>z. z \<notin> s \<longrightarrow> winding_number g z = 0\<close> + and h = \<open>\<forall>pa\<in>s. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s \<and> (w \<noteq> pa \<longrightarrow> w \<notin> insert p pts))\<close> + have "h p>0" and "p\<in>s" + and h_p: "\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> insert p pts)" + using h \<open>insert p pts \<subseteq> s\<close> by auto + obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" + "path_image pg \<subseteq> s-insert p pts" "f contour_integrable_on pg" + proof - + have "p + h p\<in>cball p (h p)" using h[rule_format,of p] + by (simp add: \<open>p \<in> s\<close> dist_norm) + then have "p + h p \<in> s - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> + by fastforce + moreover have "pathstart g \<in> s - insert p pts " using path_img by auto + ultimately show ?thesis + using get_integrable_path[OF \<open>open s\<close> connected fin holo,of "pathstart g" "p+h p"] that + by blast + qed + obtain n::int where "n=winding_number g p" + using integer_winding_number[OF _ g_loop,of p] valid path_img + by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) + define p_circ where "p_circ \<equiv> circlepath p (h p)" + define p_circ_pt where "p_circ_pt \<equiv> linepath (p+h p) (p+h p)" + define n_circ where "n_circ \<equiv> \<lambda>n. (op +++ p_circ ^^ n) p_circ_pt" + define cp where "cp \<equiv> if n\<ge>0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" + have n_circ:"valid_path (n_circ k)" + "winding_number (n_circ k) p = k" + "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" + "path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" + "p \<notin> path_image (n_circ k)" + "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number (n_circ k) p'=0 \<and> p'\<notin>path_image (n_circ k)" + "f contour_integrable_on (n_circ k)" + "contour_integral (n_circ k) f = k * contour_integral p_circ f" + for k + proof (induct k) + case 0 + show "valid_path (n_circ 0)" + and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" + and "winding_number (n_circ 0) p = of_nat 0" + and "pathstart (n_circ 0) = p + h p" + and "pathfinish (n_circ 0) = p + h p" + and "p \<notin> path_image (n_circ 0)" + unfolding n_circ_def p_circ_pt_def using \<open>h p > 0\<close> + by (auto simp add: dist_norm) + show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>s- pts" for p' + unfolding n_circ_def p_circ_pt_def + apply (auto intro!:winding_number_trivial) + by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ + show "f contour_integrable_on (n_circ 0)" + unfolding n_circ_def p_circ_pt_def + by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing) + show "contour_integral (n_circ 0) f = of_nat 0 * contour_integral p_circ f" + unfolding n_circ_def p_circ_pt_def by auto + next + case (Suc k) + have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto + have pcirc:"p \<notin> path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" + using Suc(3) unfolding p_circ_def using \<open>h p > 0\<close> by (auto simp add: p_circ_def) + have pcirc_image:"path_image p_circ \<subseteq> s - insert p pts" + proof - + have "path_image p_circ \<subseteq> cball p (h p)" using \<open>0 < h p\<close> p_circ_def by auto + then show ?thesis using h_p pcirc(1) by auto + qed + have pcirc_integrable:"f contour_integrable_on p_circ" + by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] + contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on + holomorphic_on_subset[OF holo]) + show "valid_path (n_circ (Suc k))" + using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto + show "path_image (n_circ (Suc k)) + = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" + proof - + have "path_image p_circ = sphere p (h p)" + unfolding p_circ_def using \<open>0 < h p\<close> by auto + then show ?thesis unfolding n_Suc using Suc.hyps(5) \<open>h p>0\<close> + by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) + qed + then show "p \<notin> path_image (n_circ (Suc k))" using \<open>h p>0\<close> by auto + show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" + proof - + have "winding_number p_circ p = 1" + by (simp add: \<open>h p > 0\<close> p_circ_def winding_number_circlepath_centre) + moreover have "p \<notin> path_image (n_circ k)" using Suc(5) \<open>h p>0\<close> by auto + then have "winding_number (p_circ +++ n_circ k) p + = winding_number p_circ p + winding_number (n_circ k) p" + using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc + apply (intro winding_number_join) + by auto + ultimately show ?thesis using Suc(2) unfolding n_circ_def + by auto + qed + show "pathstart (n_circ (Suc k)) = p + h p" + by (simp add: n_circ_def p_circ_def) + show "pathfinish (n_circ (Suc k)) = p + h p" + using Suc(4) unfolding n_circ_def by auto + show "winding_number (n_circ (Suc k)) p'=0 \<and> p'\<notin>path_image (n_circ (Suc k))" when "p'\<notin>s-pts" for p' + proof - + have " p' \<notin> path_image p_circ" using \<open>p \<in> s\<close> h p_circ_def that using pcirc_image by blast + moreover have "p' \<notin> path_image (n_circ k)" + using Suc.hyps(7) that by blast + moreover have "winding_number p_circ p' = 0" + proof - + have "path_image p_circ \<subseteq> cball p (h p)" + using h unfolding p_circ_def using \<open>p \<in> s\<close> by fastforce + moreover have "p'\<notin>cball p (h p)" using \<open>p \<in> s\<close> h that "2.hyps"(2) by fastforce + ultimately show ?thesis unfolding p_circ_def + apply (intro winding_number_zero_outside) + by auto + qed + ultimately show ?thesis + unfolding n_Suc + apply (subst winding_number_join) + by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that]) + qed + show "f contour_integrable_on (n_circ (Suc k))" + unfolding n_Suc + by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) + show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" + unfolding n_Suc + by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] + Suc(9) algebra_simps) + qed + have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" + "valid_path cp" "path_image cp \<subseteq> s - insert p pts" + "winding_number cp p = - n" + "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number cp p'=0 \<and> p' \<notin> path_image cp" + "f contour_integrable_on cp" + "contour_integral cp f = - n * contour_integral p_circ f" + proof - + show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" + using n_circ unfolding cp_def by auto + next + have "sphere p (h p) \<subseteq> s - insert p pts" + using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> by force + moreover have "p + complex_of_real (h p) \<in> s - insert p pts" + using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) + ultimately show "path_image cp \<subseteq> s - insert p pts" unfolding cp_def + using n_circ(5) by auto + next + show "winding_number cp p = - n" + unfolding cp_def using winding_number_reversepath n_circ \<open>h p>0\<close> + by (auto simp: valid_path_imp_path) + next + show "winding_number cp p'=0 \<and> p' \<notin> path_image cp" when "p'\<notin>s - pts" for p' + unfolding cp_def + apply (auto) + apply (subst winding_number_reversepath) + by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1)) + next + show "f contour_integrable_on cp" unfolding cp_def + using contour_integrable_reversepath_eq n_circ(1,8) by auto + next + show "contour_integral cp f = - n * contour_integral p_circ f" + unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) + by auto + qed + def g'\<equiv>"g +++ pg +++ cp +++ (reversepath pg)" + have "contour_integral g' f = (\<Sum>p\<in>pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" + proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ]) + show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) + show "open (s - {p})" using \<open>open s\<close> by auto + show " pts \<subseteq> s - {p}" using \<open>insert p pts \<subseteq> s\<close> \<open> p \<notin> pts\<close> by blast + show "f holomorphic_on s - {p} - pts" using holo \<open>p \<notin> pts\<close> by (metis Diff_insert2) + show "valid_path g'" + unfolding g'_def cp_def using n_circ valid pg g_loop + by (auto intro!:valid_path_join ) + show "pathfinish g' = pathstart g'" + unfolding g'_def cp_def using pg(2) by simp + show "path_image g' \<subseteq> s - {p} - pts" + proof - + def s'\<equiv>"s - {p} - pts" + have s':"s' = s-insert p pts " unfolding s'_def by auto + then show ?thesis using path_img pg(4) cp(4) + unfolding g'_def + apply (fold s'_def s') + apply (intro subset_path_image_join) + by auto + qed + note path_join_imp[simp] + show "\<forall>z. z \<notin> s - {p} \<longrightarrow> winding_number g' z = 0" + proof clarify + fix z assume z:"z\<notin>s - {p}" + have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z + + winding_number (pg +++ cp +++ (reversepath pg)) z" + proof (rule winding_number_join) + show "path g" using \<open>valid_path g\<close> by (simp add: valid_path_imp_path) + show "z \<notin> path_image g" using z path_img by auto + show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp + by (simp add: valid_path_imp_path) + next + have "path_image (pg +++ cp +++ reversepath pg) \<subseteq> s - insert p pts" + using pg(4) cp(4) by (auto simp:subset_path_image_join) + then show "z \<notin> path_image (pg +++ cp +++ reversepath pg)" using z by auto + next + show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto + qed + also have "... = winding_number g z + (winding_number pg z + + winding_number (cp +++ (reversepath pg)) z)" + proof (subst add_left_cancel,rule winding_number_join) + show "path pg" and "path (cp +++ reversepath pg)" + and "pathfinish pg = pathstart (cp +++ reversepath pg)" + by (auto simp add: valid_path_imp_path) + show "z \<notin> path_image pg" using pg(4) z by blast + show "z \<notin> path_image (cp +++ reversepath pg)" using z + by (metis Diff_iff \<open>z \<notin> path_image pg\<close> contra_subsetD cp(4) insertI1 + not_in_path_image_join path_image_reversepath singletonD) + qed + also have "... = winding_number g z + (winding_number pg z + + (winding_number cp z + winding_number (reversepath pg) z))" + apply (auto intro!:winding_number_join simp: valid_path_imp_path) + apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) + by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) + also have "... = winding_number g z + winding_number cp z" + apply (subst winding_number_reversepath) + apply (auto simp: valid_path_imp_path) + by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) + finally have "winding_number g' z = winding_number g z + winding_number cp z" + unfolding g'_def . + moreover have "winding_number g z + winding_number cp z = 0" + using winding z \<open>n=winding_number g p\<close> by auto + ultimately show "winding_number g' z = 0" unfolding g'_def by auto + qed + show "\<forall>pa\<in>s - {p}. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s - {p} \<and> (w \<noteq> pa \<longrightarrow> w \<notin> pts))" + using h by fastforce + qed + moreover have "contour_integral g' f = contour_integral g f + - winding_number g p * contour_integral p_circ f" + proof - + have "contour_integral g' f = contour_integral g f + + contour_integral (pg +++ cp +++ reversepath pg) f" + unfolding g'_def + apply (subst contour_integral_join) + by (auto simp add:open_Diff[OF \<open>open s\<close>,OF finite_imp_closed[OF fin]] + intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] + contour_integrable_reversepath) + also have "... = contour_integral g f + contour_integral pg f + + contour_integral (cp +++ reversepath pg) f" + apply (subst contour_integral_join) + by (auto simp add:contour_integrable_reversepath) + also have "... = contour_integral g f + contour_integral pg f + + contour_integral cp f + contour_integral (reversepath pg) f" + apply (subst contour_integral_join) + by (auto simp add:contour_integrable_reversepath) + also have "... = contour_integral g f + contour_integral cp f" + using contour_integral_reversepath + by (auto simp add:contour_integrable_reversepath) + also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f" + using \<open>n=winding_number g p\<close> by auto + finally show ?thesis . + qed + moreover have "winding_number g' p' = winding_number g p'" when "p'\<in>pts" for p' + proof - + have [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp" + using "2.prems"(8) that + apply blast + apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) + by (meson DiffD2 cp(4) set_rev_mp subset_insertI that) + have "winding_number g' p' = winding_number g p' + + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def + apply (subst winding_number_join) + apply (simp_all add: valid_path_imp_path) + apply (intro not_in_path_image_join) + by auto + also have "... = winding_number g p' + winding_number pg p' + + winding_number (cp +++ reversepath pg) p'" + apply (subst winding_number_join) + apply (simp_all add: valid_path_imp_path) + apply (intro not_in_path_image_join) + by auto + also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p' + + winding_number (reversepath pg) p'" + apply (subst winding_number_join) + by (simp_all add: valid_path_imp_path) + also have "... = winding_number g p' + winding_number cp p'" + apply (subst winding_number_reversepath) + by (simp_all add: valid_path_imp_path) + also have "... = winding_number g p'" using that by auto + finally show ?thesis . + qed + ultimately show ?case unfolding p_circ_def + apply (subst (asm) setsum.cong[OF refl, + of pts _ "\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"]) + by (auto simp add:setsum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps) +qed + +lemma Cauchy_theorem_singularities: + assumes "open s" "connected s" "finite pts" and + holo:"f holomorphic_on s-pts" and + "valid_path g" and + loop:"pathfinish g = pathstart g" and + "path_image g \<subseteq> s-pts" and + homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and + avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))" + shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)" + (is "?L=?R") +proof - + define circ where "circ \<equiv> \<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f" + define pts1 where "pts1 \<equiv> pts \<inter> s" + define pts2 where "pts2 \<equiv> pts - pts1" + have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s" + unfolding pts1_def pts2_def by auto + have "contour_integral g f = (\<Sum>p\<in>pts1. circ p)" unfolding circ_def + proof (rule Cauchy_theorem_aux[OF \<open>open s\<close> _ _ \<open>pts1\<subseteq>s\<close> _ \<open>valid_path g\<close> loop _ homo]) + have "finite pts1" unfolding pts1_def using \<open>finite pts\<close> by auto + then show "connected (s - pts1)" + using \<open>open s\<close> \<open>connected s\<close> connected_open_delete_finite[of s] by auto + next + show "finite pts1" using \<open>pts = pts1 \<union> pts2\<close> assms(3) by auto + show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) + show "path_image g \<subseteq> s - pts1" using assms(7) pts1_def by auto + show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))" + by (simp add: avoid pts1_def) + qed + moreover have "setsum circ pts2=0" + proof - + have "winding_number g p=0" when "p\<in>pts2" for p + using \<open>pts2 \<inter> s={}\<close> that homo[rule_format,of p] by auto + thus ?thesis unfolding circ_def + apply (intro setsum.neutral) + by auto + qed + moreover have "?R=setsum circ pts1 + setsum circ pts2" + unfolding circ_def + using setsum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close> + by blast + ultimately show ?thesis + apply (fold circ_def) + by auto +qed + +lemma Residue_theorem: + fixes s pts::"complex set" and f::"complex \<Rightarrow> complex" + and g::"real \<Rightarrow> complex" + assumes "open s" "connected s" "finite pts" and + holo:"f holomorphic_on s-pts" and + "valid_path g" and + loop:"pathfinish g = pathstart g" and + "path_image g \<subseteq> s-pts" and + homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" + shows "contour_integral g f = 2 * pi * \<i> *(\<Sum>p\<in>pts. winding_number g p * residue f p)" +proof - + define c where "c \<equiv> 2 * pi * \<i>" + obtain h where avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))" + using finite_cball_avoid[OF \<open>open s\<close> \<open>finite pts\<close>] by metis + have "contour_integral g f + = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)" + using Cauchy_theorem_singularities[OF assms avoid] . + also have "... = (\<Sum>p\<in>pts. c * winding_number g p * residue f p)" + proof (intro setsum.cong) + show "pts = pts" by simp + next + fix x assume "x \<in> pts" + show "winding_number g x * contour_integral (circlepath x (h x)) f + = c * winding_number g x * residue f x" + proof (cases "x\<in>s") + case False + then have "winding_number g x=0" using homo by auto + thus ?thesis by auto + next + case True + have "contour_integral (circlepath x (h x)) f = c* residue f x" + using \<open>x\<in>pts\<close> \<open>finite pts\<close> avoid[rule_format,OF True] + apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) + by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \<open>open s\<close> finite_imp_closed]) + then show ?thesis by auto + qed + qed + also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)" + by (simp add: setsum_right_distrib algebra_simps) + finally show ?thesis unfolding c_def . +qed + +subsection \<open>The argument principle\<close> + +definition is_pole :: "('a::topological_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool" where + "is_pole f a = (LIM x (at a). f x :> at_infinity)" + +lemma is_pole_tendsto: + fixes f::"('a::topological_space \<Rightarrow> 'b::real_normed_div_algebra)" + shows "is_pole f x \<Longrightarrow> ((inverse o f) \<longlongrightarrow> 0) (at x)" +unfolding is_pole_def +by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) + +lemma is_pole_inverse_holomorphic: + assumes "open s" + and f_holo:"f holomorphic_on (s-{z})" + and pole:"is_pole f z" + and non_z:"\<forall>x\<in>s-{z}. f x\<noteq>0" + shows "(\<lambda>x. if x=z then 0 else inverse (f x)) holomorphic_on s" +proof - + define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)" + have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] + apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \<circ> f"]) + by (simp_all add:g_def) + moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto + hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def + by (auto elim!:continuous_on_inverse simp add:non_z) + hence "continuous_on (s-{z}) g" unfolding g_def + apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) + by auto + ultimately have "continuous_on s g" using open_delete[OF \<open>open s\<close>] \<open>open s\<close> + by (auto simp add:continuous_on_eq_continuous_at) + moreover have "(inverse o f) holomorphic_on (s-{z})" + unfolding comp_def using f_holo + by (auto elim!:holomorphic_on_inverse simp add:non_z) + hence "g holomorphic_on (s-{z})" + apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) + by (auto simp add:g_def) + ultimately show ?thesis unfolding g_def using \<open>open s\<close> + by (auto elim!: no_isolated_singularity) +qed + + +(*order of the zero of f at z*) +definition zorder::"(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> nat" where + "zorder f z = (THE n. n>0 \<and> (\<exists>h r. r>0 \<and> h holomorphic_on cball z r + \<and> (\<forall>w\<in>cball z r. f w = h w * (w-z)^n \<and> h w \<noteq>0)))" + +definition zer_poly::"[complex \<Rightarrow> complex,complex]\<Rightarrow>complex \<Rightarrow> complex" where + "zer_poly f z = (SOME h. \<exists>r . r>0 \<and> h holomorphic_on cball z r + \<and> (\<forall>w\<in>cball z r. f w = h w * (w-z)^(zorder f z) \<and> h w \<noteq>0))" + +(*order of the pole of f at z*) +definition porder::"(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> nat" where + "porder f z = (let f'=(\<lambda>x. if x=z then 0 else inverse (f x)) in zorder f' z)" + +definition pol_poly::"[complex \<Rightarrow> complex,complex]\<Rightarrow>complex \<Rightarrow> complex" where + "pol_poly f z = (let f'=(\<lambda> x. if x=z then 0 else inverse (f x)) + in inverse o zer_poly f' z)" + + +lemma holomorphic_factor_zero_unique: + fixes f::"complex \<Rightarrow> complex" and z::complex and r::real + assumes "r>0" + and asm:"\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0" + and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" + shows "n=m" +proof - + have "n>m \<Longrightarrow> False" + proof - + assume "n>m" + have "(h \<longlongrightarrow> 0) (at z within ball z r)" + proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (n - m) * g w"]) + have "\<forall>w\<in>ball z r. w\<noteq>z \<longrightarrow> h w = (w-z)^(n-m) * g w" using \<open>n>m\<close> asm + by (auto simp add:field_simps power_diff) + then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk> + \<Longrightarrow> (x' - z) ^ (n - m) * g x' = h x'" for x' by auto + next + define F where "F \<equiv> at z within ball z r" + define f' where "f' \<equiv> \<lambda>x. (x - z) ^ (n-m)" + have "f' z=0" using \<open>n>m\<close> unfolding f'_def by auto + moreover have "continuous F f'" unfolding f'_def F_def + by (intro continuous_intros) + ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def + by (simp add: continuous_within) + moreover have "(g \<longlongrightarrow> g z) F" + using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \<open>r>0\<close> + unfolding F_def by auto + ultimately show " ((\<lambda>w. f' w * g w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce + qed + moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)" + using holomorphic_on_imp_continuous_on[OF h_holo] + by (auto simp add:continuous_on_def \<open>r>0\<close>) + moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close> + by (auto simp add:trivial_limit_within islimpt_ball) + ultimately have "h z=0" by (auto intro: tendsto_unique) + thus False using asm \<open>r>0\<close> by auto + qed + moreover have "m>n \<Longrightarrow> False" + proof - + assume "m>n" + have "(g \<longlongrightarrow> 0) (at z within ball z r)" + proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (m - n) * h w"]) + have "\<forall>w\<in>ball z r. w\<noteq>z \<longrightarrow> g w = (w-z)^(m-n) * h w" using \<open>m>n\<close> asm + by (auto simp add:field_simps power_diff) + then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk> + \<Longrightarrow> (x' - z) ^ (m - n) * h x' = g x'" for x' by auto + next + define F where "F \<equiv> at z within ball z r" + define f' where "f' \<equiv>\<lambda>x. (x - z) ^ (m-n)" + have "f' z=0" using \<open>m>n\<close> unfolding f'_def by auto + moreover have "continuous F f'" unfolding f'_def F_def + by (intro continuous_intros) + ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def + by (simp add: continuous_within) + moreover have "(h \<longlongrightarrow> h z) F" + using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \<open>r>0\<close> + unfolding F_def by auto + ultimately show " ((\<lambda>w. f' w * h w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce + qed + moreover have "(g \<longlongrightarrow> g z) (at z within ball z r)" + using holomorphic_on_imp_continuous_on[OF g_holo] + by (auto simp add:continuous_on_def \<open>r>0\<close>) + moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close> + by (auto simp add:trivial_limit_within islimpt_ball) + ultimately have "g z=0" by (auto intro: tendsto_unique) + thus False using asm \<open>r>0\<close> by auto + qed + ultimately show "n=m" by fastforce +qed + + +lemma holomorphic_factor_zero_Ex1: + assumes "open s" "connected s" "z \<in> s" and + holo:"f holomorphic_on s" + and "f z = 0" and "\<exists>w\<in>s. f w \<noteq> 0" + shows "\<exists>!n. \<exists>g r. 0 < n \<and> 0 < r \<and> + g holomorphic_on cball z r + \<and> (\<forall>w\<in>cball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0)" +proof (rule ex_ex1I) + obtain g r n where "0 < n" "0 < r" "ball z r \<subseteq> s" and + g:"g holomorphic_on ball z r" + "\<And>w. w \<in> ball z r \<Longrightarrow> f w = (w - z) ^ n * g w" + "\<And>w. w \<in> ball z r \<Longrightarrow> g w \<noteq> 0" + using holomorphic_factor_zero_nonconstant[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> \<open>f z=0\<close>] + by (metis assms(3) assms(5) assms(6)) + def r'\<equiv>"r/2" + have "cball z r' \<subseteq> ball z r" unfolding r'_def by (simp add: \<open>0 < r\<close> cball_subset_ball_iff) + hence "cball z r' \<subseteq> s" "g holomorphic_on cball z r'" + "(\<forall>w\<in>cball z r'. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)" + using g \<open>ball z r \<subseteq> s\<close> by auto + moreover have "r'>0" unfolding r'_def using \<open>0<r\<close> by auto + ultimately show "\<exists>n g r. 0 < n \<and> 0 < r \<and> g holomorphic_on cball z r + \<and> (\<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)" + apply (intro exI[of _ n] exI[of _ g] exI[of _ r']) + by (simp add:\<open>0 < n\<close>) +next + fix m n + define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0" + assume n_asm:"\<exists>g r1. 0 < n \<and> 0 < r1 \<and> g holomorphic_on cball z r1 \<and> fac n g r1" + and m_asm:"\<exists>h r2. 0 < m \<and> 0 < r2 \<and> h holomorphic_on cball z r2 \<and> fac m h r2" + obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1" + and "fac n g r1" using n_asm by auto + obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2" + and "fac m h r2" using m_asm by auto + define r where "r \<equiv> min r1 r2" + have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto + moreover have "\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0" + using \<open>fac m h r2\<close> \<open>fac n g r1\<close> unfolding fac_def r_def + by fastforce + ultimately show "m=n" using g_holo h_holo + apply (elim holomorphic_factor_zero_unique[of r z f n g m h,symmetric,rotated]) + by (auto simp add:r_def) +qed + +lemma zorder_exist: + fixes f::"complex \<Rightarrow> complex" and z::complex + defines "n\<equiv>zorder f z" and "h\<equiv>zer_poly f z" + assumes "open s" "connected s" "z\<in>s" + and holo: "f holomorphic_on s" + and "f z=0" "\<exists>w\<in>s. f w\<noteq>0" + shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r + \<and> (\<forall>w\<in>cball z r. f w = h w * (w-z)^n \<and> h w \<noteq>0) " +proof - + define P where "P \<equiv> \<lambda>h r n. r>0 \<and> h holomorphic_on cball z r + \<and> (\<forall>w\<in>cball z r. ( f w = h w * (w-z)^n) \<and> h w \<noteq>0)" + have "(\<exists>!n. n>0 \<and> (\<exists> h r. P h r n))" + proof - + have "\<exists>!n. \<exists>h r. n>0 \<and> P h r n" + using holomorphic_factor_zero_Ex1[OF \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> holo \<open>f z=0\<close> + \<open>\<exists>w\<in>s. f w\<noteq>0\<close>] unfolding P_def + apply (subst mult.commute) + by auto + thus ?thesis by auto + qed + moreover have n:"n=(THE n. n>0 \<and> (\<exists>h r. P h r n))" + unfolding n_def zorder_def P_def by simp + ultimately have "n>0 \<and> (\<exists>h r. P h r n)" + apply (drule_tac theI') + by simp + then have "n>0" and "\<exists>h r. P h r n" by auto + moreover have "h=(SOME h. \<exists>r. P h r n)" + unfolding h_def P_def zer_poly_def[of f z,folded n_def P_def] by simp + ultimately have "\<exists>r. P h r n" + apply (drule_tac someI_ex) + by simp + then obtain r1 where "P h r1 n" by auto + obtain r2 where "r2>0" "cball z r2 \<subseteq> s" + using assms(3) assms(5) open_contains_cball_eq by blast + define r3 where "r3 \<equiv> min r1 r2" + have "P h r3 n" using \<open>P h r1 n\<close> \<open>r2>0\<close> unfolding P_def r3_def + by auto + moreover have "cball z r3 \<subseteq> s" using \<open>cball z r2 \<subseteq> s\<close> unfolding r3_def by auto + ultimately show ?thesis using \<open>n>0\<close> unfolding P_def by auto +qed + +lemma porder_exist: + fixes f::"complex \<Rightarrow> complex" and z::complex + defines "n \<equiv> porder f z" and "h \<equiv> pol_poly f z" + assumes "open s" "z \<in> s" + and holo:"f holomorphic_on s-{z}" + and "is_pole f z" + shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r + \<and> (\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w-z)^n) \<and> h w \<noteq>0)" +proof - + obtain e where "e>0" and e_ball:"ball z e \<subseteq> s"and e_def: "\<forall>x\<in>ball z e-{z}. f x\<noteq>0" + proof - + have "\<forall>\<^sub>F z in at z. f z \<noteq> 0" + using \<open>is_pole f z\<close> filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def + by auto + then obtain e1 where "e1>0" and e1_def: "\<forall>x. x \<noteq> z \<and> dist x z < e1 \<longrightarrow> f x \<noteq> 0" + using eventually_at[of "\<lambda>x. f x\<noteq>0" z,simplified] by auto + obtain e2 where "e2>0" and "ball z e2 \<subseteq>s" using \<open>open s\<close> \<open>z\<in>s\<close> openE by auto + define e where "e \<equiv> min e1 e2" + have "e>0" using \<open>e1>0\<close> \<open>e2>0\<close> unfolding e_def by auto + moreover have "ball z e \<subseteq> s" unfolding e_def using \<open>ball z e2 \<subseteq> s\<close> by auto + moreover have "\<forall>x\<in>ball z e-{z}. f x\<noteq>0" using e1_def \<open>e1>0\<close> \<open>e2>0\<close> unfolding e_def + by (simp add: DiffD1 DiffD2 dist_commute singletonI) + ultimately show ?thesis using that by auto + qed + define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)" + define zo where "zo \<equiv> zorder g z" + define zp where "zp \<equiv> zer_poly g z" + have "\<exists>w\<in>ball z e. g w \<noteq> 0" + proof - + obtain w where w:"w\<in>ball z e-{z}" using \<open>0 < e\<close> + by (metis open_ball all_not_in_conv centre_in_ball insert_Diff_single + insert_absorb not_open_singleton) + hence "w\<noteq>z" "f w\<noteq>0" using e_def[rule_format,of w] mem_ball + by (auto simp add:dist_commute) + then show ?thesis unfolding g_def using w by auto + qed + moreover have "g holomorphic_on ball z e" + apply (intro is_pole_inverse_holomorphic[of "ball z e",OF _ _ \<open>is_pole f z\<close> e_def,folded g_def]) + using holo e_ball by auto + moreover have "g z=0" unfolding g_def by auto + ultimately obtain r where "0 < zo" "0 < r" "cball z r \<subseteq> ball z e" + and zp_holo: "zp holomorphic_on cball z r" and + zp_fac: "\<forall>w\<in>cball z r. g w = zp w * (w - z) ^ zo \<and> zp w \<noteq> 0" + using zorder_exist[of "ball z e" z g,simplified,folded zo_def zp_def] \<open>e>0\<close> + by auto + have n:"n=zo" and h:"h=inverse o zp" + unfolding n_def zo_def porder_def h_def zp_def pol_poly_def g_def by simp_all + have "h holomorphic_on cball z r" + using zp_holo zp_fac holomorphic_on_inverse unfolding h comp_def by blast + moreover have "\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w-z)^n) \<and> h w \<noteq>0" + using zp_fac unfolding h n comp_def g_def + by (metis divide_inverse_commute field_class.field_inverse_zero inverse_inverse_eq + inverse_mult_distrib mult.commute) + moreover have "0 < n" unfolding n using \<open>zo>0\<close> by simp + ultimately show ?thesis using \<open>0 < r\<close> \<open>cball z r \<subseteq> ball z e\<close> e_ball by auto +qed + +lemma residue_porder: + fixes f::"complex \<Rightarrow> complex" and z::complex + defines "n \<equiv> porder f z" and "h \<equiv> pol_poly f z" + assumes "open s" "z \<in> s" + and holo:"f holomorphic_on s - {z}" + and pole:"is_pole f z" + shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" +proof - + define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)" + obtain r where "0 < n" "0 < r" and r_cball:"cball z r \<subseteq> s" and h_holo: "h holomorphic_on cball z r" + and h_divide:"(\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w - z) ^ n) \<and> h w \<noteq> 0)" + using porder_exist[OF \<open>open s\<close> \<open>z \<in> s\<close> holo pole, folded n_def h_def] by blast + have r_nonzero:"\<And>w. w \<in> ball z r - {z} \<Longrightarrow> f w \<noteq> 0" + using h_divide by simp + define c where "c \<equiv> 2 * pi * \<i>" + define der_f where "der_f \<equiv> ((deriv ^^ (n - 1)) h z / fact (n-1))" + def h'\<equiv>"\<lambda>u. h u / (u - z) ^ n" + have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" + unfolding h'_def + proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", + folded c_def Suc_pred'[OF \<open>n>0\<close>]]) + show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp + show "h holomorphic_on ball z r" using h_holo by auto + show " z \<in> ball z r" using \<open>r>0\<close> by auto + qed + then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto + then have "(f has_contour_integral c * der_f) (circlepath z r)" + proof (elim has_contour_integral_eq) + fix x assume "x \<in> path_image (circlepath z r)" + hence "x\<in>cball z r - {z}" using \<open>r>0\<close> by auto + then show "h' x = f x" using h_divide unfolding h'_def by auto + qed + moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" + using base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>r>0\<close> holo r_cball,folded c_def] . + ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast + hence "der_f = residue f z" unfolding c_def by auto + thus ?thesis unfolding der_f_def by auto +qed + +theorem argument_principle: + fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set" + defines "zeros\<equiv>{p. f p=0} - poles" + assumes "open s" and + "connected s" and + f_holo:"f holomorphic_on s-poles" and + h_holo:"h holomorphic_on s" and + "valid_path g" and + loop:"pathfinish g = pathstart g" and + path_img:"path_image g \<subseteq> s - (zeros \<union> poles)" and + homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and + finite:"finite (zeros \<union> poles)" and + poles:"\<forall>p\<in>poles. is_pole f p" + shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> * + ((\<Sum>p\<in>zeros. winding_number g p * h p * zorder f p) + - (\<Sum>p\<in>poles. winding_number g p * h p * porder f p))" + (is "?L=?R") +proof - + define c where "c \<equiv> 2 * complex_of_real pi * \<i> " + define ff where "ff \<equiv> (\<lambda>x. deriv f x * h x / f x)" + define cont_pole where "cont_pole \<equiv> \<lambda>ff p e. (ff has_contour_integral - c * porder f p * h p) (circlepath p e)" + define cont_zero where "cont_zero \<equiv> \<lambda>ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" + define avoid where "avoid \<equiv> \<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> zeros \<union> poles)" + have "\<exists>e>0. avoid p e \<and> (p\<in>poles \<longrightarrow> cont_pole ff p e) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p e)" + when "p\<in>s" for p + proof - + obtain e1 where "e1>0" and e1_avoid:"avoid p e1" + using finite_cball_avoid[OF \<open>open s\<close> finite] \<open>p\<in>s\<close> unfolding avoid_def by auto + have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> cont_pole ff p e2" + when "p\<in>poles" + proof - + define po where "po \<equiv> porder f p" + define pp where "pp \<equiv> pol_poly f p" + def f'\<equiv>"\<lambda>w. pp w / (w - p) ^ po" + def ff'\<equiv>"(\<lambda>x. deriv f' x * h x / f' x)" + have "f holomorphic_on ball p e1 - {p}" + apply (intro holomorphic_on_subset[OF f_holo]) + using e1_avoid \<open>p\<in>poles\<close> unfolding avoid_def by auto + then obtain r where + "0 < po" "r>0" + "cball p r \<subseteq> ball p e1" and + pp_holo:"pp holomorphic_on cball p r" and + pp_po:"(\<forall>w\<in>cball p r. (w\<noteq>p \<longrightarrow> f w = pp w / (w - p) ^ po) \<and> pp w \<noteq> 0)" + using porder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] poles \<open>p\<in>poles\<close> + unfolding po_def pp_def + by auto + define e2 where "e2 \<equiv> r/2" + have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto + define anal where "anal \<equiv> \<lambda>w. deriv pp w * h w / pp w" + define prin where "prin \<equiv> \<lambda>w. - of_nat po * h w / (w - p)" + have "((\<lambda>w. prin w + anal w) has_contour_integral - c * po * h p) (circlepath p e2)" + proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) + have "ball p r \<subseteq> s" + using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e1_avoid by blast + then have "cball p e2 \<subseteq> s" + using \<open>r>0\<close> unfolding e2_def by auto + then have "(\<lambda>w. - of_nat po * h w) holomorphic_on cball p e2" + using h_holo + by (auto intro!: holomorphic_intros) + then show "(prin has_contour_integral - c * of_nat po * h p ) (circlepath p e2)" + using Cauchy_integral_circlepath_simple[folded c_def, of "\<lambda>w. - of_nat po * h w"] + \<open>e2>0\<close> + unfolding prin_def + by (auto simp add: mult.assoc) + have "anal holomorphic_on ball p r" unfolding anal_def + using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close> + by (auto intro!: holomorphic_intros) + then show "(anal has_contour_integral 0) (circlepath p e2)" + using e2_def \<open>r>0\<close> + by (auto elim!: Cauchy_theorem_disc_simple) + qed + then have "cont_pole ff' p e2" unfolding cont_pole_def po_def + proof (elim has_contour_integral_eq) + fix w assume "w \<in> path_image (circlepath p e2)" + then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto + define wp where "wp \<equiv> w-p" + have "wp\<noteq>0" and "pp w \<noteq>0" + unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto + moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" + proof (rule DERIV_imp_deriv) + define der where "der \<equiv> - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" + have po:"po = Suc (po - Suc 0) " using \<open>po>0\<close> by auto + have "(pp has_field_derivative (deriv pp w)) (at w)" + using DERIV_deriv_iff_has_field_derivative pp_holo \<open>w\<noteq>p\<close> + by (meson open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset) + then show "(f' has_field_derivative der) (at w)" + using \<open>w\<noteq>p\<close> \<open>po>0\<close> unfolding der_def f'_def + apply (auto intro!: derivative_eq_intros simp add:field_simps) + apply (subst (4) po) + apply (subst power_Suc) + by (auto simp add:field_simps) + qed + ultimately show "prin w + anal w = ff' w" + unfolding ff'_def prin_def anal_def + apply simp + apply (unfold f'_def) + apply (fold wp_def) + by (auto simp add:field_simps) + qed + then have "cont_pole ff p e2" unfolding cont_pole_def + proof (elim has_contour_integral_eq) + fix w assume "w \<in> path_image (circlepath p e2)" + then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto + have "deriv f' w = deriv f w" + proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) + show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo + by (auto intro!: holomorphic_intros) + next + have "ball p e1 - {p} \<subseteq> s - poles" + using avoid_def ball_subset_cball e1_avoid + by auto + then have "ball p r - {p} \<subseteq> s - poles" using \<open>cball p r \<subseteq> ball p e1\<close> + using ball_subset_cball by blast + then show "f holomorphic_on ball p r - {p}" using f_holo + by auto + next + show "open (ball p r - {p})" by auto + next + show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto + next + fix x assume "x \<in> ball p r - {p}" + then show "f' x = f x" + using pp_po unfolding f'_def by auto + qed + moreover have " f' w = f w " + using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po \<open>w\<noteq>p\<close> + unfolding f'_def by auto + ultimately show "ff' w = ff w" + unfolding ff'_def ff_def by simp + qed + moreover have "cball p e2 \<subseteq> ball p e1" + using \<open>0 < r\<close> \<open>cball p r \<subseteq> ball p e1\<close> e2_def by auto + ultimately show ?thesis using \<open>e2>0\<close> by auto + qed + then obtain e2 where e2:"p\<in>poles \<longrightarrow> e2>0 \<and> cball p e2 \<subseteq> ball p e1 \<and> cont_pole ff p e2" + by auto + have "\<exists>e3>0. cball p e3 \<subseteq> ball p e1 \<and> cont_zero ff p e3" + when "p\<in>zeros" + proof - + define zo where "zo \<equiv> zorder f p" + define zp where "zp \<equiv> zer_poly f p" + def f'\<equiv>"\<lambda>w. zp w * (w - p) ^ zo" + def ff'\<equiv>"(\<lambda>x. deriv f' x * h x / f' x)" + have "f holomorphic_on ball p e1" + proof - + have "ball p e1 \<subseteq> s - poles" + using avoid_def ball_subset_cball e1_avoid that zeros_def by fastforce + thus ?thesis using f_holo by auto + qed + moreover have "f p = 0" using \<open>p\<in>zeros\<close> + using DiffD1 mem_Collect_eq zeros_def by blast + moreover have "\<exists>w\<in>ball p e1. f w \<noteq> 0" + proof - + def p'\<equiv>"p+e1/2" + have "p'\<in>ball p e1" and "p'\<noteq>p" using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm) + then show "\<exists>w\<in>ball p e1. f w \<noteq> 0" using e1_avoid unfolding avoid_def + apply (rule_tac x=p' in bexI) + by (auto simp add:zeros_def) + qed + ultimately obtain r where + "0 < zo" "r>0" + "cball p r \<subseteq> ball p e1" and + pp_holo:"zp holomorphic_on cball p r" and + pp_po:"(\<forall>w\<in>cball p r. f w = zp w * (w - p) ^ zo \<and> zp w \<noteq> 0)" + using zorder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] unfolding zo_def zp_def + by auto + define e2 where "e2 \<equiv> r/2" + have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto + define anal where "anal \<equiv> \<lambda>w. deriv zp w * h w / zp w" + define prin where "prin \<equiv> \<lambda>w. of_nat zo * h w / (w - p)" + have "((\<lambda>w. prin w + anal w) has_contour_integral c * zo * h p) (circlepath p e2)" + proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) + have "ball p r \<subseteq> s" + using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e1_avoid by blast + then have "cball p e2 \<subseteq> s" + using \<open>r>0\<close> unfolding e2_def by auto + then have "(\<lambda>w. of_nat zo * h w) holomorphic_on cball p e2" + using h_holo + by (auto intro!: holomorphic_intros) + then show "(prin has_contour_integral c * of_nat zo * h p ) (circlepath p e2)" + using Cauchy_integral_circlepath_simple[folded c_def, of "\<lambda>w. of_nat zo * h w"] + \<open>e2>0\<close> + unfolding prin_def + by (auto simp add: mult.assoc) + have "anal holomorphic_on ball p r" unfolding anal_def + using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close> + by (auto intro!: holomorphic_intros) + then show "(anal has_contour_integral 0) (circlepath p e2)" + using e2_def \<open>r>0\<close> + by (auto elim!: Cauchy_theorem_disc_simple) + qed + then have "cont_zero ff' p e2" unfolding cont_zero_def zo_def + proof (elim has_contour_integral_eq) + fix w assume "w \<in> path_image (circlepath p e2)" + then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto + define wp where "wp \<equiv> w-p" + have "wp\<noteq>0" and "zp w \<noteq>0" + unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto + moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" + proof (rule DERIV_imp_deriv) + define der where "der \<equiv> zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" + have po:"zo = Suc (zo - Suc 0) " using \<open>zo>0\<close> by auto + have "(zp has_field_derivative (deriv zp w)) (at w)" + using DERIV_deriv_iff_has_field_derivative pp_holo + by (meson Topology_Euclidean_Space.open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset) + then show "(f' has_field_derivative der) (at w)" + using \<open>w\<noteq>p\<close> \<open>zo>0\<close> unfolding der_def f'_def + by (auto intro!: derivative_eq_intros simp add:field_simps) + qed + ultimately show "prin w + anal w = ff' w" + unfolding ff'_def prin_def anal_def + apply simp + apply (unfold f'_def) + apply (fold wp_def) + apply (auto simp add:field_simps) + by (metis Suc_diff_Suc minus_nat.diff_0 power_Suc) + qed + then have "cont_zero ff p e2" unfolding cont_zero_def + proof (elim has_contour_integral_eq) + fix w assume "w \<in> path_image (circlepath p e2)" + then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto + have "deriv f' w = deriv f w" + proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) + show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo + by (auto intro!: holomorphic_intros) + next + have "ball p e1 - {p} \<subseteq> s - poles" + using avoid_def ball_subset_cball e1_avoid by auto + then have "ball p r - {p} \<subseteq> s - poles" using \<open>cball p r \<subseteq> ball p e1\<close> + using ball_subset_cball by blast + then show "f holomorphic_on ball p r - {p}" using f_holo + by auto + next + show "open (ball p r - {p})" by auto + next + show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto + next + fix x assume "x \<in> ball p r - {p}" + then show "f' x = f x" + using pp_po unfolding f'_def by auto + qed + moreover have " f' w = f w " + using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po unfolding f'_def by auto + ultimately show "ff' w = ff w" + unfolding ff'_def ff_def by simp + qed + moreover have "cball p e2 \<subseteq> ball p e1" + using \<open>0 < r\<close> \<open>cball p r \<subseteq> ball p e1\<close> e2_def by auto + ultimately show ?thesis using \<open>e2>0\<close> by auto + qed + then obtain e3 where e3:"p\<in>zeros \<longrightarrow> e3>0 \<and> cball p e3 \<subseteq> ball p e1 \<and> cont_zero ff p e3" + by auto + define e4 where "e4 \<equiv> if p\<in>poles then e2 else if p\<in>zeros then e3 else e1" + have "e4>0" using e2 e3 \<open>e1>0\<close> unfolding e4_def by auto + moreover have "avoid p e4" using e2 e3 \<open>e1>0\<close> e1_avoid unfolding e4_def avoid_def by auto + moreover have "p\<in>poles \<longrightarrow> cont_pole ff p e4" and "p\<in>zeros \<longrightarrow> cont_zero ff p e4" + by (auto simp add: e2 e3 e4_def zeros_def) + ultimately show ?thesis by auto + qed + then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p) + \<and> (p\<in>poles \<longrightarrow> cont_pole ff p (get_e p)) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p (get_e p))" + by metis + define cont where "cont \<equiv> \<lambda>p. contour_integral (circlepath p (get_e p)) ff" + define w where "w \<equiv> \<lambda>p. winding_number g p" + have "contour_integral g ff = (\<Sum>p\<in>zeros \<union> poles. w p * cont p)" + unfolding cont_def w_def + proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected s\<close> finite _ \<open>valid_path g\<close> loop + path_img homo]) + have "open (s - (zeros \<union> poles))" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto + then show "ff holomorphic_on s - (zeros \<union> poles)" unfolding ff_def using f_holo h_holo + by (auto intro!: holomorphic_intros simp add:zeros_def) + next + show "\<forall>p\<in>s. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> zeros \<union> poles))" + using get_e using avoid_def by blast + qed + also have "... = (\<Sum>p\<in>zeros. w p * cont p) + (\<Sum>p\<in>poles. w p * cont p)" + using finite + apply (subst setsum.union_disjoint) + by (auto simp add:zeros_def) + also have "... = c * ((\<Sum>p\<in>zeros. w p * h p * zorder f p) - (\<Sum>p\<in>poles. w p * h p * porder f p))" + proof - + have "(\<Sum>p\<in>zeros. w p * cont p) = (\<Sum>p\<in>zeros. c * w p * h p * zorder f p)" + proof (rule setsum.cong[of zeros zeros,simplified]) + fix p assume "p \<in> zeros" + show "w p * cont p = c * w p * h p * (zorder f p)" + proof (cases "p\<in>s") + assume "p \<in> s" + have "cont p = c * h p * (zorder f p)" unfolding cont_def + apply (rule contour_integral_unique) + using get_e \<open>p\<in>s\<close> \<open>p\<in>zeros\<close> unfolding cont_zero_def + by (metis mult.assoc mult.commute) + thus ?thesis by auto + next + assume "p\<notin>s" + then have "w p=0" using homo unfolding w_def by auto + then show ?thesis by auto + qed + qed + then have "(\<Sum>p\<in>zeros. w p * cont p) = c * (\<Sum>p\<in>zeros. w p * h p * zorder f p)" + apply (subst setsum_right_distrib) + by (simp add:algebra_simps) + moreover have "(\<Sum>p\<in>poles. w p * cont p) = (\<Sum>p\<in>poles. - c * w p * h p * porder f p)" + proof (rule setsum.cong[of poles poles,simplified]) + fix p assume "p \<in> poles" + show "w p * cont p = - c * w p * h p * (porder f p)" + proof (cases "p\<in>s") + assume "p \<in> s" + have "cont p = - c * h p * (porder f p)" unfolding cont_def + apply (rule contour_integral_unique) + using get_e \<open>p\<in>s\<close> \<open>p\<in>poles\<close> unfolding cont_pole_def + by (metis mult.assoc mult.commute) + thus ?thesis by auto + next + assume "p\<notin>s" + then have "w p=0" using homo unfolding w_def by auto + then show ?thesis by auto + qed + qed + then have "(\<Sum>p\<in>poles. w p * cont p) = - c * (\<Sum>p\<in>poles. w p * h p * porder f p)" + apply (subst setsum_right_distrib) + by (simp add:algebra_simps) + ultimately show ?thesis by (simp add: right_diff_distrib) + qed + finally show ?thesis unfolding w_def ff_def c_def by auto +qed + +subsection \<open>Rouche's theorem \<close> + +theorem Rouche_theorem: + fixes f g::"complex \<Rightarrow> complex" and s:: "complex set" + defines "fg\<equiv>(\<lambda>p. f p+ g p)" + defines "zeros_fg\<equiv>{p. fg p =0}" and "zeros_f\<equiv>{p. f p=0}" + assumes + "open s" and "connected s" and + "finite zeros_fg" and + "finite zeros_f" and + f_holo:"f holomorphic_on s" and + g_holo:"g holomorphic_on s" and + "valid_path \<gamma>" and + loop:"pathfinish \<gamma> = pathstart \<gamma>" and + path_img:"path_image \<gamma> \<subseteq> s " and + path_less:"\<forall>z\<in>path_image \<gamma>. cmod(f z) > cmod(g z)" and + homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number \<gamma> z = 0" + shows "(\<Sum>p\<in>zeros_fg. winding_number \<gamma> p * zorder fg p) + = (\<Sum>p\<in>zeros_f. winding_number \<gamma> p * zorder f p)" +proof - + have path_fg:"path_image \<gamma> \<subseteq> s - zeros_fg" + proof - + have False when "z\<in>path_image \<gamma>" and "f z + g z=0" for z + proof - + have "cmod (f z) > cmod (g z)" using \<open>z\<in>path_image \<gamma>\<close> path_less by auto + moreover have "f z = - g z" using \<open>f z + g z =0\<close> by (simp add: eq_neg_iff_add_eq_0) + then have "cmod (f z) = cmod (g z)" by auto + ultimately show False by auto + qed + then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto + qed + have path_f:"path_image \<gamma> \<subseteq> s - zeros_f" + proof - + have False when "z\<in>path_image \<gamma>" and "f z =0" for z + proof - + have "cmod (g z) < cmod (f z) " using \<open>z\<in>path_image \<gamma>\<close> path_less by auto + then have "cmod (g z) < 0" using \<open>f z=0\<close> by auto + then show False by auto + qed + then show ?thesis unfolding zeros_f_def using path_img by auto + qed + define w where "w \<equiv> \<lambda>p. winding_number \<gamma> p" + define c where "c \<equiv> 2 * complex_of_real pi * \<i>" + define h where "h \<equiv> \<lambda>p. g p / f p + 1" + obtain spikes + where "finite spikes" and spikes: "\<forall>x\<in>{0..1} - spikes. \<gamma> differentiable at x" + using \<open>valid_path \<gamma>\<close> + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have h_contour:"((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>" + proof - + have outside_img:"0 \<in> outside (path_image (h o \<gamma>))" + proof - + have "h p \<in> ball 1 1" when "p\<in>path_image \<gamma>" for p + proof - + have "cmod (g p/f p) <1" using path_less[rule_format,OF that] + apply (cases "cmod (f p) = 0") + by (auto simp add: norm_divide) + then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) + qed + then have "path_image (h o \<gamma>) \<subseteq> ball 1 1" + by (simp add: image_subset_iff path_image_compose) + moreover have " (0::complex) \<notin> ball 1 1" by (simp add: dist_norm) + ultimately show "?thesis" + using convex_in_outside[of "ball 1 1" 0] outside_mono by blast + qed + have valid_h:"valid_path (h \<circ> \<gamma>)" + proof (rule valid_path_compose_holomorphic[OF \<open>valid_path \<gamma>\<close> _ _ path_f]) + show "h holomorphic_on s - zeros_f" + unfolding h_def using f_holo g_holo + by (auto intro!: holomorphic_intros simp add:zeros_f_def) + next + show "open (s - zeros_f)" using \<open>finite zeros_f\<close> \<open>open s\<close> finite_imp_closed + by auto + qed + have "((\<lambda>z. 1/z) has_contour_integral 0) (h \<circ> \<gamma>)" + proof - + have "0 \<notin> path_image (h \<circ> \<gamma>)" using outside_img by (simp add: outside_def) + then have "((\<lambda>z. 1/z) has_contour_integral c * winding_number (h \<circ> \<gamma>) 0) (h \<circ> \<gamma>)" + using has_contour_integral_winding_number[of "h o \<gamma>" 0,simplified] valid_h + unfolding c_def by auto + moreover have "winding_number (h o \<gamma>) 0 = 0" + proof - + have "0 \<in> outside (path_image (h \<circ> \<gamma>))" using outside_img . + moreover have "path (h o \<gamma>)" + using valid_h by (simp add: valid_path_imp_path) + moreover have "pathfinish (h o \<gamma>) = pathstart (h o \<gamma>)" + by (simp add: loop pathfinish_compose pathstart_compose) + ultimately show ?thesis using winding_number_zero_in_outside by auto + qed + ultimately show ?thesis by auto + qed + moreover have "vector_derivative (h \<circ> \<gamma>) (at x) = vector_derivative \<gamma> (at x) * deriv h (\<gamma> x)" + when "x\<in>{0..1} - spikes" for x + proof (rule vector_derivative_chain_at_general) + show "\<gamma> differentiable at x" using that \<open>valid_path \<gamma>\<close> spikes by auto + next + define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" + define t where "t \<equiv> \<gamma> x" + have "f t\<noteq>0" unfolding zeros_f_def t_def + by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) + moreover have "t\<in>s" + using contra_subsetD path_image_def path_fg t_def that by fastforce + ultimately have "(h has_field_derivative der t) (at t)" + unfolding h_def der_def using g_holo f_holo \<open>open s\<close> + by (auto intro!: holomorphic_derivI derivative_eq_intros ) + then show "\<exists>g'. (h has_field_derivative g') (at (\<gamma> x))" unfolding t_def by auto + qed + then have " (op / 1 has_contour_integral 0) (h \<circ> \<gamma>) + = ((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>" + unfolding has_contour_integral + apply (intro has_integral_spike_eq[OF negligible_finite, OF \<open>finite spikes\<close>]) + by auto + ultimately show ?thesis by auto + qed + then have "contour_integral \<gamma> (\<lambda>x. deriv h x / h x) = 0" + using contour_integral_unique by simp + moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = contour_integral \<gamma> (\<lambda>x. deriv f x / f x) + + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)" + proof - + have "(\<lambda>p. deriv f p / f p) contour_integrable_on \<gamma>" + proof (rule contour_integrable_holomorphic_simple[OF _ _ \<open>valid_path \<gamma>\<close> path_f]) + show "open (s - zeros_f)" using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close> + by auto + then show "(\<lambda>p. deriv f p / f p) holomorphic_on s - zeros_f" + using f_holo + by (auto intro!: holomorphic_intros simp add:zeros_f_def) + qed + moreover have "(\<lambda>p. deriv h p / h p) contour_integrable_on \<gamma>" + using h_contour + by (simp add: has_contour_integral_integrable) + ultimately have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x + deriv h x / h x) = + contour_integral \<gamma> (\<lambda>p. deriv f p / f p) + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)" + using contour_integral_add[of "(\<lambda>p. deriv f p / f p)" \<gamma> "(\<lambda>p. deriv h p / h p)" ] + by auto + moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" + when "p\<in> path_image \<gamma>" for p + proof - + have "fg p\<noteq>0" and "f p\<noteq>0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def + by auto + have "h p\<noteq>0" + proof (rule ccontr) + assume "\<not> h p \<noteq> 0" + then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) + then have "cmod (g p/f p) = 1" by auto + moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] + apply (cases "cmod (f p) = 0") + by (auto simp add: norm_divide) + ultimately show False by auto + qed + have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def + using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \<open>open s\<close>] path_img that + by auto + have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" + proof - + define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" + have "p\<in>s" using path_img that by auto + then have "(h has_field_derivative der p) (at p)" + unfolding h_def der_def using g_holo f_holo \<open>open s\<close> \<open>f p\<noteq>0\<close> + by (auto intro!: derivative_eq_intros holomorphic_derivI) + then show ?thesis unfolding der_def using DERIV_imp_deriv by auto + qed + show ?thesis + apply (simp only:der_fg der_h) + apply (auto simp add:field_simps \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>) + by (auto simp add:field_simps h_def \<open>f p\<noteq>0\<close> fg_def) + qed + then have "contour_integral \<gamma> (\<lambda>p. deriv fg p / fg p) + = contour_integral \<gamma> (\<lambda>p. deriv f p / f p + deriv h p / h p)" + by (elim contour_integral_eq) + ultimately show ?thesis by auto + qed + moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = c * (\<Sum>p\<in>zeros_fg. w p * zorder fg p)" + unfolding c_def zeros_fg_def w_def + proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo + , of _ "{}" "\<lambda>_. 1",simplified]) + show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto + show "path_image \<gamma> \<subseteq> s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def . + show " finite {p. fg p = 0}" using \<open>finite zeros_fg\<close> unfolding zeros_fg_def . + qed + moreover have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x) = c * (\<Sum>p\<in>zeros_f. w p * zorder f p)" + unfolding c_def zeros_f_def w_def + proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo + , of _ "{}" "\<lambda>_. 1",simplified]) + show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto + show "path_image \<gamma> \<subseteq> s - {p. f p = 0}" using path_f unfolding zeros_f_def . + show " finite {p. f p = 0}" using \<open>finite zeros_f\<close> unfolding zeros_f_def . + qed + ultimately have " c* (\<Sum>p\<in>zeros_fg. w p * (zorder fg p)) = c* (\<Sum>p\<in>zeros_f. w p * (zorder f p))" + by auto + then show ?thesis unfolding c_def using w_def by auto +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Continuous_Extension.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,547 @@ +(* Title: HOL/Analysis/Extension.thy + Authors: LC Paulson, based on material from HOL Light +*) + +section \<open>Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\<close> + +theory Continuous_Extension +imports Convex_Euclidean_Space +begin + +subsection\<open>Partitions of unity subordinate to locally finite open coverings\<close> + +text\<open>A difference from HOL Light: all summations over infinite sets equal zero, + so the "support" must be made explicit in the summation below!\<close> + +proposition subordinate_partition_of_unity: + fixes S :: "'a :: euclidean_space set" + assumes "S \<subseteq> \<Union>\<C>" and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T" + and fin: "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. U \<inter> V \<noteq> {}}" + obtains F :: "['a set, 'a] \<Rightarrow> real" + where "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x \<in> S. 0 \<le> F U x)" + and "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0" + and "\<And>x. x \<in> S \<Longrightarrow> supp_setsum (\<lambda>W. F W x) \<C> = 1" + and "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}" +proof (cases "\<exists>W. W \<in> \<C> \<and> S \<subseteq> W") + case True + then obtain W where "W \<in> \<C>" "S \<subseteq> W" by metis + then show ?thesis + apply (rule_tac F = "\<lambda>V x. if V = W then 1 else 0" in that) + apply (auto simp: continuous_on_const supp_setsum_def support_on_def) + done +next + case False + have nonneg: "0 \<le> supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" for x + by (simp add: supp_setsum_def setsum_nonneg) + have sd_pos: "0 < setdist {x} (S - V)" if "V \<in> \<C>" "x \<in> S" "x \<in> V" for V x + proof - + have "closedin (subtopology euclidean S) (S - V)" + by (simp add: Diff_Diff_Int Diff_subset closedin_def opC openin_open_Int \<open>V \<in> \<C>\<close>) + with that False setdist_eq_0_closedin [of S "S-V" x] setdist_pos_le [of "{x}" "S - V"] + show ?thesis + by (simp add: order_class.order.order_iff_strict) + qed + have ss_pos: "0 < supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" if "x \<in> S" for x + proof - + obtain U where "U \<in> \<C>" "x \<in> U" using \<open>x \<in> S\<close> \<open>S \<subseteq> \<Union>\<C>\<close> + by blast + obtain V where "open V" "x \<in> V" "finite {U \<in> \<C>. U \<inter> V \<noteq> {}}" + using \<open>x \<in> S\<close> fin by blast + then have *: "finite {A \<in> \<C>. \<not> S \<subseteq> A \<and> x \<notin> closure (S - A)}" + using closure_def that by (blast intro: rev_finite_subset) + have "x \<notin> closure (S - U)" + by (metis \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> less_irrefl sd_pos setdist_eq_0_sing_1 that) + then show ?thesis + apply (simp add: setdist_eq_0_sing_1 supp_setsum_def support_on_def) + apply (rule ordered_comm_monoid_add_class.setsum_pos2 [OF *, of U]) + using \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> False + apply (auto simp: setdist_pos_le sd_pos that) + done + qed + define F where + "F \<equiv> \<lambda>W x. if x \<in> S then setdist {x} (S - W) / supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C> + else 0" + show ?thesis + proof (rule_tac F = F in that) + have "continuous_on S (F U)" if "U \<in> \<C>" for U + proof - + have *: "continuous_on S (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)" + proof (clarsimp simp add: continuous_on_eq_continuous_within) + fix x assume "x \<in> S" + then obtain X where "open X" and x: "x \<in> S \<inter> X" and finX: "finite {U \<in> \<C>. U \<inter> X \<noteq> {}}" + using assms by blast + then have OSX: "openin (subtopology euclidean S) (S \<inter> X)" by blast + have sumeq: "\<And>x. x \<in> S \<inter> X \<Longrightarrow> + (\<Sum>V | V \<in> \<C> \<and> V \<inter> X \<noteq> {}. setdist {x} (S - V)) + = supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" + apply (simp add: supp_setsum_def) + apply (rule setsum.mono_neutral_right [OF finX]) + apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff) + apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE) + done + show "continuous (at x within S) (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)" + apply (rule continuous_transform_within_openin + [where f = "\<lambda>x. (setsum (\<lambda>V. setdist {x} (S - V)) {V \<in> \<C>. V \<inter> X \<noteq> {}})" + and S ="S \<inter> X"]) + apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+ + apply (simp add: sumeq) + done + qed + show ?thesis + apply (simp add: F_def) + apply (rule continuous_intros *)+ + using ss_pos apply force + done + qed + moreover have "\<lbrakk>U \<in> \<C>; x \<in> S\<rbrakk> \<Longrightarrow> 0 \<le> F U x" for U x + using nonneg [of x] by (simp add: F_def divide_simps setdist_pos_le) + ultimately show "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x\<in>S. 0 \<le> F U x)" + by metis + next + show "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0" + by (simp add: setdist_eq_0_sing_1 closure_def F_def) + next + show "supp_setsum (\<lambda>W. F W x) \<C> = 1" if "x \<in> S" for x + using that ss_pos [OF that] + by (simp add: F_def divide_simps supp_setsum_divide_distrib [symmetric]) + next + show "\<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}" if "x \<in> S" for x + using fin [OF that] that + by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset) + qed +qed + + +subsection\<open>Urysohn's lemma (for Euclidean spaces, where the proof is easy using distances)\<close> + +lemma Urysohn_both_ne: + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \<inter> T = {}" "S \<noteq> {}" "T \<noteq> {}" "a \<noteq> b" + obtains f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + where "continuous_on U f" + "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b" + "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)" + "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)" +proof - + have S0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} S = 0 \<longleftrightarrow> x \<in> S" + using \<open>S \<noteq> {}\<close> US setdist_eq_0_closedin by auto + have T0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} T = 0 \<longleftrightarrow> x \<in> T" + using \<open>T \<noteq> {}\<close> UT setdist_eq_0_closedin by auto + have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \<in> U" for x + proof - + have "~ (setdist {x} S = 0 \<and> setdist {x} T = 0)" + using assms by (metis IntI empty_iff setdist_eq_0_closedin that) + then show ?thesis + by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le) + qed + define f where "f \<equiv> \<lambda>x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)" + show ?thesis + proof (rule_tac f = f in that) + show "continuous_on U f" + using sdpos unfolding f_def + by (intro continuous_intros | force)+ + show "f x \<in> closed_segment a b" if "x \<in> U" for x + unfolding f_def + apply (simp add: closed_segment_def) + apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI) + using sdpos that apply (simp add: algebra_simps) + done + show "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)" + using S0 \<open>a \<noteq> b\<close> f_def sdpos by force + show "(f x = b \<longleftrightarrow> x \<in> T)" if "x \<in> U" for x + proof - + have "f x = b \<longleftrightarrow> (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1" + unfolding f_def + apply (rule iffI) + apply (metis \<open>a \<noteq> b\<close> add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force) + done + also have "... \<longleftrightarrow> setdist {x} T = 0 \<and> setdist {x} S \<noteq> 0" + using sdpos that + by (simp add: divide_simps) linarith + also have "... \<longleftrightarrow> x \<in> T" + using \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>S \<inter> T = {}\<close> that + by (force simp: S0 T0) + finally show ?thesis . + qed + qed +qed + +proposition Urysohn_local_strong: + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \<inter> T = {}" "a \<noteq> b" + obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + where "continuous_on U f" + "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b" + "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)" + "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)" +proof (cases "S = {}") + case True show ?thesis + proof (cases "T = {}") + case True show ?thesis + proof (rule_tac f = "\<lambda>x. midpoint a b" in that) + show "continuous_on U (\<lambda>x. midpoint a b)" + by (intro continuous_intros) + show "midpoint a b \<in> closed_segment a b" + using csegment_midpoint_subset by blast + show "(midpoint a b = a) = (x \<in> S)" for x + using \<open>S = {}\<close> \<open>a \<noteq> b\<close> by simp + show "(midpoint a b = b) = (x \<in> T)" for x + using \<open>T = {}\<close> \<open>a \<noteq> b\<close> by simp + qed + next + case False + show ?thesis + proof (cases "T = U") + case True with \<open>S = {}\<close> \<open>a \<noteq> b\<close> show ?thesis + by (rule_tac f = "\<lambda>x. b" in that) (auto simp: continuous_on_const) + next + case False + with UT closedin_subset obtain c where c: "c \<in> U" "c \<notin> T" + by fastforce + obtain f where f: "continuous_on U f" + "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment (midpoint a b) b" + "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)" + "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)" + apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"]) + using c \<open>T \<noteq> {}\<close> assms apply simp_all + done + show ?thesis + apply (rule_tac f=f in that) + using \<open>S = {}\<close> \<open>T \<noteq> {}\<close> f csegment_midpoint_subset notin_segment_midpoint [OF \<open>a \<noteq> b\<close>] + apply force+ + done + qed + qed +next + case False + show ?thesis + proof (cases "T = {}") + case True show ?thesis + proof (cases "S = U") + case True with \<open>T = {}\<close> \<open>a \<noteq> b\<close> show ?thesis + by (rule_tac f = "\<lambda>x. a" in that) (auto simp: continuous_on_const) + next + case False + with US closedin_subset obtain c where c: "c \<in> U" "c \<notin> S" + by fastforce + obtain f where f: "continuous_on U f" + "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a (midpoint a b)" + "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)" + "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)" + apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"]) + using c \<open>S \<noteq> {}\<close> assms apply simp_all + apply (metis midpoint_eq_endpoint) + done + show ?thesis + apply (rule_tac f=f in that) + using \<open>S \<noteq> {}\<close> \<open>T = {}\<close> f \<open>a \<noteq> b\<close> + apply simp_all + apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff) + apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint) + done + qed + next + case False + show ?thesis + using Urysohn_both_ne [OF US UT \<open>S \<inter> T = {}\<close> \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>a \<noteq> b\<close>] that + by blast + qed +qed + +lemma Urysohn_local: + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \<inter> T = {}" + obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + where "continuous_on U f" + "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b" + "\<And>x. x \<in> S \<Longrightarrow> f x = a" + "\<And>x. x \<in> T \<Longrightarrow> f x = b" +proof (cases "a = b") + case True then show ?thesis + by (rule_tac f = "\<lambda>x. b" in that) (auto simp: continuous_on_const) +next + case False + then show ?thesis + apply (rule Urysohn_local_strong [OF assms]) + apply (erule that, assumption) + apply (meson US closedin_singleton closedin_trans) + apply (meson UT closedin_singleton closedin_trans) + done +qed + +lemma Urysohn_strong: + assumes US: "closed S" + and UT: "closed T" + and "S \<inter> T = {}" "a \<noteq> b" + obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + where "continuous_on UNIV f" + "\<And>x. f x \<in> closed_segment a b" + "\<And>x. f x = a \<longleftrightarrow> x \<in> S" + "\<And>x. f x = b \<longleftrightarrow> x \<in> T" +apply (rule Urysohn_local_strong [of UNIV S T]) +using assms +apply (auto simp: closed_closedin) +done + +proposition Urysohn: + assumes US: "closed S" + and UT: "closed T" + and "S \<inter> T = {}" + obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + where "continuous_on UNIV f" + "\<And>x. f x \<in> closed_segment a b" + "\<And>x. x \<in> S \<Longrightarrow> f x = a" + "\<And>x. x \<in> T \<Longrightarrow> f x = b" +apply (rule Urysohn_local [of UNIV S T a b]) +using assms +apply (auto simp: closed_closedin) +done + + +subsection\<open> The Dugundji extension theorem, and Tietze variants as corollaries.\<close> + +text\<open>J. Dugundji. An extension of Tietze's theorem. Pacific J. Math. Volume 1, Number 3 (1951), 353-367. +http://projecteuclid.org/euclid.pjm/1103052106\<close> + +theorem Dugundji: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner" + assumes "convex C" "C \<noteq> {}" + and cloin: "closedin (subtopology euclidean U) S" + and contf: "continuous_on S f" and "f ` S \<subseteq> C" + obtains g where "continuous_on U g" "g ` U \<subseteq> C" + "\<And>x. x \<in> S \<Longrightarrow> g x = f x" +proof (cases "S = {}") + case True then show thesis + apply (rule_tac g="\<lambda>x. @y. y \<in> C" in that) + apply (rule continuous_intros) + apply (meson all_not_in_conv \<open>C \<noteq> {}\<close> image_subsetI someI_ex, simp) + done +next + case False + then have sd_pos: "\<And>x. \<lbrakk>x \<in> U; x \<notin> S\<rbrakk> \<Longrightarrow> 0 < setdist {x} S" + using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce + define \<B> where "\<B> = {ball x (setdist {x} S / 2) |x. x \<in> U - S}" + have [simp]: "\<And>T. T \<in> \<B> \<Longrightarrow> open T" + by (auto simp: \<B>_def) + have USS: "U - S \<subseteq> \<Union>\<B>" + by (auto simp: sd_pos \<B>_def) + obtain \<C> where USsub: "U - S \<subseteq> \<Union>\<C>" + and nbrhd: "\<And>U. U \<in> \<C> \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<B> \<and> U \<subseteq> T)" + and fin: "\<And>x. x \<in> U - S + \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U. U \<in> \<C> \<and> U \<inter> V \<noteq> {}}" + using paracompact [OF USS] by auto + have "\<exists>v a. v \<in> U \<and> v \<notin> S \<and> a \<in> S \<and> + T \<subseteq> ball v (setdist {v} S / 2) \<and> + dist v a \<le> 2 * setdist {v} S" if "T \<in> \<C>" for T + proof - + obtain v where v: "T \<subseteq> ball v (setdist {v} S / 2)" "v \<in> U" "v \<notin> S" + using \<open>T \<in> \<C>\<close> nbrhd by (force simp: \<B>_def) + then obtain a where "a \<in> S" "dist v a < 2 * setdist {v} S" + using setdist_ltE [of "{v}" S "2 * setdist {v} S"] + using False sd_pos by force + with v show ?thesis + apply (rule_tac x=v in exI) + apply (rule_tac x=a in exI, auto) + done + qed + then obtain \<V> \<A> where + VA: "\<And>T. T \<in> \<C> \<Longrightarrow> \<V> T \<in> U \<and> \<V> T \<notin> S \<and> \<A> T \<in> S \<and> + T \<subseteq> ball (\<V> T) (setdist {\<V> T} S / 2) \<and> + dist (\<V> T) (\<A> T) \<le> 2 * setdist {\<V> T} S" + by metis + have sdle: "setdist {\<V> T} S \<le> 2 * setdist {v} S" if "T \<in> \<C>" "v \<in> T" for T v + using setdist_Lipschitz [of "\<V> T" S v] VA [OF \<open>T \<in> \<C>\<close>] \<open>v \<in> T\<close> by auto + have d6: "dist a (\<A> T) \<le> 6 * dist a v" if "T \<in> \<C>" "v \<in> T" "a \<in> S" for T v a + proof - + have "dist (\<V> T) v < setdist {\<V> T} S / 2" + using that VA mem_ball by blast + also have "... \<le> setdist {v} S" + using sdle [OF \<open>T \<in> \<C>\<close> \<open>v \<in> T\<close>] by simp + also have vS: "setdist {v} S \<le> dist a v" + by (simp add: setdist_le_dist setdist_sym \<open>a \<in> S\<close>) + finally have VTV: "dist (\<V> T) v < dist a v" . + have VTS: "setdist {\<V> T} S \<le> 2 * dist a v" + using sdle that vS by force + have "dist a (\<A> T) \<le> dist a v + dist v (\<V> T) + dist (\<V> T) (\<A> T)" + by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le) + also have "... \<le> dist a v + dist a v + dist (\<V> T) (\<A> T)" + using VTV by (simp add: dist_commute) + also have "... \<le> 2 * dist a v + 2 * setdist {\<V> T} S" + using VA [OF \<open>T \<in> \<C>\<close>] by auto + finally show ?thesis + using VTS by linarith + qed + obtain H :: "['a set, 'a] \<Rightarrow> real" + where Hcont: "\<And>Z. Z \<in> \<C> \<Longrightarrow> continuous_on (U-S) (H Z)" + and Hge0: "\<And>Z x. \<lbrakk>Z \<in> \<C>; x \<in> U-S\<rbrakk> \<Longrightarrow> 0 \<le> H Z x" + and Heq0: "\<And>x Z. \<lbrakk>Z \<in> \<C>; x \<in> U-S; x \<notin> Z\<rbrakk> \<Longrightarrow> H Z x = 0" + and H1: "\<And>x. x \<in> U-S \<Longrightarrow> supp_setsum (\<lambda>W. H W x) \<C> = 1" + and Hfin: "\<And>x. x \<in> U-S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. H U x \<noteq> 0}" + apply (rule subordinate_partition_of_unity [OF USsub _ fin]) + using nbrhd by auto + define g where "g \<equiv> \<lambda>x. if x \<in> S then f x else supp_setsum (\<lambda>T. H T x *\<^sub>R f(\<A> T)) \<C>" + show ?thesis + proof (rule that) + show "continuous_on U g" + proof (clarsimp simp: continuous_on_eq_continuous_within) + fix a assume "a \<in> U" + show "continuous (at a within U) g" + proof (cases "a \<in> S") + case True show ?thesis + proof (clarsimp simp add: continuous_within_topological) + fix W + assume "open W" "g a \<in> W" + then obtain e where "0 < e" and e: "ball (f a) e \<subseteq> W" + using openE True g_def by auto + have "continuous (at a within S) f" + using True contf continuous_on_eq_continuous_within by blast + then obtain d where "0 < d" + and d: "\<And>x. \<lbrakk>x \<in> S; dist x a < d\<rbrakk> \<Longrightarrow> dist (f x) (f a) < e" + using continuous_within_eps_delta \<open>0 < e\<close> by force + have "g y \<in> ball (f a) e" if "y \<in> U" and y: "y \<in> ball a (d / 6)" for y + proof (cases "y \<in> S") + case True + then have "dist (f a) (f y) < e" + by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d) + then show ?thesis + by (simp add: True g_def) + next + case False + have *: "dist (f (\<A> T)) (f a) < e" if "T \<in> \<C>" "H T y \<noteq> 0" for T + proof - + have "y \<in> T" + using Heq0 that False \<open>y \<in> U\<close> by blast + have "dist (\<A> T) a < d" + using d6 [OF \<open>T \<in> \<C>\<close> \<open>y \<in> T\<close> \<open>a \<in> S\<close>] y + by (simp add: dist_commute mult.commute) + then show ?thesis + using VA [OF \<open>T \<in> \<C>\<close>] by (auto simp: d) + qed + have "supp_setsum (\<lambda>T. H T y *\<^sub>R f (\<A> T)) \<C> \<in> ball (f a) e" + apply (rule convex_supp_setsum [OF convex_ball]) + apply (simp_all add: False H1 Hge0 \<open>y \<in> U\<close>) + by (metis dist_commute *) + then show ?thesis + by (simp add: False g_def) + qed + then show "\<exists>A. open A \<and> a \<in> A \<and> (\<forall>y\<in>U. y \<in> A \<longrightarrow> g y \<in> W)" + apply (rule_tac x = "ball a (d / 6)" in exI) + using e \<open>0 < d\<close> by fastforce + qed + next + case False + obtain N where N: "open N" "a \<in> N" + and finN: "finite {U \<in> \<C>. \<exists>a\<in>N. H U a \<noteq> 0}" + using Hfin False \<open>a \<in> U\<close> by auto + have oUS: "openin (subtopology euclidean U) (U - S)" + using cloin by (simp add: openin_diff) + have HcontU: "continuous (at a within U) (H T)" if "T \<in> \<C>" for T + using Hcont [OF \<open>T \<in> \<C>\<close>] False \<open>a \<in> U\<close> \<open>T \<in> \<C>\<close> + apply (simp add: continuous_on_eq_continuous_within continuous_within) + apply (rule Lim_transform_within_set) + using oUS + apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+ + done + show ?thesis + proof (rule continuous_transform_within_openin [OF _ oUS]) + show "continuous (at a within U) (\<lambda>x. supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>)" + proof (rule continuous_transform_within_openin) + show "continuous (at a within U) + (\<lambda>x. \<Sum>T\<in>{U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))" + by (force intro: continuous_intros HcontU)+ + next + show "openin (subtopology euclidean U) ((U - S) \<inter> N)" + using N oUS openin_trans by blast + next + show "a \<in> (U - S) \<inter> N" using False \<open>a \<in> U\<close> N by blast + next + show "\<And>x. x \<in> (U - S) \<inter> N \<Longrightarrow> + (\<Sum>T \<in> {U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T)) + = supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>" + by (auto simp: supp_setsum_def support_on_def + intro: setsum.mono_neutral_right [OF finN]) + qed + next + show "a \<in> U - S" using False \<open>a \<in> U\<close> by blast + next + show "\<And>x. x \<in> U - S \<Longrightarrow> supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C> = g x" + by (simp add: g_def) + qed + qed + qed + show "g ` U \<subseteq> C" + using \<open>f ` S \<subseteq> C\<close> VA + by (fastforce simp: g_def Hge0 intro!: convex_supp_setsum [OF \<open>convex C\<close>] H1) + show "\<And>x. x \<in> S \<Longrightarrow> g x = f x" + by (simp add: g_def) + qed +qed + + +corollary Tietze: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "0 \<le> B" + and "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> B" + obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" + "\<And>x. x \<in> U \<Longrightarrow> norm(g x) \<le> B" +using assms +by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f]) + +corollary Tietze_closed_interval: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "cbox a b \<noteq> {}" + and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b" + obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" + "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b" +apply (rule Dugundji [of "cbox a b" U S f]) +using assms by auto + +corollary Tietze_closed_interval_1: + fixes f :: "'a::euclidean_space \<Rightarrow> real" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "a \<le> b" + and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b" + obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" + "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b" +apply (rule Dugundji [of "cbox a b" U S f]) +using assms by (auto simp: image_subset_iff) + +corollary Tietze_open_interval: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "box a b \<noteq> {}" + and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b" + obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" + "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b" +apply (rule Dugundji [of "box a b" U S f]) +using assms by auto + +corollary Tietze_open_interval_1: + fixes f :: "'a::euclidean_space \<Rightarrow> real" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "a < b" + and no: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b" + obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" + "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b" +apply (rule Dugundji [of "box a b" U S f]) +using assms by (auto simp: image_subset_iff) + +corollary Tietze_unbounded: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" +apply (rule Dugundji [of UNIV U S f]) +using assms by auto + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,12687 @@ +(* Title: HOL/Analysis/Convex_Euclidean_Space.thy + Authors: Robert Himmelmann, TU Muenchen; Bogdan Grechuk, University of Edinburgh; LCP +*) + +section \<open>Convex sets, functions and related things\<close> + +theory Convex_Euclidean_Space +imports + Topology_Euclidean_Space + "~~/src/HOL/Library/Convex" + "~~/src/HOL/Library/Set_Algebras" +begin + +(*FIXME move up e.g. to Library/Convex.thy, but requires the "support" primitive*) +lemma convex_supp_setsum: + assumes "convex S" and 1: "supp_setsum u I = 1" + and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)" + shows "supp_setsum (\<lambda>i. u i *\<^sub>R f i) I \<in> S" +proof - + have fin: "finite {i \<in> I. u i \<noteq> 0}" + using 1 setsum.infinite by (force simp: supp_setsum_def support_on_def) + then have eq: "supp_setsum (\<lambda>i. u i *\<^sub>R f i) I = setsum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}" + by (force intro: setsum.mono_neutral_left simp: supp_setsum_def support_on_def) + show ?thesis + apply (simp add: eq) + apply (rule convex_setsum [OF fin \<open>convex S\<close>]) + using 1 assms apply (auto simp: supp_setsum_def support_on_def) + done +qed + + +lemma dim_image_eq: + fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" + assumes lf: "linear f" + and fi: "inj_on f (span S)" + shows "dim (f ` S) = dim (S::'n::euclidean_space set)" +proof - + obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" + using basis_exists[of S] by auto + then have "span S = span B" + using span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto + then have "independent (f ` B)" + using independent_inj_on_image[of B f] B assms by auto + moreover have "card (f ` B) = card B" + using assms card_image[of f B] subset_inj_on[of f "span S" B] B span_inc by auto + moreover have "(f ` B) \<subseteq> (f ` S)" + using B by auto + ultimately have "dim (f ` S) \<ge> dim S" + using independent_card_le_dim[of "f ` B" "f ` S"] B by auto + then show ?thesis + using dim_image_le[of f S] assms by auto +qed + +lemma linear_injective_on_subspace_0: + assumes lf: "linear f" + and "subspace S" + shows "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)" +proof - + have "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x = f y \<longrightarrow> x = y)" + by (simp add: inj_on_def) + also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x - f y = 0 \<longrightarrow> x - y = 0)" + by simp + also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f (x - y) = 0 \<longrightarrow> x - y = 0)" + by (simp add: linear_diff[OF lf]) + also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)" + using \<open>subspace S\<close> subspace_def[of S] subspace_diff[of S] by auto + finally show ?thesis . +qed + +lemma subspace_Inter: "\<forall>s \<in> f. subspace s \<Longrightarrow> subspace (\<Inter>f)" + unfolding subspace_def by auto + +lemma span_eq[simp]: "span s = s \<longleftrightarrow> subspace s" + unfolding span_def by (rule hull_eq) (rule subspace_Inter) + +lemma substdbasis_expansion_unique: + assumes d: "d \<subseteq> Basis" + shows "(\<Sum>i\<in>d. f i *\<^sub>R i) = (x::'a::euclidean_space) \<longleftrightarrow> + (\<forall>i\<in>Basis. (i \<in> d \<longrightarrow> f i = x \<bullet> i) \<and> (i \<notin> d \<longrightarrow> x \<bullet> i = 0))" +proof - + have *: "\<And>x a b P. x * (if P then a else b) = (if P then x * a else x * b)" + by auto + have **: "finite d" + by (auto intro: finite_subset[OF assms]) + have ***: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>i\<in>d. f i *\<^sub>R i) \<bullet> i = (\<Sum>x\<in>d. if x = i then f x else 0)" + using d + by (auto intro!: setsum.cong simp: inner_Basis inner_setsum_left) + show ?thesis + unfolding euclidean_eq_iff[where 'a='a] by (auto simp: setsum.delta[OF **] ***) +qed + +lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d" + by (rule independent_mono[OF independent_Basis]) + +lemma dim_cball: + assumes "e > 0" + shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)" +proof - + { + fix x :: "'n::euclidean_space" + define y where "y = (e / norm x) *\<^sub>R x" + then have "y \<in> cball 0 e" + using assms by auto + moreover have *: "x = (norm x / e) *\<^sub>R y" + using y_def assms by simp + moreover from * have "x = (norm x/e) *\<^sub>R y" + by auto + ultimately have "x \<in> span (cball 0 e)" + using span_mul[of y "cball 0 e" "norm x/e"] span_inc[of "cball 0 e"] + by (simp add: span_superset) + } + then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" + by auto + then show ?thesis + using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp add: dim_UNIV) +qed + +lemma indep_card_eq_dim_span: + fixes B :: "'n::euclidean_space set" + assumes "independent B" + shows "finite B \<and> card B = dim (span B)" + using assms basis_card_eq_dim[of B "span B"] span_inc by auto + +lemma setsum_not_0: "setsum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0" + by (rule ccontr) auto + +lemma subset_translation_eq [simp]: + fixes a :: "'a::real_vector" shows "op + a ` s \<subseteq> op + a ` t \<longleftrightarrow> s \<subseteq> t" + by auto + +lemma translate_inj_on: + fixes A :: "'a::ab_group_add set" + shows "inj_on (\<lambda>x. a + x) A" + unfolding inj_on_def by auto + +lemma translation_assoc: + fixes a b :: "'a::ab_group_add" + shows "(\<lambda>x. b + x) ` ((\<lambda>x. a + x) ` S) = (\<lambda>x. (a + b) + x) ` S" + by auto + +lemma translation_invert: + fixes a :: "'a::ab_group_add" + assumes "(\<lambda>x. a + x) ` A = (\<lambda>x. a + x) ` B" + shows "A = B" +proof - + have "(\<lambda>x. -a + x) ` ((\<lambda>x. a + x) ` A) = (\<lambda>x. - a + x) ` ((\<lambda>x. a + x) ` B)" + using assms by auto + then show ?thesis + using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto +qed + +lemma translation_galois: + fixes a :: "'a::ab_group_add" + shows "T = ((\<lambda>x. a + x) ` S) \<longleftrightarrow> S = ((\<lambda>x. (- a) + x) ` T)" + using translation_assoc[of "-a" a S] + apply auto + using translation_assoc[of a "-a" T] + apply auto + done + +lemma convex_translation_eq [simp]: "convex ((\<lambda>x. a + x) ` s) \<longleftrightarrow> convex s" + by (metis convex_translation translation_galois) + +lemma translation_inverse_subset: + assumes "((\<lambda>x. - a + x) ` V) \<le> (S :: 'n::ab_group_add set)" + shows "V \<le> ((\<lambda>x. a + x) ` S)" +proof - + { + fix x + assume "x \<in> V" + then have "x-a \<in> S" using assms by auto + then have "x \<in> {a + v |v. v \<in> S}" + apply auto + apply (rule exI[of _ "x-a"]) + apply simp + done + then have "x \<in> ((\<lambda>x. a+x) ` S)" by auto + } + then show ?thesis by auto +qed + +lemma convex_linear_image_eq [simp]: + fixes f :: "'a::real_vector \<Rightarrow> 'b::real_vector" + shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> convex (f ` s) \<longleftrightarrow> convex s" + by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq) + +lemma basis_to_basis_subspace_isomorphism: + assumes s: "subspace (S:: ('n::euclidean_space) set)" + and t: "subspace (T :: ('m::euclidean_space) set)" + and d: "dim S = dim T" + and B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" + and C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T" + shows "\<exists>f. linear f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S" +proof - + from B independent_bound have fB: "finite B" + by blast + from C independent_bound have fC: "finite C" + by blast + from B(4) C(4) card_le_inj[of B C] d obtain f where + f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto + from linear_independent_extend[OF B(2)] obtain g where + g: "linear g" "\<forall>x \<in> B. g x = f x" by blast + from inj_on_iff_eq_card[OF fB, of f] f(2) + have "card (f ` B) = card B" by simp + with B(4) C(4) have ceq: "card (f ` B) = card C" using d + by simp + have "g ` B = f ` B" using g(2) + by (auto simp add: image_iff) + also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] . + finally have gBC: "g ` B = C" . + have gi: "inj_on g B" using f(2) g(2) + by (auto simp add: inj_on_def) + note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] + { + fix x y + assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y" + from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" + by blast+ + from gxy have th0: "g (x - y) = 0" + by (simp add: linear_diff[OF g(1)]) + have th1: "x - y \<in> span B" using x' y' + by (metis span_sub) + have "x = y" using g0[OF th1 th0] by simp + } + then have giS: "inj_on g S" unfolding inj_on_def by blast + from span_subspace[OF B(1,3) s] + have "g ` S = span (g ` B)" + by (simp add: span_linear_image[OF g(1)]) + also have "\<dots> = span C" + unfolding gBC .. + also have "\<dots> = T" + using span_subspace[OF C(1,3) t] . + finally have gS: "g ` S = T" . + from g(1) gS giS gBC show ?thesis + by blast +qed + +lemma closure_bounded_linear_image_subset: + assumes f: "bounded_linear f" + shows "f ` closure S \<subseteq> closure (f ` S)" + using linear_continuous_on [OF f] closed_closure closure_subset + by (rule image_closure_subset) + +lemma closure_linear_image_subset: + fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector" + assumes "linear f" + shows "f ` (closure S) \<subseteq> closure (f ` S)" + using assms unfolding linear_conv_bounded_linear + by (rule closure_bounded_linear_image_subset) + +lemma closed_injective_linear_image: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes S: "closed S" and f: "linear f" "inj f" + shows "closed (f ` S)" +proof - + obtain g where g: "linear g" "g \<circ> f = id" + using linear_injective_left_inverse [OF f] by blast + then have confg: "continuous_on (range f) g" + using linear_continuous_on linear_conv_bounded_linear by blast + have [simp]: "g ` f ` S = S" + using g by (simp add: image_comp) + have cgf: "closed (g ` f ` S)" + by (simp add: \<open>g \<circ> f = id\<close> S image_comp) + have [simp]: "{x \<in> range f. g x \<in> S} = f ` S" + using g by (simp add: o_def id_def image_def) metis + show ?thesis + apply (rule closedin_closed_trans [of "range f"]) + apply (rule continuous_closedin_preimage [OF confg cgf, simplified]) + apply (rule closed_injective_image_subspace) + using f + apply (auto simp: linear_linear linear_injective_0) + done +qed + +lemma closed_injective_linear_image_eq: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes f: "linear f" "inj f" + shows "(closed(image f s) \<longleftrightarrow> closed s)" + by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) + +lemma closure_injective_linear_image: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" + apply (rule subset_antisym) + apply (simp add: closure_linear_image_subset) + by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) + +lemma closure_bounded_linear_image: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" + apply (rule subset_antisym, simp add: closure_linear_image_subset) + apply (rule closure_minimal, simp add: closure_subset image_mono) + by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) + +lemma closure_scaleR: + fixes S :: "'a::real_normed_vector set" + shows "(op *\<^sub>R c) ` (closure S) = closure ((op *\<^sub>R c) ` S)" +proof + show "(op *\<^sub>R c) ` (closure S) \<subseteq> closure ((op *\<^sub>R c) ` S)" + using bounded_linear_scaleR_right + by (rule closure_bounded_linear_image_subset) + show "closure ((op *\<^sub>R c) ` S) \<subseteq> (op *\<^sub>R c) ` (closure S)" + by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) +qed + +lemma fst_linear: "linear fst" + unfolding linear_iff by (simp add: algebra_simps) + +lemma snd_linear: "linear snd" + unfolding linear_iff by (simp add: algebra_simps) + +lemma fst_snd_linear: "linear (\<lambda>(x,y). x + y)" + unfolding linear_iff by (simp add: algebra_simps) + +lemma scaleR_2: + fixes x :: "'a::real_vector" + shows "scaleR 2 x = x + x" + unfolding one_add_one [symmetric] scaleR_left_distrib by simp + +lemma scaleR_half_double [simp]: + fixes a :: "'a::real_normed_vector" + shows "(1 / 2) *\<^sub>R (a + a) = a" +proof - + have "\<And>r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" + by (metis scaleR_2 scaleR_scaleR) + then show ?thesis + by simp +qed + +lemma vector_choose_size: + assumes "0 \<le> c" + obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c" +proof - + obtain a::'a where "a \<noteq> 0" + using UNIV_not_singleton UNIV_eq_I set_zero singletonI by fastforce + then show ?thesis + by (rule_tac x="scaleR (c / norm a) a" in that) (simp add: assms) +qed + +lemma vector_choose_dist: + assumes "0 \<le> c" + obtains y :: "'a::{real_normed_vector, perfect_space}" where "dist x y = c" +by (metis add_diff_cancel_left' assms dist_commute dist_norm vector_choose_size) + +lemma sphere_eq_empty [simp]: + fixes a :: "'a::{real_normed_vector, perfect_space}" + shows "sphere a r = {} \<longleftrightarrow> r < 0" +by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) + +lemma setsum_delta_notmem: + assumes "x \<notin> s" + shows "setsum (\<lambda>y. if (y = x) then P x else Q y) s = setsum Q s" + and "setsum (\<lambda>y. if (x = y) then P x else Q y) s = setsum Q s" + and "setsum (\<lambda>y. if (y = x) then P y else Q y) s = setsum Q s" + and "setsum (\<lambda>y. if (x = y) then P y else Q y) s = setsum Q s" + apply (rule_tac [!] setsum.cong) + using assms + apply auto + done + +lemma setsum_delta'': + fixes s::"'a::real_vector set" + assumes "finite s" + shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)" +proof - + have *: "\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)" + by auto + show ?thesis + unfolding * using setsum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto +qed + +lemma if_smult: "(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)" + by (fact if_distrib) + +lemma dist_triangle_eq: + fixes x y z :: "'a::real_inner" + shows "dist x z = dist x y + dist y z \<longleftrightarrow> + norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" +proof - + have *: "x - y + (y - z) = x - z" by auto + show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] + by (auto simp add:norm_minus_commute) +qed + +lemma norm_minus_eqI: "x = - y \<Longrightarrow> norm x = norm y" by auto + +lemma Min_grI: + assumes "finite A" "A \<noteq> {}" "\<forall>a\<in>A. x < a" + shows "x < Min A" + unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto + +lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y" + unfolding norm_eq_sqrt_inner by simp + +lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y" + unfolding norm_eq_sqrt_inner by simp + + +subsection \<open>Affine set and affine hull\<close> + +definition affine :: "'a::real_vector set \<Rightarrow> bool" + where "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)" + +lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)" + unfolding affine_def by (metis eq_diff_eq') + +lemma affine_empty [iff]: "affine {}" + unfolding affine_def by auto + +lemma affine_sing [iff]: "affine {x}" + unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) + +lemma affine_UNIV [iff]: "affine UNIV" + unfolding affine_def by auto + +lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)" + unfolding affine_def by auto + +lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)" + unfolding affine_def by auto + +lemma affine_scaling: "affine s \<Longrightarrow> affine (image (\<lambda>x. c *\<^sub>R x) s)" + apply (clarsimp simp add: affine_def) + apply (rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in image_eqI) + apply (auto simp: algebra_simps) + done + +lemma affine_affine_hull [simp]: "affine(affine hull s)" + unfolding hull_def + using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto + +lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s" + by (metis affine_affine_hull hull_same) + +lemma affine_hyperplane: "affine {x. a \<bullet> x = b}" + by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral) + + +subsubsection \<open>Some explicit formulations (from Lars Schewe)\<close> + +lemma affine: + fixes V::"'a::real_vector set" + shows "affine V \<longleftrightarrow> + (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (setsum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)" + unfolding affine_def + apply rule + apply(rule, rule, rule) + apply(erule conjE)+ + defer + apply (rule, rule, rule, rule, rule) +proof - + fix x y u v + assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)" + "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" + then show "u *\<^sub>R x + v *\<^sub>R y \<in> V" + apply (cases "x = y") + using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]] + and as(1-3) + apply (auto simp add: scaleR_left_distrib[symmetric]) + done +next + fix s u + assume as: "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V" + "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = (1::real)" + define n where "n = card s" + have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto + then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" + proof (auto simp only: disjE) + assume "card s = 2" + then have "card s = Suc (Suc 0)" + by auto + then obtain a b where "s = {a, b}" + unfolding card_Suc_eq by auto + then show ?thesis + using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5) + by (auto simp add: setsum_clauses(2)) + next + assume "card s > 2" + then show ?thesis using as and n_def + proof (induct n arbitrary: u s) + case 0 + then show ?case by auto + next + case (Suc n) + fix s :: "'a set" and u :: "'a \<Rightarrow> real" + assume IA: + "\<And>u s. \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s; + s \<noteq> {}; s \<subseteq> V; setsum u s = 1; n = card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" + and as: + "Suc n = card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V" + "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = 1" + have "\<exists>x\<in>s. u x \<noteq> 1" + proof (rule ccontr) + assume "\<not> ?thesis" + then have "setsum u s = real_of_nat (card s)" + unfolding card_eq_setsum by auto + then show False + using as(7) and \<open>card s > 2\<close> + by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2) + qed + then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto + + have c: "card (s - {x}) = card s - 1" + apply (rule card_Diff_singleton) + using \<open>x\<in>s\<close> as(4) + apply auto + done + have *: "s = insert x (s - {x})" "finite (s - {x})" + using \<open>x\<in>s\<close> and as(4) by auto + have **: "setsum u (s - {x}) = 1 - u x" + using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto + have ***: "inverse (1 - u x) * setsum u (s - {x}) = 1" + unfolding ** using \<open>u x \<noteq> 1\<close> by auto + have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V" + proof (cases "card (s - {x}) > 2") + case True + then have "s - {x} \<noteq> {}" "card (s - {x}) = n" + unfolding c and as(1)[symmetric] + proof (rule_tac ccontr) + assume "\<not> s - {x} \<noteq> {}" + then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp + then show False using True by auto + qed auto + then show ?thesis + apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"]) + unfolding setsum_right_distrib[symmetric] + using as and *** and True + apply auto + done + next + case False + then have "card (s - {x}) = Suc (Suc 0)" + using as(2) and c by auto + then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b" + unfolding card_Suc_eq by auto + then show ?thesis + using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] + using *** *(2) and \<open>s \<subseteq> V\<close> + unfolding setsum_right_distrib + by (auto simp add: setsum_clauses(2)) + qed + then have "u x + (1 - u x) = 1 \<Longrightarrow> + u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V" + apply - + apply (rule as(3)[rule_format]) + unfolding Real_Vector_Spaces.scaleR_right.setsum + using x(1) as(6) + apply auto + done + then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" + unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric] + apply (subst *) + unfolding setsum_clauses(2)[OF *(2)] + using \<open>u x \<noteq> 1\<close> + apply auto + done + qed + next + assume "card s = 1" + then obtain a where "s={a}" + by (auto simp add: card_Suc_eq) + then show ?thesis + using as(4,5) by simp + qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto) +qed + +lemma affine_hull_explicit: + "affine hull p = + {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> setsum (\<lambda>v. (u v) *\<^sub>R v) s = y}" + apply (rule hull_unique) + apply (subst subset_eq) + prefer 3 + apply rule + unfolding mem_Collect_eq + apply (erule exE)+ + apply (erule conjE)+ + prefer 2 + apply rule +proof - + fix x + assume "x\<in>p" + then show "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" + apply (rule_tac x="{x}" in exI) + apply (rule_tac x="\<lambda>x. 1" in exI) + apply auto + done +next + fix t x s u + assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}" + "s \<subseteq> p" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" + then show "x \<in> t" + using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]] + by auto +next + show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}" + unfolding affine_def + apply (rule, rule, rule, rule, rule) + unfolding mem_Collect_eq + proof - + fix u v :: real + assume uv: "u + v = 1" + fix x + assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" + then obtain sx ux where + x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "setsum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x" + by auto + fix y + assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" + then obtain sy uy where + y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "setsum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto + have xy: "finite (sx \<union> sy)" + using x(1) y(1) by auto + have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy" + by auto + show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> + setsum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" + apply (rule_tac x="sx \<union> sy" in exI) + apply (rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI) + unfolding scaleR_left_distrib setsum.distrib if_smult scaleR_zero_left + ** setsum.inter_restrict[OF xy, symmetric] + unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric] + and setsum_right_distrib[symmetric] + unfolding x y + using x(1-3) y(1-3) uv + apply simp + done + qed +qed + +lemma affine_hull_finite: + assumes "finite s" + shows "affine hull s = {y. \<exists>u. setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}" + unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq + apply (rule, rule) + apply (erule exE)+ + apply (erule conjE)+ + defer + apply (erule exE) + apply (erule conjE) +proof - + fix x u + assume "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" + then show "\<exists>sa u. finite sa \<and> + \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x" + apply (rule_tac x=s in exI, rule_tac x=u in exI) + using assms + apply auto + done +next + fix x t u + assume "t \<subseteq> s" + then have *: "s \<inter> t = t" + by auto + assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x" + then show "\<exists>u. setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" + apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) + unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms, symmetric] and * + apply auto + done +qed + + +subsubsection \<open>Stepping theorems and hence small special cases\<close> + +lemma affine_hull_empty[simp]: "affine hull {} = {}" + by (rule hull_unique) auto + +lemma affine_hull_finite_step: + fixes y :: "'a::real_vector" + shows + "(\<exists>u. setsum u {} = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1) + and + "finite s \<Longrightarrow> + (\<exists>u. setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow> + (\<exists>v u. setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs") +proof - + show ?th1 by simp + assume fin: "finite s" + show "?lhs = ?rhs" + proof + assume ?lhs + then obtain u where u: "setsum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" + by auto + show ?rhs + proof (cases "a \<in> s") + case True + then have *: "insert a s = s" by auto + show ?thesis + using u[unfolded *] + apply(rule_tac x=0 in exI) + apply auto + done + next + case False + then show ?thesis + apply (rule_tac x="u a" in exI) + using u and fin + apply auto + done + qed + next + assume ?rhs + then obtain v u where vu: "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" + by auto + have *: "\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" + by auto + show ?lhs + proof (cases "a \<in> s") + case True + then show ?thesis + apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI) + unfolding setsum_clauses(2)[OF fin] + apply simp + unfolding scaleR_left_distrib and setsum.distrib + unfolding vu and * and scaleR_zero_left + apply (auto simp add: setsum.delta[OF fin]) + done + next + case False + then have **: + "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)" + "\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto + from False show ?thesis + apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI) + unfolding setsum_clauses(2)[OF fin] and * using vu + using setsum.cong [of s _ "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)] + using setsum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)] + apply auto + done + qed + qed +qed + +lemma affine_hull_2: + fixes a b :: "'a::real_vector" + shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}" + (is "?lhs = ?rhs") +proof - + have *: + "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" + "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto + have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}" + using affine_hull_finite[of "{a,b}"] by auto + also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}" + by (simp add: affine_hull_finite_step(2)[of "{b}" a]) + also have "\<dots> = ?rhs" unfolding * by auto + finally show ?thesis by auto +qed + +lemma affine_hull_3: + fixes a b c :: "'a::real_vector" + shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}" +proof - + have *: + "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" + "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto + show ?thesis + apply (simp add: affine_hull_finite affine_hull_finite_step) + unfolding * + apply auto + apply (rule_tac x=v in exI) + apply (rule_tac x=va in exI) + apply auto + apply (rule_tac x=u in exI) + apply force + done +qed + +lemma mem_affine: + assumes "affine S" "x \<in> S" "y \<in> S" "u + v = 1" + shows "u *\<^sub>R x + v *\<^sub>R y \<in> S" + using assms affine_def[of S] by auto + +lemma mem_affine_3: + assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" "u + v + w = 1" + shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> S" +proof - + have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> affine hull {x, y, z}" + using affine_hull_3[of x y z] assms by auto + moreover + have "affine hull {x, y, z} \<subseteq> affine hull S" + using hull_mono[of "{x, y, z}" "S"] assms by auto + moreover + have "affine hull S = S" + using assms affine_hull_eq[of S] by auto + ultimately show ?thesis by auto +qed + +lemma mem_affine_3_minus: + assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" + shows "x + v *\<^sub>R (y-z) \<in> S" + using mem_affine_3[of S x y z 1 v "-v"] assms + by (simp add: algebra_simps) + +corollary mem_affine_3_minus2: + "\<lbrakk>affine S; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> x - v *\<^sub>R (y-z) \<in> S" + by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left) + + +subsubsection \<open>Some relations between affine hull and subspaces\<close> + +lemma affine_hull_insert_subset_span: + "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}" + unfolding subset_eq Ball_def + unfolding affine_hull_explicit span_explicit mem_Collect_eq + apply (rule, rule) + apply (erule exE)+ + apply (erule conjE)+ +proof - + fix x t u + assume as: "finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x" + have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}" + using as(3) by auto + then show "\<exists>v. x = a + v \<and> (\<exists>S u. finite S \<and> S \<subseteq> {x - a |x. x \<in> s} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = v)" + apply (rule_tac x="x - a" in exI) + apply (rule conjI, simp) + apply (rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI) + apply (rule_tac x="\<lambda>x. u (x + a)" in exI) + apply (rule conjI) using as(1) apply simp + apply (erule conjI) + using as(1) + apply (simp add: setsum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib + setsum_subtractf scaleR_left.setsum[symmetric] setsum_diff1 scaleR_left_diff_distrib) + unfolding as + apply simp + done +qed + +lemma affine_hull_insert_span: + assumes "a \<notin> s" + shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x. x \<in> s}}" + apply (rule, rule affine_hull_insert_subset_span) + unfolding subset_eq Ball_def + unfolding affine_hull_explicit and mem_Collect_eq +proof (rule, rule, erule exE, erule conjE) + fix y v + assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}" + then obtain t u where obt: "finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y" + unfolding span_explicit by auto + define f where "f = (\<lambda>x. x + a) ` t" + have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a" + unfolding f_def using obt by (auto simp add: setsum.reindex[unfolded inj_on_def]) + have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f" + using f(2) assms by auto + show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y" + apply (rule_tac x = "insert a f" in exI) + apply (rule_tac x = "\<lambda>x. if x=a then 1 - setsum (\<lambda>x. u (x - a)) f else u (x - a)" in exI) + using assms and f + unfolding setsum_clauses(2)[OF f(1)] and if_smult + unfolding setsum.If_cases[OF f(1), of "\<lambda>x. x = a"] + apply (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps *) + done +qed + +lemma affine_hull_span: + assumes "a \<in> s" + shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}" + using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto + + +subsubsection \<open>Parallel affine sets\<close> + +definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool" + where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)" + +lemma affine_parallel_expl_aux: + fixes S T :: "'a::real_vector set" + assumes "\<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T" + shows "T = (\<lambda>x. a + x) ` S" +proof - + { + fix x + assume "x \<in> T" + then have "( - a) + x \<in> S" + using assms by auto + then have "x \<in> ((\<lambda>x. a + x) ` S)" + using imageI[of "-a+x" S "(\<lambda>x. a+x)"] by auto + } + moreover have "T \<ge> (\<lambda>x. a + x) ` S" + using assms by auto + ultimately show ?thesis by auto +qed + +lemma affine_parallel_expl: "affine_parallel S T \<longleftrightarrow> (\<exists>a. \<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T)" + unfolding affine_parallel_def + using affine_parallel_expl_aux[of S _ T] by auto + +lemma affine_parallel_reflex: "affine_parallel S S" + unfolding affine_parallel_def + apply (rule exI[of _ "0"]) + apply auto + done + +lemma affine_parallel_commut: + assumes "affine_parallel A B" + shows "affine_parallel B A" +proof - + from assms obtain a where B: "B = (\<lambda>x. a + x) ` A" + unfolding affine_parallel_def by auto + have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff) + from B show ?thesis + using translation_galois [of B a A] + unfolding affine_parallel_def by auto +qed + +lemma affine_parallel_assoc: + assumes "affine_parallel A B" + and "affine_parallel B C" + shows "affine_parallel A C" +proof - + from assms obtain ab where "B = (\<lambda>x. ab + x) ` A" + unfolding affine_parallel_def by auto + moreover + from assms obtain bc where "C = (\<lambda>x. bc + x) ` B" + unfolding affine_parallel_def by auto + ultimately show ?thesis + using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto +qed + +lemma affine_translation_aux: + fixes a :: "'a::real_vector" + assumes "affine ((\<lambda>x. a + x) ` S)" + shows "affine S" +proof - + { + fix x y u v + assume xy: "x \<in> S" "y \<in> S" "(u :: real) + v = 1" + then have "(a + x) \<in> ((\<lambda>x. a + x) ` S)" "(a + y) \<in> ((\<lambda>x. a + x) ` S)" + by auto + then have h1: "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) \<in> (\<lambda>x. a + x) ` S" + using xy assms unfolding affine_def by auto + have "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) = (u + v) *\<^sub>R a + (u *\<^sub>R x + v *\<^sub>R y)" + by (simp add: algebra_simps) + also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)" + using \<open>u + v = 1\<close> by auto + ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S" + using h1 by auto + then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto + } + then show ?thesis unfolding affine_def by auto +qed + +lemma affine_translation: + fixes a :: "'a::real_vector" + shows "affine S \<longleftrightarrow> affine ((\<lambda>x. a + x) ` S)" +proof - + have "affine S \<Longrightarrow> affine ((\<lambda>x. a + x) ` S)" + using affine_translation_aux[of "-a" "((\<lambda>x. a + x) ` S)"] + using translation_assoc[of "-a" a S] by auto + then show ?thesis using affine_translation_aux by auto +qed + +lemma parallel_is_affine: + fixes S T :: "'a::real_vector set" + assumes "affine S" "affine_parallel S T" + shows "affine T" +proof - + from assms obtain a where "T = (\<lambda>x. a + x) ` S" + unfolding affine_parallel_def by auto + then show ?thesis + using affine_translation assms by auto +qed + +lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s" + unfolding subspace_def affine_def by auto + + +subsubsection \<open>Subspace parallel to an affine set\<close> + +lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S" +proof - + have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S" + using subspace_imp_affine[of S] subspace_0 by auto + { + assume assm: "affine S \<and> 0 \<in> S" + { + fix c :: real + fix x + assume x: "x \<in> S" + have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto + moreover + have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S" + using affine_alt[of S] assm x by auto + ultimately have "c *\<^sub>R x \<in> S" by auto + } + then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto + + { + fix x y + assume xy: "x \<in> S" "y \<in> S" + define u where "u = (1 :: real)/2" + have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)" + by auto + moreover + have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y" + by (simp add: algebra_simps) + moreover + have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S" + using affine_alt[of S] assm xy by auto + ultimately + have "(1/2) *\<^sub>R (x+y) \<in> S" + using u_def by auto + moreover + have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))" + by auto + ultimately + have "x + y \<in> S" + using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto + } + then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S" + by auto + then have "subspace S" + using h1 assm unfolding subspace_def by auto + } + then show ?thesis using h0 by metis +qed + +lemma affine_diffs_subspace: + assumes "affine S" "a \<in> S" + shows "subspace ((\<lambda>x. (-a)+x) ` S)" +proof - + have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff) + have "affine ((\<lambda>x. (-a)+x) ` S)" + using affine_translation assms by auto + moreover have "0 : ((\<lambda>x. (-a)+x) ` S)" + using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto + ultimately show ?thesis using subspace_affine by auto +qed + +lemma parallel_subspace_explicit: + assumes "affine S" + and "a \<in> S" + assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}" + shows "subspace L \<and> affine_parallel S L" +proof - + from assms have "L = plus (- a) ` S" by auto + then have par: "affine_parallel S L" + unfolding affine_parallel_def .. + then have "affine L" using assms parallel_is_affine by auto + moreover have "0 \<in> L" + using assms by auto + ultimately show ?thesis + using subspace_affine par by auto +qed + +lemma parallel_subspace_aux: + assumes "subspace A" + and "subspace B" + and "affine_parallel A B" + shows "A \<supseteq> B" +proof - + from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B" + using affine_parallel_expl[of A B] by auto + then have "-a \<in> A" + using assms subspace_0[of B] by auto + then have "a \<in> A" + using assms subspace_neg[of A "-a"] by auto + then show ?thesis + using assms a unfolding subspace_def by auto +qed + +lemma parallel_subspace: + assumes "subspace A" + and "subspace B" + and "affine_parallel A B" + shows "A = B" +proof + show "A \<supseteq> B" + using assms parallel_subspace_aux by auto + show "A \<subseteq> B" + using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto +qed + +lemma affine_parallel_subspace: + assumes "affine S" "S \<noteq> {}" + shows "\<exists>!L. subspace L \<and> affine_parallel S L" +proof - + have ex: "\<exists>L. subspace L \<and> affine_parallel S L" + using assms parallel_subspace_explicit by auto + { + fix L1 L2 + assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2" + then have "affine_parallel L1 L2" + using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto + then have "L1 = L2" + using ass parallel_subspace by auto + } + then show ?thesis using ex by auto +qed + + +subsection \<open>Cones\<close> + +definition cone :: "'a::real_vector set \<Rightarrow> bool" + where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)" + +lemma cone_empty[intro, simp]: "cone {}" + unfolding cone_def by auto + +lemma cone_univ[intro, simp]: "cone UNIV" + unfolding cone_def by auto + +lemma cone_Inter[intro]: "\<forall>s\<in>f. cone s \<Longrightarrow> cone (\<Inter>f)" + unfolding cone_def by auto + +lemma subspace_imp_cone: "subspace S \<Longrightarrow> cone S" + by (simp add: cone_def subspace_mul) + + +subsubsection \<open>Conic hull\<close> + +lemma cone_cone_hull: "cone (cone hull s)" + unfolding hull_def by auto + +lemma cone_hull_eq: "cone hull s = s \<longleftrightarrow> cone s" + apply (rule hull_eq) + using cone_Inter + unfolding subset_eq + apply auto + done + +lemma mem_cone: + assumes "cone S" "x \<in> S" "c \<ge> 0" + shows "c *\<^sub>R x : S" + using assms cone_def[of S] by auto + +lemma cone_contains_0: + assumes "cone S" + shows "S \<noteq> {} \<longleftrightarrow> 0 \<in> S" +proof - + { + assume "S \<noteq> {}" + then obtain a where "a \<in> S" by auto + then have "0 \<in> S" + using assms mem_cone[of S a 0] by auto + } + then show ?thesis by auto +qed + +lemma cone_0: "cone {0}" + unfolding cone_def by auto + +lemma cone_Union[intro]: "(\<forall>s\<in>f. cone s) \<longrightarrow> cone (\<Union>f)" + unfolding cone_def by blast + +lemma cone_iff: + assumes "S \<noteq> {}" + shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" +proof - + { + assume "cone S" + { + fix c :: real + assume "c > 0" + { + fix x + assume "x \<in> S" + then have "x \<in> (op *\<^sub>R c) ` S" + unfolding image_def + using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"] + exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] + by auto + } + moreover + { + fix x + assume "x \<in> (op *\<^sub>R c) ` S" + then have "x \<in> S" + using \<open>cone S\<close> \<open>c > 0\<close> + unfolding cone_def image_def \<open>c > 0\<close> by auto + } + ultimately have "(op *\<^sub>R c) ` S = S" by auto + } + then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" + using \<open>cone S\<close> cone_contains_0[of S] assms by auto + } + moreover + { + assume a: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" + { + fix x + assume "x \<in> S" + fix c1 :: real + assume "c1 \<ge> 0" + then have "c1 = 0 \<or> c1 > 0" by auto + then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto + } + then have "cone S" unfolding cone_def by auto + } + ultimately show ?thesis by blast +qed + +lemma cone_hull_empty: "cone hull {} = {}" + by (metis cone_empty cone_hull_eq) + +lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}" + by (metis bot_least cone_hull_empty hull_subset xtrans(5)) + +lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S" + using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S] + by auto + +lemma mem_cone_hull: + assumes "x \<in> S" "c \<ge> 0" + shows "c *\<^sub>R x \<in> cone hull S" + by (metis assms cone_cone_hull hull_inc mem_cone) + +lemma cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}" + (is "?lhs = ?rhs") +proof - + { + fix x + assume "x \<in> ?rhs" + then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" + by auto + fix c :: real + assume c: "c \<ge> 0" + then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx" + using x by (simp add: algebra_simps) + moreover + have "c * cx \<ge> 0" using c x by auto + ultimately + have "c *\<^sub>R x \<in> ?rhs" using x by auto + } + then have "cone ?rhs" + unfolding cone_def by auto + then have "?rhs \<in> Collect cone" + unfolding mem_Collect_eq by auto + { + fix x + assume "x \<in> S" + then have "1 *\<^sub>R x \<in> ?rhs" + apply auto + apply (rule_tac x = 1 in exI) + apply auto + done + then have "x \<in> ?rhs" by auto + } + then have "S \<subseteq> ?rhs" by auto + then have "?lhs \<subseteq> ?rhs" + using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto + moreover + { + fix x + assume "x \<in> ?rhs" + then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" + by auto + then have "xx \<in> cone hull S" + using hull_subset[of S] by auto + then have "x \<in> ?lhs" + using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto + } + ultimately show ?thesis by auto +qed + +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 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" + using cone_iff[of S] assms by auto + then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` closure S = closure S)" + using closure_subset by (auto simp add: closure_scaleR) + then show ?thesis + using False cone_iff[of "closure S"] by auto +qed + + +subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close> + +definition affine_dependent :: "'a::real_vector set \<Rightarrow> bool" + where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))" + +lemma affine_dependent_subset: + "\<lbrakk>affine_dependent s; s \<subseteq> t\<rbrakk> \<Longrightarrow> affine_dependent t" +apply (simp add: affine_dependent_def Bex_def) +apply (blast dest: hull_mono [OF Diff_mono [OF _ subset_refl]]) +done + +lemma affine_independent_subset: + shows "\<lbrakk>~ affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> ~ affine_dependent s" +by (metis affine_dependent_subset) + +lemma affine_independent_Diff: + "~ affine_dependent s \<Longrightarrow> ~ affine_dependent(s - t)" +by (meson Diff_subset affine_dependent_subset) + +lemma affine_dependent_explicit: + "affine_dependent p \<longleftrightarrow> + (\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> + (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)" + unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq + apply rule + apply (erule bexE, erule exE, erule exE) + apply (erule conjE)+ + defer + apply (erule exE, erule exE) + apply (erule conjE)+ + apply (erule bexE) +proof - + fix x s u + assume as: "x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" + have "x \<notin> s" using as(1,4) by auto + show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0" + apply (rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI) + unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF \<open>x\<notin>s\<close>] and as + using as + apply auto + done +next + fix s u v + assume as: "finite s" "s \<subseteq> p" "setsum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0" + have "s \<noteq> {v}" + using as(3,6) by auto + then show "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" + apply (rule_tac x=v in bexI) + apply (rule_tac x="s - {v}" in exI) + apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI) + unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric] + unfolding setsum_right_distrib[symmetric] and setsum_diff1[OF as(1)] + using as + apply auto + done +qed + +lemma affine_dependent_explicit_finite: + fixes s :: "'a::real_vector set" + assumes "finite s" + shows "affine_dependent s \<longleftrightarrow> + (\<exists>u. setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)" + (is "?lhs = ?rhs") +proof + have *: "\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)" + by auto + assume ?lhs + then obtain t u v where + "finite t" "t \<subseteq> s" "setsum u t = 0" "v\<in>t" "u v \<noteq> 0" "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0" + unfolding affine_dependent_explicit by auto + then show ?rhs + apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) + apply auto unfolding * and setsum.inter_restrict[OF assms, symmetric] + unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>] + apply auto + done +next + assume ?rhs + then obtain u v where "setsum u s = 0" "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" + by auto + then show ?lhs unfolding affine_dependent_explicit + using assms by auto +qed + + +subsection \<open>Connectedness of convex sets\<close> + +lemma connectedD: + "connected S \<Longrightarrow> open A \<Longrightarrow> open B \<Longrightarrow> S \<subseteq> A \<union> B \<Longrightarrow> A \<inter> B \<inter> S = {} \<Longrightarrow> A \<inter> S = {} \<or> B \<inter> S = {}" + by (rule Topological_Spaces.topological_space_class.connectedD) + +lemma convex_connected: + fixes s :: "'a::real_normed_vector set" + assumes "convex s" + shows "connected s" +proof (rule connectedI) + fix A B + assume "open A" "open B" "A \<inter> B \<inter> s = {}" "s \<subseteq> A \<union> B" + moreover + assume "A \<inter> s \<noteq> {}" "B \<inter> s \<noteq> {}" + then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto + define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u + then have "continuous_on {0 .. 1} f" + by (auto intro!: continuous_intros) + then have "connected (f ` {0 .. 1})" + by (auto intro!: connected_continuous_image) + note connectedD[OF this, of A B] + moreover have "a \<in> A \<inter> f ` {0 .. 1}" + using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def) + moreover have "b \<in> B \<inter> f ` {0 .. 1}" + using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def) + moreover have "f ` {0 .. 1} \<subseteq> s" + using \<open>convex s\<close> a b unfolding convex_def f_def by auto + ultimately show False by auto +qed + +corollary connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)" + by(simp add: convex_connected) + +proposition clopen: + fixes s :: "'a :: real_normed_vector set" + shows "closed s \<and> open s \<longleftrightarrow> s = {} \<or> s = UNIV" +apply (rule iffI) + apply (rule connected_UNIV [unfolded connected_clopen, rule_format]) + apply (force simp add: open_openin closed_closedin, force) +done + +corollary compact_open: + fixes s :: "'a :: euclidean_space set" + shows "compact s \<and> open s \<longleftrightarrow> s = {}" + by (auto simp: compact_eq_bounded_closed clopen) + +corollary finite_imp_not_open: + fixes S :: "'a::{real_normed_vector, perfect_space} set" + shows "\<lbrakk>finite S; open S\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> interior S = {}" + by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open) + +text \<open>Balls, being convex, are connected.\<close> + +lemma convex_prod: + assumes "\<And>i. i \<in> Basis \<Longrightarrow> convex {x. P i x}" + shows "convex {x. \<forall>i\<in>Basis. P i (x\<bullet>i)}" + using assms unfolding convex_def + by (auto simp: inner_add_left) + +lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}" + by (rule convex_prod) (simp add: atLeast_def[symmetric] convex_real_interval) + +lemma convex_local_global_minimum: + fixes s :: "'a::real_normed_vector set" + assumes "e > 0" + and "convex_on s f" + and "ball x e \<subseteq> s" + and "\<forall>y\<in>ball x e. f x \<le> f y" + shows "\<forall>y\<in>s. f x \<le> f y" +proof (rule ccontr) + have "x \<in> s" using assms(1,3) by auto + assume "\<not> ?thesis" + then obtain y where "y\<in>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 \<le> 1" and u: "u < e / dist x y" + using real_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto + then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y" + using \<open>x\<in>s\<close> \<open>y\<in>s\<close> + 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 \<in> ball x e" + unfolding mem_ball dist_norm + unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>] + unfolding dist_norm[symmetric] + using u + unfolding pos_less_divide_eq[OF xy] + by auto + then have "f x \<le> 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 \<open>u>0\<close>] + 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 add: convex_def) + fix y z + assume yz: "dist x y < e" "dist x z < e" + fix u v :: real + assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" + have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> 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 \<le> e" "dist x z \<le> e" + fix u v :: real + assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" + have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> 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) \<le> e" + using convex_bound_le[OF yz uv] by auto + } + then show ?thesis by (auto simp add: 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 + + +subsection \<open>Convex hull\<close> + +lemma convex_convex_hull [iff]: "convex (convex hull s)" + unfolding hull_def + using convex_Inter[of "{t. convex t \<and> s \<subseteq> t}"] + by auto + +lemma convex_hull_subset: + "s \<subseteq> convex hull t \<Longrightarrow> convex hull s \<subseteq> convex hull t" + by (simp add: convex_convex_hull subset_hull) + +lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s" + by (metis convex_convex_hull hull_same) + +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: "\<forall>x\<in>s. norm x \<le> B" + unfolding bounded_iff by auto + show ?thesis + apply (rule bounded_subset[OF bounded_cball, of _ 0 B]) + unfolding subset_hull[of convex, OF convex_cball] + unfolding subset_eq mem_cball dist_norm using B + apply auto + done +qed + +lemma finite_imp_bounded_convex_hull: + fixes s :: "'a::real_normed_vector set" + shows "finite s \<Longrightarrow> bounded (convex hull s)" + using bounded_convex_hull finite_imp_bounded + by auto + + +subsubsection \<open>Convex hull is "preserved" by a linear function\<close> + +lemma convex_hull_linear_image: + assumes f: "linear f" + shows "f ` (convex hull s) = convex hull (f ` s)" +proof + show "convex hull (f ` s) \<subseteq> f ` (convex hull s)" + by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull) + show "f ` (convex hull s) \<subseteq> convex hull (f ` s)" + proof (unfold image_subset_iff_subset_vimage, rule hull_minimal) + show "s \<subseteq> f -` (convex hull (f ` s))" + by (fast intro: hull_inc) + show "convex (f -` (convex hull (f ` s)))" + by (intro convex_linear_vimage [OF f] convex_convex_hull) + qed +qed + +lemma in_convex_hull_linear_image: + assumes "linear f" + and "x \<in> convex hull s" + shows "f x \<in> convex hull (f ` s)" + using convex_hull_linear_image[OF assms(1)] assms(2) by auto + +lemma convex_hull_Times: + "convex hull (s \<times> t) = (convex hull s) \<times> (convex hull t)" +proof + show "convex hull (s \<times> t) \<subseteq> (convex hull s) \<times> (convex hull t)" + by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull) + have "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)" + proof (intro hull_induct) + fix x y assume "x \<in> s" and "y \<in> t" + then show "(x, y) \<in> convex hull (s \<times> t)" + by (simp add: hull_inc) + next + fix x let ?S = "((\<lambda>y. (0, y)) -` (\<lambda>p. (- x, 0) + p) ` (convex hull s \<times> t))" + have "convex ?S" + by (intro convex_linear_vimage convex_translation convex_convex_hull, + simp add: linear_iff) + also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}" + by (auto simp add: image_def Bex_def) + finally show "convex {y. (x, y) \<in> convex hull (s \<times> t)}" . + next + show "convex {x. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)}" + proof (unfold Collect_ball_eq, rule convex_INT [rule_format]) + fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull s \<times> t))" + have "convex ?S" + by (intro convex_linear_vimage convex_translation convex_convex_hull, + simp add: linear_iff) + also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}" + by (auto simp add: image_def Bex_def) + finally show "convex {x. (x, y) \<in> convex hull (s \<times> t)}" . + qed + qed + then show "(convex hull s) \<times> (convex hull t) \<subseteq> convex hull (s \<times> t)" + unfolding subset_eq split_paired_Ball_Sigma . +qed + + +subsubsection \<open>Stepping theorems for convex hulls of finite sets\<close> + +lemma convex_hull_empty[simp]: "convex hull {} = {}" + by (rule hull_unique) auto + +lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" + by (rule hull_unique) auto + +lemma convex_hull_insert: + fixes s :: "'a::real_vector set" + assumes "s \<noteq> {}" + shows "convex hull (insert a s) = + {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and> b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}" + (is "_ = ?hull") + apply (rule, rule hull_minimal, rule) + unfolding insert_iff + prefer 3 + apply rule +proof - + fix x + assume x: "x = a \<or> x \<in> s" + then show "x \<in> ?hull" + apply rule + unfolding mem_Collect_eq + apply (rule_tac x=1 in exI) + defer + apply (rule_tac x=0 in exI) + using assms hull_subset[of s convex] + apply auto + done +next + fix x + assume "x \<in> ?hull" + then obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b" + by auto + have "a \<in> convex hull insert a s" "b \<in> convex hull insert a s" + using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4) + by auto + then show "x \<in> convex hull insert a s" + unfolding obt(5) using obt(1-3) + by (rule convexD [OF convex_convex_hull]) +next + show "convex ?hull" + proof (rule convexI) + fix x y u v + assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull" + from as(4) obtain u1 v1 b1 where + obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" + by auto + from as(5) obtain u2 v2 b2 where + obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" + by auto + have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" + by (auto simp add: algebra_simps) + have **: "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y = + (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" + proof (cases "u * v1 + v * v2 = 0") + case True + have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" + by (auto simp add: algebra_simps) + from True have ***: "u * v1 = 0" "v * v2 = 0" + using mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>] + by arith+ + then have "u * u1 + v * u2 = 1" + using as(3) obt1(3) obt2(3) by auto + then show ?thesis + unfolding obt1(5) obt2(5) * + using assms hull_subset[of s convex] + by (auto simp add: *** scaleR_right_distrib) + next + case False + have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" + using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) + also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" + using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) + also have "\<dots> = u * v1 + v * v2" + by simp + finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto + have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" + using as(1,2) obt1(1,2) obt2(1,2) by auto + then show ?thesis + unfolding obt1(5) obt2(5) + unfolding * and ** + using False + apply (rule_tac + x = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) + defer + apply (rule convexD [OF convex_convex_hull]) + using obt1(4) obt2(4) + unfolding add_divide_distrib[symmetric] and zero_le_divide_iff + apply (auto simp add: scaleR_left_distrib scaleR_right_distrib) + done + qed + have u1: "u1 \<le> 1" + unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto + have u2: "u2 \<le> 1" + unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto + have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v" + apply (rule add_mono) + apply (rule_tac [!] mult_right_mono) + using as(1,2) obt1(1,2) obt2(1,2) + apply auto + done + also have "\<dots> \<le> 1" + unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto + finally show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" + unfolding mem_Collect_eq + apply (rule_tac x="u * u1 + v * u2" in exI) + apply (rule conjI) + defer + apply (rule_tac x="1 - u * u1 - v * u2" in exI) + unfolding Bex_def + using as(1,2) obt1(1,2) obt2(1,2) ** + apply (auto simp add: algebra_simps) + done + qed +qed + + +subsubsection \<open>Explicit expression for convex hull\<close> + +lemma convex_hull_indexed: + fixes s :: "'a::real_vector set" + shows "convex hull s = + {y. \<exists>k u x. + (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and> + (setsum u {1..k} = 1) \<and> (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}" + (is "?xyz = ?hull") + apply (rule hull_unique) + apply rule + defer + apply (rule convexI) +proof - + fix x + assume "x\<in>s" + then show "x \<in> ?hull" + unfolding mem_Collect_eq + apply (rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI) + apply auto + done +next + fix t + assume as: "s \<subseteq> t" "convex t" + show "?hull \<subseteq> t" + apply rule + unfolding mem_Collect_eq + apply (elim exE conjE) + proof - + fix x k u y + assume assm: + "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s" + "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x" + show "x\<in>t" + unfolding assm(3) [symmetric] + apply (rule as(2)[unfolded convex, rule_format]) + using assm(1,2) as(1) apply auto + done + qed +next + fix x y u v + assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)" + assume xy: "x \<in> ?hull" "y \<in> ?hull" + from xy obtain k1 u1 x1 where + x: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "setsum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x" + by auto + from xy obtain k2 u2 x2 where + y: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "setsum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" + by auto + have *: "\<And>P (x1::'a) x2 s1 s2 i. + (if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)" + "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}" + prefer 3 + apply (rule, rule) + unfolding image_iff + apply (rule_tac x = "x - k1" in bexI) + apply (auto simp add: not_le) + done + have inj: "inj_on (\<lambda>i. i + k1) {1..k2}" + unfolding inj_on_def by auto + show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" + apply rule + apply (rule_tac x="k1 + k2" in exI) + apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI) + apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI) + apply (rule, rule) + defer + apply rule + unfolding * and setsum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and + setsum.reindex[OF inj] and o_def Collect_mem_eq + unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_right_distrib[symmetric] + proof - + fix i + assume i: "i \<in> {1..k1+k2}" + show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and> + (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s" + proof (cases "i\<in>{1..k1}") + case True + then show ?thesis + using uv(1) x(1)[THEN bspec[where x=i]] by auto + next + case False + define j where "j = i - k1" + from i False have "j \<in> {1..k2}" + unfolding j_def by auto + then show ?thesis + using False uv(2) y(1)[THEN bspec[where x=j]] + by (auto simp: j_def[symmetric]) + qed + qed (auto simp add: not_le x(2,3) y(2,3) uv(3)) +qed + +lemma convex_hull_finite: + fixes s :: "'a::real_vector set" + assumes "finite s" + shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> + setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}" + (is "?HULL = ?set") +proof (rule hull_unique, auto simp add: convex_def[of ?set]) + fix x + assume "x \<in> s" + then show "\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x" + apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) + apply auto + unfolding setsum.delta'[OF assms] and setsum_delta''[OF assms] + apply auto + done +next + fix u v :: real + assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" + fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)" + fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)" + { + fix x + assume "x\<in>s" + then have "0 \<le> u * ux x + v * uy x" + using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2) + by auto + } + moreover + have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1" + unfolding setsum.distrib and setsum_right_distrib[symmetric] and ux(2) uy(2) + using uv(3) by auto + moreover + have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)" + unfolding scaleR_left_distrib and setsum.distrib and scaleR_scaleR[symmetric] + and scaleR_right.setsum [symmetric] + by auto + ultimately + show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and> + (\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)" + apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) + apply auto + done +next + fix t + assume t: "s \<subseteq> t" "convex t" + fix u + assume u: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)" + then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" + using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]] + using assms and t(1) by auto +qed + + +subsubsection \<open>Another formulation from Lars Schewe\<close> + +lemma convex_hull_explicit: + fixes p :: "'a::real_vector set" + shows "convex hull p = + {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}" + (is "?lhs = ?rhs") +proof - + { + fix x + assume "x\<in>?lhs" + then obtain k u y where + obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x" + unfolding convex_hull_indexed by auto + + have fin: "finite {1..k}" by auto + have fin': "\<And>v. finite {i \<in> {1..k}. y i = v}" by auto + { + fix j + assume "j\<in>{1..k}" + then have "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}" + using obt(1)[THEN bspec[where x=j]] and obt(2) + apply simp + apply (rule setsum_nonneg) + using obt(1) + apply auto + done + } + moreover + have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1" + unfolding setsum_image_gen[OF fin, symmetric] using obt(2) by auto + moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x" + using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric] + unfolding scaleR_left.setsum using obt(3) by auto + ultimately + have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" + apply (rule_tac x="y ` {1..k}" in exI) + apply (rule_tac x="\<lambda>v. setsum u {i\<in>{1..k}. y i = v}" in exI) + apply auto + done + then have "x\<in>?rhs" by auto + } + moreover + { + fix y + assume "y\<in>?rhs" + then obtain s u where + obt: "finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" + by auto + + obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s" + using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto + + { + fix i :: nat + assume "i\<in>{1..card s}" + then have "f i \<in> s" + apply (subst f(2)[symmetric]) + apply auto + done + then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto + } + moreover have *: "finite {1..card s}" by auto + { + fix y + assume "y\<in>s" + then obtain i where "i\<in>{1..card s}" "f i = y" + using f using image_iff[of y f "{1..card s}"] + by auto + then have "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}" + apply auto + using f(1)[unfolded inj_on_def] + apply(erule_tac x=x in ballE) + apply auto + done + then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto + then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y" + "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" + by (auto simp add: setsum_constant_scaleR) + } + then have "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y" + unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] + and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f] + unfolding f + using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"] + using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u] + unfolding obt(4,5) + by auto + ultimately + have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> setsum u {1..k} = 1 \<and> + (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y" + apply (rule_tac x="card s" in exI) + apply (rule_tac x="u \<circ> f" in exI) + apply (rule_tac x=f in exI) + apply fastforce + done + then have "y \<in> ?lhs" + unfolding convex_hull_indexed by auto + } + ultimately show ?thesis + unfolding set_eq_iff by blast +qed + + +subsubsection \<open>A stepping theorem for that expansion\<close> + +lemma convex_hull_finite_step: + fixes s :: "'a::real_vector set" + assumes "finite s" + shows + "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) + \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" + (is "?lhs = ?rhs") +proof (rule, case_tac[!] "a\<in>s") + assume "a \<in> s" + then have *: "insert a s = s" by auto + assume ?lhs + then show ?rhs + unfolding * + apply (rule_tac x=0 in exI) + apply auto + done +next + assume ?lhs + then obtain u where + u: "\<forall>x\<in>insert a s. 0 \<le> u x" "setsum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" + by auto + assume "a \<notin> s" + then show ?rhs + apply (rule_tac x="u a" in exI) + using u(1)[THEN bspec[where x=a]] + apply simp + apply (rule_tac x=u in exI) + using u[unfolded setsum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close> + apply auto + done +next + assume "a \<in> s" + then have *: "insert a s = s" by auto + have fin: "finite (insert a s)" using assms by auto + assume ?rhs + then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" + by auto + show ?lhs + apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI) + unfolding scaleR_left_distrib and setsum.distrib and setsum_delta''[OF fin] and setsum.delta'[OF fin] + unfolding setsum_clauses(2)[OF assms] + using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close> + apply auto + done +next + assume ?rhs + then obtain v u where + uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" + by auto + moreover + assume "a \<notin> s" + moreover + have "(\<Sum>x\<in>s. if a = x then v else u x) = setsum u s" + and "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" + apply (rule_tac setsum.cong) apply rule + defer + apply (rule_tac setsum.cong) apply rule + using \<open>a \<notin> s\<close> + apply auto + done + ultimately show ?lhs + apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI) + unfolding setsum_clauses(2)[OF assms] + apply auto + done +qed + + +subsubsection \<open>Hence some special cases\<close> + +lemma convex_hull_2: + "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}" +proof - + have *: "\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b" + by auto + have **: "finite {b}" by auto + show ?thesis + apply (simp add: convex_hull_finite) + unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc] + apply auto + apply (rule_tac x=v in exI) + apply (rule_tac x="1 - v" in exI) + apply simp + apply (rule_tac x=u in exI) + apply simp + apply (rule_tac x="\<lambda>x. v" in exI) + apply simp + done +qed + +lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u. 0 \<le> u \<and> u \<le> 1}" + unfolding convex_hull_2 +proof (rule Collect_cong) + have *: "\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y" + by auto + fix x + show "(\<exists>v u. x = v *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> v \<and> 0 \<le> u \<and> v + u = 1) \<longleftrightarrow> + (\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)" + unfolding * + apply auto + apply (rule_tac[!] x=u in exI) + apply (auto simp add: algebra_simps) + done +qed + +lemma convex_hull_3: + "convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \<le> u \<and> 0 \<le> v \<and> 0 \<le> w \<and> u + v + w = 1}" +proof - + have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}" + by auto + have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" + by (auto simp add: field_simps) + show ?thesis + unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and * + unfolding convex_hull_finite_step[OF fin(3)] + apply (rule Collect_cong) + apply simp + apply auto + apply (rule_tac x=va in exI) + apply (rule_tac x="u c" in exI) + apply simp + apply (rule_tac x="1 - v - w" in exI) + apply simp + apply (rule_tac x=v in exI) + apply simp + apply (rule_tac x="\<lambda>x. w" in exI) + apply simp + done +qed + +lemma convex_hull_3_alt: + "convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v \<le> 1}" +proof - + have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" + by auto + show ?thesis + unfolding convex_hull_3 + apply (auto simp add: *) + apply (rule_tac x=v in exI) + apply (rule_tac x=w in exI) + apply (simp add: algebra_simps) + apply (rule_tac x=u in exI) + apply (rule_tac x=v in exI) + apply (simp add: algebra_simps) + done +qed + + +subsection \<open>Relations among closure notions and corresponding hulls\<close> + +lemma affine_imp_convex: "affine s \<Longrightarrow> convex s" + unfolding affine_def convex_def by auto + +lemma subspace_imp_convex: "subspace s \<Longrightarrow> convex s" + using subspace_imp_affine affine_imp_convex by auto + +lemma affine_hull_subset_span: "(affine hull s) \<subseteq> (span s)" + by (metis hull_minimal span_inc subspace_imp_affine subspace_span) + +lemma convex_hull_subset_span: "(convex hull s) \<subseteq> (span s)" + by (metis hull_minimal span_inc subspace_imp_convex subspace_span) + +lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)" + by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset) + +lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s" + unfolding affine_dependent_def dependent_def + using affine_hull_subset_span by auto + +lemma dependent_imp_affine_dependent: + assumes "dependent {x - a| x . x \<in> s}" + and "a \<notin> s" + shows "affine_dependent (insert a s)" +proof - + from assms(1)[unfolded dependent_explicit] obtain S u v + where obt: "finite S" "S \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0" + by auto + define t where "t = (\<lambda>x. x + a) ` S" + + have inj: "inj_on (\<lambda>x. x + a) S" + unfolding inj_on_def by auto + have "0 \<notin> S" + using obt(2) assms(2) unfolding subset_eq by auto + have fin: "finite t" and "t \<subseteq> s" + unfolding t_def using obt(1,2) by auto + then have "finite (insert a t)" and "insert a t \<subseteq> insert a s" + by auto + moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)" + apply (rule setsum.cong) + using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> + apply auto + done + have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0" + unfolding setsum_clauses(2)[OF fin] + using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> + apply auto + unfolding * + apply auto + done + moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0" + apply (rule_tac x="v + a" in bexI) + using obt(3,4) and \<open>0\<notin>S\<close> + unfolding t_def + apply auto + done + moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)" + apply (rule setsum.cong) + using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> + apply auto + done + have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)" + unfolding scaleR_left.setsum + unfolding t_def and setsum.reindex[OF inj] and o_def + using obt(5) + by (auto simp add: setsum.distrib scaleR_right_distrib) + then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0" + unfolding setsum_clauses(2)[OF fin] + using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> + by (auto simp add: *) + ultimately show ?thesis + unfolding affine_dependent_explicit + apply (rule_tac x="insert a t" in exI) + apply auto + done +qed + +lemma convex_cone: + "convex s \<and> cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. (x + y) \<in> s) \<and> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)" + (is "?lhs = ?rhs") +proof - + { + fix x y + assume "x\<in>s" "y\<in>s" and ?lhs + then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s" + unfolding cone_def by auto + then have "x + y \<in> s" + using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1] + apply (erule_tac x="2*\<^sub>R x" in ballE) + apply (erule_tac x="2*\<^sub>R y" in ballE) + apply (erule_tac x="1/2" in allE) + apply simp + apply (erule_tac x="1/2" in allE) + apply auto + done + } + then show ?thesis + unfolding convex_def cone_def by blast +qed + +lemma affine_dependent_biggerset: + fixes s :: "'a::euclidean_space set" + assumes "finite s" "card s \<ge> DIM('a) + 2" + shows "affine_dependent s" +proof - + have "s \<noteq> {}" using assms by auto + then obtain a where "a\<in>s" by auto + have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" + by auto + have "card {x - a |x. x \<in> s - {a}} = card (s - {a})" + unfolding * + apply (rule card_image) + unfolding inj_on_def + apply auto + done + also have "\<dots> > DIM('a)" using assms(2) + unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto + finally show ?thesis + apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) + apply (rule dependent_imp_affine_dependent) + apply (rule dependent_biggerset) + apply auto + done +qed + +lemma affine_dependent_biggerset_general: + assumes "finite (s :: 'a::euclidean_space set)" + and "card s \<ge> dim s + 2" + shows "affine_dependent s" +proof - + from assms(2) have "s \<noteq> {}" by auto + then obtain a where "a\<in>s" by auto + have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" + by auto + have **: "card {x - a |x. x \<in> s - {a}} = card (s - {a})" + unfolding * + apply (rule card_image) + unfolding inj_on_def + apply auto + done + have "dim {x - a |x. x \<in> s - {a}} \<le> dim s" + apply (rule subset_le_dim) + unfolding subset_eq + using \<open>a\<in>s\<close> + apply (auto simp add:span_superset span_sub) + done + also have "\<dots> < dim s + 1" by auto + also have "\<dots> \<le> card (s - {a})" + using assms + using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] + by auto + finally show ?thesis + apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) + apply (rule dependent_imp_affine_dependent) + apply (rule dependent_biggerset_general) + unfolding ** + apply auto + done +qed + + +subsection \<open>Some Properties of Affine Dependent Sets\<close> + +lemma affine_independent_0: "\<not> affine_dependent {}" + by (simp add: affine_dependent_def) + +lemma affine_independent_1: "\<not> affine_dependent {a}" + by (simp add: affine_dependent_def) + +lemma affine_independent_2: "\<not> affine_dependent {a,b}" + by (simp add: affine_dependent_def insert_Diff_if hull_same) + +lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (affine hull S)" +proof - + have "affine ((\<lambda>x. a + x) ` (affine hull S))" + using affine_translation affine_affine_hull by blast + moreover have "(\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" + using hull_subset[of S] by auto + ultimately have h1: "affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" + by (metis hull_minimal) + have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)))" + using affine_translation affine_affine_hull by blast + moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S))" + using hull_subset[of "(\<lambda>x. a + x) ` S"] by auto + moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S" + using translation_assoc[of "-a" a] by auto + ultimately have "(\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)) >= (affine hull S)" + by (metis hull_minimal) + then have "affine hull ((\<lambda>x. a + x) ` S) >= (\<lambda>x. a + x) ` (affine hull S)" + by auto + then show ?thesis using h1 by auto +qed + +lemma affine_dependent_translation: + assumes "affine_dependent S" + shows "affine_dependent ((\<lambda>x. a + x) ` S)" +proof - + obtain x where x: "x \<in> S \<and> x \<in> affine hull (S - {x})" + using assms affine_dependent_def by auto + have "op + a ` (S - {x}) = op + a ` S - {a + x}" + by auto + then have "a + x \<in> affine hull ((\<lambda>x. a + x) ` S - {a + x})" + using affine_hull_translation[of a "S - {x}"] x by auto + moreover have "a + x \<in> (\<lambda>x. a + x) ` S" + using x by auto + ultimately show ?thesis + unfolding affine_dependent_def by auto +qed + +lemma affine_dependent_translation_eq: + "affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)" +proof - + { + assume "affine_dependent ((\<lambda>x. a + x) ` S)" + then have "affine_dependent S" + using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a] + by auto + } + then show ?thesis + using affine_dependent_translation by auto +qed + +lemma affine_hull_0_dependent: + assumes "0 \<in> affine hull S" + shows "dependent S" +proof - + obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0" + using assms affine_hull_explicit[of S] by auto + then have "\<exists>v\<in>s. u v \<noteq> 0" + using setsum_not_0[of "u" "s"] by auto + then have "finite s \<and> s \<subseteq> S \<and> (\<exists>v\<in>s. u v \<noteq> 0 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0)" + using s_u by auto + then show ?thesis + unfolding dependent_explicit[of S] by auto +qed + +lemma affine_dependent_imp_dependent2: + assumes "affine_dependent (insert 0 S)" + shows "dependent S" +proof - + obtain x where x: "x \<in> insert 0 S \<and> x \<in> affine hull (insert 0 S - {x})" + using affine_dependent_def[of "(insert 0 S)"] assms by blast + then have "x \<in> span (insert 0 S - {x})" + using affine_hull_subset_span by auto + moreover have "span (insert 0 S - {x}) = span (S - {x})" + using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto + ultimately have "x \<in> span (S - {x})" by auto + then have "x \<noteq> 0 \<Longrightarrow> dependent S" + using x dependent_def by auto + moreover + { + assume "x = 0" + then have "0 \<in> affine hull S" + using x hull_mono[of "S - {0}" S] by auto + then have "dependent S" + using affine_hull_0_dependent by auto + } + ultimately show ?thesis by auto +qed + +lemma affine_dependent_iff_dependent: + assumes "a \<notin> S" + shows "affine_dependent (insert a S) \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` S)" +proof - + have "(op + (- a) ` S) = {x - a| x . x : S}" by auto + then show ?thesis + using affine_dependent_translation_eq[of "(insert a S)" "-a"] + affine_dependent_imp_dependent2 assms + dependent_imp_affine_dependent[of a S] + by (auto simp del: uminus_add_conv_diff) +qed + +lemma affine_dependent_iff_dependent2: + assumes "a \<in> S" + shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))" +proof - + have "insert a (S - {a}) = S" + using assms by auto + then show ?thesis + using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto +qed + +lemma affine_hull_insert_span_gen: + "affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)" +proof - + have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)" + by auto + { + assume "a \<notin> s" + then have ?thesis + using affine_hull_insert_span[of a s] h1 by auto + } + moreover + { + assume a1: "a \<in> s" + have "\<exists>x. x \<in> s \<and> -a+x=0" + apply (rule exI[of _ a]) + using a1 + apply auto + done + then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s" + by auto + then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)" + using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff) + moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))" + by auto + moreover have "insert a (s - {a}) = insert a s" + by auto + ultimately have ?thesis + using affine_hull_insert_span[of "a" "s-{a}"] by auto + } + ultimately show ?thesis by auto +qed + +lemma affine_hull_span2: + assumes "a \<in> s" + shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))" + using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]] + by auto + +lemma affine_hull_span_gen: + assumes "a \<in> affine hull s" + shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)" +proof - + have "affine hull (insert a s) = affine hull s" + using hull_redundant[of a affine s] assms by auto + then show ?thesis + using affine_hull_insert_span_gen[of a "s"] by auto +qed + +lemma affine_hull_span_0: + assumes "0 \<in> affine hull S" + shows "affine hull S = span S" + using affine_hull_span_gen[of "0" S] assms by auto + +lemma extend_to_affine_basis_nonempty: + fixes S V :: "'n::euclidean_space set" + assumes "\<not> affine_dependent S" "S \<subseteq> V" "S \<noteq> {}" + shows "\<exists>T. \<not> affine_dependent T \<and> S \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V" +proof - + obtain a where a: "a \<in> S" + using assms by auto + then have h0: "independent ((\<lambda>x. -a + x) ` (S-{a}))" + using affine_dependent_iff_dependent2 assms by auto + then obtain B where B: + "(\<lambda>x. -a+x) ` (S - {a}) \<subseteq> B \<and> B \<subseteq> (\<lambda>x. -a+x) ` V \<and> independent B \<and> (\<lambda>x. -a+x) ` V \<subseteq> span B" + using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms + by blast + define T where "T = (\<lambda>x. a+x) ` insert 0 B" + then have "T = insert a ((\<lambda>x. a+x) ` B)" + by auto + then have "affine hull T = (\<lambda>x. a+x) ` span B" + using affine_hull_insert_span_gen[of a "((\<lambda>x. a+x) ` B)"] translation_assoc[of "-a" a B] + by auto + then have "V \<subseteq> affine hull T" + using B assms translation_inverse_subset[of a V "span B"] + by auto + moreover have "T \<subseteq> V" + using T_def B a assms by auto + ultimately have "affine hull T = affine hull V" + by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono) + moreover have "S \<subseteq> T" + using T_def B translation_inverse_subset[of a "S-{a}" B] + by auto + moreover have "\<not> affine_dependent T" + using T_def affine_dependent_translation_eq[of "insert 0 B"] + affine_dependent_imp_dependent2 B + by auto + ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto +qed + +lemma affine_basis_exists: + fixes V :: "'n::euclidean_space set" + shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B" +proof (cases "V = {}") + case True + then show ?thesis + using affine_independent_0 by auto +next + case False + then obtain x where "x \<in> V" by auto + then show ?thesis + using affine_dependent_def[of "{x}"] extend_to_affine_basis_nonempty[of "{x}" V] + by auto +qed + +proposition extend_to_affine_basis: + fixes S V :: "'n::euclidean_space set" + assumes "\<not> affine_dependent S" "S \<subseteq> V" + obtains T where "\<not> affine_dependent T" "S \<subseteq> T" "T \<subseteq> V" "affine hull T = affine hull V" +proof (cases "S = {}") + case True then show ?thesis + using affine_basis_exists by (metis empty_subsetI that) +next + case False + then show ?thesis by (metis assms extend_to_affine_basis_nonempty that) +qed + + +subsection \<open>Affine Dimension of a Set\<close> + +definition aff_dim :: "('a::euclidean_space) set \<Rightarrow> int" + where "aff_dim V = + (SOME d :: int. + \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)" + +lemma aff_dim_basis_exists: + fixes V :: "('n::euclidean_space) set" + shows "\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" +proof - + obtain B where "\<not> affine_dependent B \<and> affine hull B = affine hull V" + using affine_basis_exists[of V] by auto + then show ?thesis + unfolding aff_dim_def + some_eq_ex[of "\<lambda>d. \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1"] + apply auto + apply (rule exI[of _ "int (card B) - (1 :: int)"]) + apply (rule exI[of _ "B"]) + apply auto + done +qed + +lemma affine_hull_nonempty: "S \<noteq> {} \<longleftrightarrow> affine hull S \<noteq> {}" +proof - + have "S = {} \<Longrightarrow> affine hull S = {}" + using affine_hull_empty by auto + moreover have "affine hull S = {} \<Longrightarrow> S = {}" + unfolding hull_def by auto + ultimately show ?thesis by blast +qed + +lemma aff_dim_parallel_subspace_aux: + fixes B :: "'n::euclidean_space set" + assumes "\<not> affine_dependent B" "a \<in> B" + shows "finite B \<and> ((card B) - 1 = dim (span ((\<lambda>x. -a+x) ` (B-{a}))))" +proof - + have "independent ((\<lambda>x. -a + x) ` (B-{a}))" + using affine_dependent_iff_dependent2 assms by auto + then have fin: "dim (span ((\<lambda>x. -a+x) ` (B-{a}))) = card ((\<lambda>x. -a + x) ` (B-{a}))" + "finite ((\<lambda>x. -a + x) ` (B - {a}))" + using indep_card_eq_dim_span[of "(\<lambda>x. -a+x) ` (B-{a})"] by auto + show ?thesis + proof (cases "(\<lambda>x. -a + x) ` (B - {a}) = {}") + case True + have "B = insert a ((\<lambda>x. a + x) ` (\<lambda>x. -a + x) ` (B - {a}))" + using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto + then have "B = {a}" using True by auto + then show ?thesis using assms fin by auto + next + case False + then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0" + using fin by auto + moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})" + apply (rule card_image) + using translate_inj_on + apply (auto simp del: uminus_add_conv_diff) + done + ultimately have "card (B-{a}) > 0" by auto + then have *: "finite (B - {a})" + using card_gt_0_iff[of "(B - {a})"] by auto + then have "card (B - {a}) = card B - 1" + using card_Diff_singleton assms by auto + with * show ?thesis using fin h1 by auto + qed +qed + +lemma aff_dim_parallel_subspace: + fixes V L :: "'n::euclidean_space set" + assumes "V \<noteq> {}" + and "subspace L" + and "affine_parallel (affine hull V) L" + shows "aff_dim V = int (dim L)" +proof - + obtain B where + B: "affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> int (card B) = aff_dim V + 1" + using aff_dim_basis_exists by auto + then have "B \<noteq> {}" + using assms B affine_hull_nonempty[of V] affine_hull_nonempty[of B] + by auto + then obtain a where a: "a \<in> B" by auto + define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))" + moreover have "affine_parallel (affine hull B) Lb" + using Lb_def B assms affine_hull_span2[of a B] a + affine_parallel_commut[of "Lb" "(affine hull B)"] + unfolding affine_parallel_def + by auto + moreover have "subspace Lb" + using Lb_def subspace_span by auto + moreover have "affine hull B \<noteq> {}" + using assms B affine_hull_nonempty[of V] by auto + ultimately have "L = Lb" + using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B + by auto + then have "dim L = dim Lb" + by auto + moreover have "card B - 1 = dim Lb" and "finite B" + using Lb_def aff_dim_parallel_subspace_aux a B by auto + ultimately show ?thesis + using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto +qed + +lemma aff_independent_finite: + fixes B :: "'n::euclidean_space set" + assumes "\<not> affine_dependent B" + shows "finite B" +proof - + { + assume "B \<noteq> {}" + then obtain a where "a \<in> B" by auto + then have ?thesis + using aff_dim_parallel_subspace_aux assms by auto + } + then show ?thesis by auto +qed + +lemma independent_finite: + fixes B :: "'n::euclidean_space set" + assumes "independent B" + shows "finite B" + using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms + by auto + +lemma subspace_dim_equal: + assumes "subspace (S :: ('n::euclidean_space) set)" + and "subspace T" + and "S \<subseteq> T" + and "dim S \<ge> dim T" + shows "S = T" +proof - + obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S" + using basis_exists[of S] by auto + then have "span B \<subseteq> S" + using span_mono[of B S] span_eq[of S] assms by metis + then have "span B = S" + using B by auto + have "dim S = dim T" + using assms dim_subset[of S T] by auto + then have "T \<subseteq> span B" + using card_eq_dim[of B T] B independent_finite assms by auto + then show ?thesis + using assms \<open>span B = S\<close> by auto +qed + +corollary dim_eq_span: + fixes S :: "'a::euclidean_space set" + shows "\<lbrakk>S \<subseteq> T; dim T \<le> dim S\<rbrakk> \<Longrightarrow> span S = span T" +by (simp add: span_mono subspace_dim_equal subspace_span) + +lemma dim_eq_full: + fixes S :: "'a :: euclidean_space set" + shows "dim S = DIM('a) \<longleftrightarrow> span S = UNIV" +apply (rule iffI) + apply (metis dim_eq_span dim_subset_UNIV span_Basis span_span subset_UNIV) +by (metis dim_UNIV dim_span) + +lemma span_substd_basis: + assumes d: "d \<subseteq> Basis" + shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}" + (is "_ = ?B") +proof - + have "d \<subseteq> ?B" + using d by (auto simp: inner_Basis) + moreover have s: "subspace ?B" + using subspace_substandard[of "\<lambda>i. i \<notin> d"] . + ultimately have "span d \<subseteq> ?B" + using span_mono[of d "?B"] span_eq[of "?B"] by blast + moreover have *: "card d \<le> dim (span d)" + using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_inc[of d] + by auto + moreover from * have "dim ?B \<le> dim (span d)" + using dim_substandard[OF assms] by auto + ultimately show ?thesis + using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto +qed + +lemma basis_to_substdbasis_subspace_isomorphism: + fixes B :: "'a::euclidean_space set" + assumes "independent B" + shows "\<exists>f d::'a set. card d = card B \<and> linear f \<and> f ` B = d \<and> + f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} \<and> inj_on f (span B) \<and> d \<subseteq> Basis" +proof - + have B: "card B = dim B" + using dim_unique[of B B "card B"] assms span_inc[of B] by auto + have "dim B \<le> card (Basis :: 'a set)" + using dim_subset_UNIV[of B] by simp + from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B" + by auto + let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}" + have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)" + apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"]) + apply (rule subspace_span) + apply (rule subspace_substandard) + defer + apply (rule span_inc) + apply (rule assms) + defer + unfolding dim_span[of B] + apply(rule B) + unfolding span_substd_basis[OF d, symmetric] + apply (rule span_inc) + apply (rule independent_substdbasis[OF d]) + apply rule + apply assumption + unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] + apply auto + done + with t \<open>card B = dim B\<close> d show ?thesis by auto +qed + +lemma aff_dim_empty: + fixes S :: "'n::euclidean_space set" + shows "S = {} \<longleftrightarrow> aff_dim S = -1" +proof - + obtain B where *: "affine hull B = affine hull S" + and "\<not> affine_dependent B" + and "int (card B) = aff_dim S + 1" + using aff_dim_basis_exists by auto + moreover + from * have "S = {} \<longleftrightarrow> B = {}" + using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto + ultimately show ?thesis + using aff_independent_finite[of B] card_gt_0_iff[of B] by auto +qed + +lemma aff_dim_empty_eq [simp]: "aff_dim ({}::'a::euclidean_space set) = -1" + by (simp add: aff_dim_empty [symmetric]) + +lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S" + unfolding aff_dim_def using hull_hull[of _ S] by auto + +lemma aff_dim_affine_hull2: + assumes "affine hull S = affine hull T" + shows "aff_dim S = aff_dim T" + unfolding aff_dim_def using assms by auto + +lemma aff_dim_unique: + fixes B V :: "'n::euclidean_space set" + assumes "affine hull B = affine hull V \<and> \<not> affine_dependent B" + shows "of_nat (card B) = aff_dim V + 1" +proof (cases "B = {}") + case True + then have "V = {}" + using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms + by auto + then have "aff_dim V = (-1::int)" + using aff_dim_empty by auto + then show ?thesis + using \<open>B = {}\<close> by auto +next + case False + then obtain a where a: "a \<in> B" by auto + define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))" + have "affine_parallel (affine hull B) Lb" + using Lb_def affine_hull_span2[of a B] a + affine_parallel_commut[of "Lb" "(affine hull B)"] + unfolding affine_parallel_def by auto + moreover have "subspace Lb" + using Lb_def subspace_span by auto + ultimately have "aff_dim B = int(dim Lb)" + using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto + moreover have "(card B) - 1 = dim Lb" "finite B" + using Lb_def aff_dim_parallel_subspace_aux a assms by auto + ultimately have "of_nat (card B) = aff_dim B + 1" + using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto + then show ?thesis + using aff_dim_affine_hull2 assms by auto +qed + +lemma aff_dim_affine_independent: + fixes B :: "'n::euclidean_space set" + assumes "\<not> affine_dependent B" + shows "of_nat (card B) = aff_dim B + 1" + using aff_dim_unique[of B B] assms by auto + +lemma affine_independent_iff_card: + fixes s :: "'a::euclidean_space set" + shows "~ affine_dependent s \<longleftrightarrow> finite s \<and> aff_dim s = int(card s) - 1" + apply (rule iffI) + apply (simp add: aff_dim_affine_independent aff_independent_finite) + by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff) + +lemma aff_dim_sing [simp]: + fixes a :: "'n::euclidean_space" + shows "aff_dim {a} = 0" + using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto + +lemma aff_dim_inner_basis_exists: + fixes V :: "('n::euclidean_space) set" + shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and> + \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" +proof - + obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V" + using affine_basis_exists[of V] by auto + then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto + with B show ?thesis by auto +qed + +lemma aff_dim_le_card: + fixes V :: "'n::euclidean_space set" + assumes "finite V" + shows "aff_dim V \<le> of_nat (card V) - 1" +proof - + obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1" + using aff_dim_inner_basis_exists[of V] by auto + then have "card B \<le> card V" + using assms card_mono by auto + with B show ?thesis by auto +qed + +lemma aff_dim_parallel_eq: + fixes S T :: "'n::euclidean_space set" + assumes "affine_parallel (affine hull S) (affine hull T)" + shows "aff_dim S = aff_dim T" +proof - + { + assume "T \<noteq> {}" "S \<noteq> {}" + then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L" + using affine_parallel_subspace[of "affine hull T"] + affine_affine_hull[of T] affine_hull_nonempty + by auto + then have "aff_dim T = int (dim L)" + using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto + moreover have *: "subspace L \<and> affine_parallel (affine hull S) L" + using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto + moreover from * have "aff_dim S = int (dim L)" + using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto + ultimately have ?thesis by auto + } + moreover + { + assume "S = {}" + then have "S = {}" and "T = {}" + using assms affine_hull_nonempty + unfolding affine_parallel_def + by auto + then have ?thesis using aff_dim_empty by auto + } + moreover + { + assume "T = {}" + then have "S = {}" and "T = {}" + using assms affine_hull_nonempty + unfolding affine_parallel_def + by auto + then have ?thesis + using aff_dim_empty by auto + } + ultimately show ?thesis by blast +qed + +lemma aff_dim_translation_eq: + fixes a :: "'n::euclidean_space" + shows "aff_dim ((\<lambda>x. a + x) ` S) = aff_dim S" +proof - + have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))" + unfolding affine_parallel_def + apply (rule exI[of _ "a"]) + using affine_hull_translation[of a S] + apply auto + done + then show ?thesis + using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto +qed + +lemma aff_dim_affine: + fixes S L :: "'n::euclidean_space set" + assumes "S \<noteq> {}" + and "affine S" + and "subspace L" + and "affine_parallel S L" + shows "aff_dim S = int (dim L)" +proof - + have *: "affine hull S = S" + using assms affine_hull_eq[of S] by auto + then have "affine_parallel (affine hull S) L" + using assms by (simp add: *) + then show ?thesis + using assms aff_dim_parallel_subspace[of S L] by blast +qed + +lemma dim_affine_hull: + fixes S :: "'n::euclidean_space set" + shows "dim (affine hull S) = dim S" +proof - + have "dim (affine hull S) \<ge> dim S" + using dim_subset by auto + moreover have "dim (span S) \<ge> dim (affine hull S)" + using dim_subset affine_hull_subset_span by blast + moreover have "dim (span S) = dim S" + using dim_span by auto + ultimately show ?thesis by auto +qed + +lemma aff_dim_subspace: + fixes S :: "'n::euclidean_space set" + assumes "subspace S" + shows "aff_dim S = int (dim S)" +proof (cases "S={}") + case True with assms show ?thesis + by (simp add: subspace_affine) +next + case False + with aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] subspace_affine + show ?thesis by auto +qed + +lemma aff_dim_zero: + fixes S :: "'n::euclidean_space set" + assumes "0 \<in> affine hull S" + shows "aff_dim S = int (dim S)" +proof - + have "subspace (affine hull S)" + using subspace_affine[of "affine hull S"] affine_affine_hull assms + by auto + then have "aff_dim (affine hull S) = int (dim (affine hull S))" + using assms aff_dim_subspace[of "affine hull S"] by auto + then show ?thesis + using aff_dim_affine_hull[of S] dim_affine_hull[of S] + by auto +qed + +lemma aff_dim_eq_dim: + fixes S :: "'n::euclidean_space set" + assumes "a \<in> affine hull S" + shows "aff_dim S = int (dim ((\<lambda>x. -a+x) ` S))" +proof - + have "0 \<in> affine hull ((\<lambda>x. -a+x) ` S)" + unfolding Convex_Euclidean_Space.affine_hull_translation + using assms by (simp add: ab_group_add_class.ab_left_minus image_iff) + with aff_dim_zero show ?thesis + by (metis aff_dim_translation_eq) +qed + +lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))" + using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"] + dim_UNIV[where 'a="'n::euclidean_space"] + by auto + +lemma aff_dim_geq: + fixes V :: "'n::euclidean_space set" + shows "aff_dim V \<ge> -1" +proof - + obtain B where "affine hull B = affine hull V" + and "\<not> affine_dependent B" + and "int (card B) = aff_dim V + 1" + using aff_dim_basis_exists by auto + then show ?thesis by auto +qed + +lemma aff_dim_negative_iff [simp]: + fixes S :: "'n::euclidean_space set" + shows "aff_dim S < 0 \<longleftrightarrow>S = {}" +by (metis aff_dim_empty aff_dim_geq diff_0 eq_iff zle_diff1_eq) + +lemma affine_independent_card_dim_diffs: + fixes S :: "'a :: euclidean_space set" + assumes "~ affine_dependent S" "a \<in> S" + shows "card S = dim {x - a|x. x \<in> S} + 1" +proof - + have 1: "{b - a|b. b \<in> (S - {a})} \<subseteq> {x - a|x. x \<in> S}" by auto + have 2: "x - a \<in> span {b - a |b. b \<in> S - {a}}" if "x \<in> S" for x + proof (cases "x = a") + case True then show ?thesis by simp + next + case False then show ?thesis + using assms by (blast intro: span_superset that) + qed + have "\<not> affine_dependent (insert a S)" + by (simp add: assms insert_absorb) + then have 3: "independent {b - a |b. b \<in> S - {a}}" + using dependent_imp_affine_dependent by fastforce + have "{b - a |b. b \<in> S - {a}} = (\<lambda>b. b-a) ` (S - {a})" + by blast + then have "card {b - a |b. b \<in> S - {a}} = card ((\<lambda>b. b-a) ` (S - {a}))" + by simp + also have "... = card (S - {a})" + by (metis (no_types, lifting) card_image diff_add_cancel inj_onI) + also have "... = card S - 1" + by (simp add: aff_independent_finite assms) + finally have 4: "card {b - a |b. b \<in> S - {a}} = card S - 1" . + have "finite S" + by (meson assms aff_independent_finite) + with \<open>a \<in> S\<close> have "card S \<noteq> 0" by auto + moreover have "dim {x - a |x. x \<in> S} = card S - 1" + using 2 by (blast intro: dim_unique [OF 1 _ 3 4]) + ultimately show ?thesis + by auto +qed + +lemma independent_card_le_aff_dim: + fixes B :: "'n::euclidean_space set" + assumes "B \<subseteq> V" + assumes "\<not> affine_dependent B" + shows "int (card B) \<le> aff_dim V + 1" +proof - + obtain T where T: "\<not> affine_dependent T \<and> B \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V" + by (metis assms extend_to_affine_basis[of B V]) + then have "of_nat (card T) = aff_dim V + 1" + using aff_dim_unique by auto + then show ?thesis + using T card_mono[of T B] aff_independent_finite[of T] by auto +qed + +lemma aff_dim_subset: + fixes S T :: "'n::euclidean_space set" + assumes "S \<subseteq> T" + shows "aff_dim S \<le> aff_dim T" +proof - + obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S" + "of_nat (card B) = aff_dim S + 1" + using aff_dim_inner_basis_exists[of S] by auto + then have "int (card B) \<le> aff_dim T + 1" + using assms independent_card_le_aff_dim[of B T] by auto + with B show ?thesis by auto +qed + +lemma aff_dim_le_DIM: + fixes S :: "'n::euclidean_space set" + shows "aff_dim S \<le> int (DIM('n))" +proof - + have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))" + using aff_dim_UNIV by auto + then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))" + using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto +qed + +lemma affine_dim_equal: + fixes S :: "'n::euclidean_space set" + assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T" + shows "S = T" +proof - + obtain a where "a \<in> S" using assms by auto + then have "a \<in> T" using assms by auto + define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}" + then have ls: "subspace LS" "affine_parallel S LS" + using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto + then have h1: "int(dim LS) = aff_dim S" + using assms aff_dim_affine[of S LS] by auto + have "T \<noteq> {}" using assms by auto + define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}" + then have lt: "subspace LT \<and> affine_parallel T LT" + using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto + then have "int(dim LT) = aff_dim T" + using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto + then have "dim LS = dim LT" + using h1 assms by auto + moreover have "LS \<le> LT" + using LS_def LT_def assms by auto + ultimately have "LS = LT" + using subspace_dim_equal[of LS LT] ls lt by auto + moreover have "S = {x. \<exists>y \<in> LS. a+y=x}" + using LS_def by auto + moreover have "T = {x. \<exists>y \<in> LT. a+y=x}" + using LT_def by auto + ultimately show ?thesis by auto +qed + +lemma affine_hull_UNIV: + fixes S :: "'n::euclidean_space set" + assumes "aff_dim S = int(DIM('n))" + shows "affine hull S = (UNIV :: ('n::euclidean_space) set)" +proof - + have "S \<noteq> {}" + using assms aff_dim_empty[of S] by auto + have h0: "S \<subseteq> affine hull S" + using hull_subset[of S _] by auto + have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S" + using aff_dim_UNIV assms by auto + then have h2: "aff_dim (affine hull S) \<le> aff_dim (UNIV :: ('n::euclidean_space) set)" + using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto + have h3: "aff_dim S \<le> aff_dim (affine hull S)" + using h0 aff_dim_subset[of S "affine hull S"] assms by auto + then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)" + using h0 h1 h2 by auto + then show ?thesis + using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"] + affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close> + by auto +qed + +lemma disjoint_affine_hull: + fixes s :: "'n::euclidean_space set" + assumes "~ affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}" + shows "(affine hull t) \<inter> (affine hull u) = {}" +proof - + have "finite s" using assms by (simp add: aff_independent_finite) + then have "finite t" "finite u" using assms finite_subset by blast+ + { fix y + assume yt: "y \<in> affine hull t" and yu: "y \<in> affine hull u" + then obtain a b + where a1 [simp]: "setsum a t = 1" and [simp]: "setsum (\<lambda>v. a v *\<^sub>R v) t = y" + and [simp]: "setsum b u = 1" "setsum (\<lambda>v. b v *\<^sub>R v) u = y" + by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>) + define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x + have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto + have "setsum c s = 0" + by (simp add: c_def comm_monoid_add_class.setsum.If_cases \<open>finite s\<close> setsum_negf) + moreover have "~ (\<forall>v\<in>s. c v = 0)" + by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def setsum_not_0 zero_neq_one) + moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0" + by (simp add: c_def if_smult setsum_negf + comm_monoid_add_class.setsum.If_cases \<open>finite s\<close>) + ultimately have False + using assms \<open>finite s\<close> by (auto simp: affine_dependent_explicit) + } + then show ?thesis by blast +qed + +lemma aff_dim_convex_hull: + fixes S :: "'n::euclidean_space set" + shows "aff_dim (convex hull S) = aff_dim S" + using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S] + hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"] + aff_dim_subset[of "convex hull S" "affine hull S"] + by auto + +lemma aff_dim_cball: + fixes a :: "'n::euclidean_space" + assumes "e > 0" + shows "aff_dim (cball a e) = int (DIM('n))" +proof - + have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e" + unfolding cball_def dist_norm by auto + then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)" + using aff_dim_translation_eq[of a "cball 0 e"] + aff_dim_subset[of "op + a ` cball 0 e" "cball a e"] + by auto + moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))" + using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] + centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms + by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) + ultimately show ?thesis + using aff_dim_le_DIM[of "cball a e"] by auto +qed + +lemma aff_dim_open: + fixes S :: "'n::euclidean_space set" + assumes "open S" + and "S \<noteq> {}" + shows "aff_dim S = int (DIM('n))" +proof - + obtain x where "x \<in> S" + using assms by auto + then obtain e where e: "e > 0" "cball x e \<subseteq> S" + using open_contains_cball[of S] assms by auto + then have "aff_dim (cball x e) \<le> aff_dim S" + using aff_dim_subset by auto + with e show ?thesis + using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto +qed + +lemma low_dim_interior: + fixes S :: "'n::euclidean_space set" + assumes "\<not> aff_dim S = int (DIM('n))" + shows "interior S = {}" +proof - + have "aff_dim(interior S) \<le> aff_dim S" + using interior_subset aff_dim_subset[of "interior S" S] by auto + then show ?thesis + using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto +qed + +corollary empty_interior_lowdim: + fixes S :: "'n::euclidean_space set" + shows "dim S < DIM ('n) \<Longrightarrow> interior S = {}" +by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV) + +corollary aff_dim_nonempty_interior: + fixes S :: "'a::euclidean_space set" + shows "interior S \<noteq> {} \<Longrightarrow> aff_dim S = DIM('a)" +by (metis low_dim_interior) + +subsection \<open>Caratheodory's theorem.\<close> + +lemma convex_hull_caratheodory_aff_dim: + fixes p :: "('a::euclidean_space) set" + shows "convex hull p = + {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> + (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}" + unfolding convex_hull_explicit set_eq_iff mem_Collect_eq +proof (intro allI iffI) + fix y + let ?P = "\<lambda>n. \<exists>s u. finite s \<and> card s = n \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> + setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" + assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" + then obtain N where "?P N" by auto + then have "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n" + apply (rule_tac ex_least_nat_le) + apply auto + done + then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k" + by blast + then obtain s u where obt: "finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x" + "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto + + have "card s \<le> aff_dim p + 1" + proof (rule ccontr, simp only: not_le) + assume "aff_dim p + 1 < card s" + then have "affine_dependent s" + using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3) + by blast + then obtain w v where wv: "setsum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0" + using affine_dependent_explicit_finite[OF obt(1)] by auto + define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}" + define t where "t = Min i" + have "\<exists>x\<in>s. w x < 0" + proof (rule ccontr, simp add: not_less) + assume as:"\<forall>x\<in>s. 0 \<le> w x" + then have "setsum w (s - {v}) \<ge> 0" + apply (rule_tac setsum_nonneg) + apply auto + done + then have "setsum w s > 0" + unfolding setsum.remove[OF obt(1) \<open>v\<in>s\<close>] + using as[THEN bspec[where x=v]] \<open>v\<in>s\<close> \<open>w v \<noteq> 0\<close> by auto + then show False using wv(1) by auto + qed + then have "i \<noteq> {}" unfolding i_def by auto + then have "t \<ge> 0" + using Min_ge_iff[of i 0 ] and obt(1) + unfolding t_def i_def + using obt(4)[unfolded le_less] + by (auto simp: divide_le_0_iff) + have t: "\<forall>v\<in>s. u v + t * w v \<ge> 0" + proof + fix v + assume "v \<in> s" + then have v: "0 \<le> u v" + using obt(4)[THEN bspec[where x=v]] by auto + show "0 \<le> u v + t * w v" + proof (cases "w v < 0") + case False + thus ?thesis using v \<open>t\<ge>0\<close> by auto + next + case True + then have "t \<le> u v / (- w v)" + using \<open>v\<in>s\<close> unfolding t_def i_def + apply (rule_tac Min_le) + using obt(1) apply auto + done + then show ?thesis + unfolding real_0_le_add_iff + using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[symmetric]]] + by auto + qed + qed + obtain a where "a \<in> s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0" + using Min_in[OF _ \<open>i\<noteq>{}\<close>] and obt(1) unfolding i_def t_def by auto + then have a: "a \<in> s" "u a + t * w a = 0" by auto + have *: "\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)" + unfolding setsum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto + have "(\<Sum>v\<in>s. u v + t * w v) = 1" + unfolding setsum.distrib wv(1) setsum_right_distrib[symmetric] obt(5) by auto + moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" + unfolding setsum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4) + using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp + ultimately have "?P (n - 1)" + apply (rule_tac x="(s - {a})" in exI) + apply (rule_tac x="\<lambda>v. u v + t * w v" in exI) + using obt(1-3) and t and a + apply (auto simp add: * scaleR_left_distrib) + done + then show False + using smallest[THEN spec[where x="n - 1"]] by auto + qed + then show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> + (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" + using obt by auto +qed auto + +lemma caratheodory_aff_dim: + fixes p :: "('a::euclidean_space) set" + shows "convex hull p = {x. \<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> x \<in> convex hull s}" + (is "?lhs = ?rhs") +proof + show "?lhs \<subseteq> ?rhs" + apply (subst convex_hull_caratheodory_aff_dim) + apply clarify + apply (rule_tac x="s" in exI) + apply (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull]) + done +next + show "?rhs \<subseteq> ?lhs" + using hull_mono by blast +qed + +lemma convex_hull_caratheodory: + fixes p :: "('a::euclidean_space) set" + shows "convex hull p = + {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and> + (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}" + (is "?lhs = ?rhs") +proof (intro set_eqI iffI) + fix x + assume "x \<in> ?lhs" then show "x \<in> ?rhs" + apply (simp only: convex_hull_caratheodory_aff_dim Set.mem_Collect_eq) + apply (erule ex_forward)+ + using aff_dim_le_DIM [of p] + apply simp + done +next + fix x + assume "x \<in> ?rhs" then show "x \<in> ?lhs" + by (auto simp add: convex_hull_explicit) +qed + +theorem caratheodory: + "convex hull p = + {x::'a::euclidean_space. \<exists>s. finite s \<and> s \<subseteq> p \<and> + card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s}" +proof safe + fix x + assume "x \<in> convex hull p" + then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1" + "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" + unfolding convex_hull_caratheodory by auto + then show "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s" + apply (rule_tac x=s in exI) + using hull_subset[of s convex] + using convex_convex_hull[simplified convex_explicit, of s, + THEN spec[where x=s], THEN spec[where x=u]] + apply auto + done +next + fix x s + assume "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1" "x \<in> convex hull s" + then show "x \<in> convex hull p" + using hull_mono[OF \<open>s\<subseteq>p\<close>] by auto +qed + + +subsection \<open>Relative interior of a set\<close> + +definition "rel_interior S = + {x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}" + +lemma rel_interior: + "rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}" + unfolding rel_interior_def[of S] openin_open[of "affine hull S"] + apply auto +proof - + fix x T + assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S" + then have **: "x \<in> T \<inter> affine hull S" + using hull_inc by auto + show "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S" + apply (rule_tac x = "T \<inter> (affine hull S)" in exI) + using * ** + apply auto + done +qed + +lemma mem_rel_interior: "x \<in> rel_interior S \<longleftrightarrow> (\<exists>T. open T \<and> x \<in> T \<inter> S \<and> T \<inter> affine hull S \<subseteq> S)" + by (auto simp add: rel_interior) + +lemma mem_rel_interior_ball: + "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S)" + apply (simp add: rel_interior, safe) + apply (force simp add: open_contains_ball) + apply (rule_tac x = "ball x e" in exI) + apply simp + done + +lemma rel_interior_ball: + "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}" + using mem_rel_interior_ball [of _ S] by auto + +lemma mem_rel_interior_cball: + "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S)" + apply (simp add: rel_interior, safe) + apply (force simp add: open_contains_cball) + apply (rule_tac x = "ball x e" in exI) + apply (simp add: subset_trans [OF ball_subset_cball]) + apply auto + done + +lemma rel_interior_cball: + "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}" + using mem_rel_interior_cball [of _ S] by auto + +lemma rel_interior_empty [simp]: "rel_interior {} = {}" + by (auto simp add: rel_interior_def) + +lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}" + by (metis affine_hull_eq affine_sing) + +lemma rel_interior_sing [simp]: + fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}" + apply (auto simp: rel_interior_ball) + apply (rule_tac x=1 in exI) + apply force + done + +lemma subset_rel_interior: + fixes S T :: "'n::euclidean_space set" + assumes "S \<subseteq> T" + and "affine hull S = affine hull T" + shows "rel_interior S \<subseteq> rel_interior T" + using assms by (auto simp add: rel_interior_def) + +lemma rel_interior_subset: "rel_interior S \<subseteq> S" + by (auto simp add: rel_interior_def) + +lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S" + using rel_interior_subset by (auto simp add: closure_def) + +lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S" + by (auto simp add: rel_interior interior_def) + +lemma interior_rel_interior: + fixes S :: "'n::euclidean_space set" + assumes "aff_dim S = int(DIM('n))" + shows "rel_interior S = interior S" +proof - + have "affine hull S = UNIV" + using assms affine_hull_UNIV[of S] by auto + then show ?thesis + unfolding rel_interior interior_def by auto +qed + +lemma rel_interior_interior: + fixes S :: "'n::euclidean_space set" + assumes "affine hull S = UNIV" + shows "rel_interior S = interior S" + using assms unfolding rel_interior interior_def by auto + +lemma rel_interior_open: + fixes S :: "'n::euclidean_space set" + assumes "open S" + shows "rel_interior S = S" + by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) + +lemma interior_ball [simp]: "interior (ball x e) = ball x e" + by (simp add: interior_open) + +lemma interior_rel_interior_gen: + fixes S :: "'n::euclidean_space set" + shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})" + by (metis interior_rel_interior low_dim_interior) + +lemma rel_interior_nonempty_interior: + fixes S :: "'n::euclidean_space set" + shows "interior S \<noteq> {} \<Longrightarrow> rel_interior S = interior S" +by (metis interior_rel_interior_gen) + +lemma affine_hull_nonempty_interior: + fixes S :: "'n::euclidean_space set" + shows "interior S \<noteq> {} \<Longrightarrow> affine hull S = UNIV" +by (metis affine_hull_UNIV interior_rel_interior_gen) + +lemma rel_interior_affine_hull [simp]: + fixes S :: "'n::euclidean_space set" + shows "rel_interior (affine hull S) = affine hull S" +proof - + have *: "rel_interior (affine hull S) \<subseteq> affine hull S" + using rel_interior_subset by auto + { + fix x + assume x: "x \<in> affine hull S" + define e :: real where "e = 1" + then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S" + using hull_hull[of _ S] by auto + then have "x \<in> rel_interior (affine hull S)" + using x rel_interior_ball[of "affine hull S"] by auto + } + then show ?thesis using * by auto +qed + +lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV" + by (metis open_UNIV rel_interior_open) + +lemma rel_interior_convex_shrink: + fixes S :: "'a::euclidean_space set" + assumes "convex S" + and "c \<in> rel_interior S" + and "x \<in> S" + and "0 < e" + and "e \<le> 1" + shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" +proof - + obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" + using assms(2) unfolding mem_rel_interior_ball by auto + { + fix y + assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S" + have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" + using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) + have "x \<in> affine hull S" + using assms hull_subset[of S] by auto + moreover have "1 / e + - ((1 - e) / e) = 1" + using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto + ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S" + using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] + by (simp add: algebra_simps) + have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \<bar>1/e\<bar> * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" + unfolding dist_norm norm_scaleR[symmetric] + apply (rule arg_cong[where f=norm]) + using \<open>e > 0\<close> + apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) + done + also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" + by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) + also have "\<dots> < d" + using as[unfolded dist_norm] and \<open>e > 0\<close> + by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) + finally have "y \<in> S" + apply (subst *) + apply (rule assms(1)[unfolded convex_alt,rule_format]) + apply (rule d[unfolded subset_eq,rule_format]) + unfolding mem_ball + using assms(3-5) ** + apply auto + done + } + then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S" + by auto + moreover have "e * d > 0" + using \<open>e > 0\<close> \<open>d > 0\<close> by simp + moreover have c: "c \<in> S" + using assms rel_interior_subset by auto + moreover from c have "x - e *\<^sub>R (x - c) \<in> S" + using convexD_alt[of S x c e] + apply (simp add: algebra_simps) + using assms + apply auto + done + ultimately show ?thesis + using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto +qed + +lemma interior_real_semiline: + fixes a :: real + shows "interior {a..} = {a<..}" +proof - + { + fix y + assume "a < y" + then have "y \<in> interior {a..}" + apply (simp add: mem_interior) + apply (rule_tac x="(y-a)" in exI) + apply (auto simp add: dist_norm) + done + } + moreover + { + fix y + assume "y \<in> interior {a..}" + then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}" + using mem_interior_cball[of y "{a..}"] by auto + moreover from e have "y - e \<in> cball y e" + by (auto simp add: cball_def dist_norm) + ultimately have "a \<le> y - e" by blast + then have "a < y" using e by auto + } + ultimately show ?thesis by auto +qed + +lemma continuous_ge_on_Ioo: + assumes "continuous_on {c..d} g" "\<And>x. x \<in> {c<..<d} \<Longrightarrow> g x \<ge> a" "c < d" "x \<in> {c..d}" + shows "g (x::real) \<ge> (a::real)" +proof- + from assms(3) have "{c..d} = closure {c<..<d}" by (rule closure_greaterThanLessThan[symmetric]) + also from assms(2) have "{c<..<d} \<subseteq> (g -` {a..} \<inter> {c..d})" by auto + hence "closure {c<..<d} \<subseteq> closure (g -` {a..} \<inter> {c..d})" by (rule closure_mono) + also from assms(1) have "closed (g -` {a..} \<inter> {c..d})" + by (auto simp: continuous_on_closed_vimage) + hence "closure (g -` {a..} \<inter> {c..d}) = g -` {a..} \<inter> {c..d}" by simp + finally show ?thesis using \<open>x \<in> {c..d}\<close> by auto +qed + +lemma interior_real_semiline': + fixes a :: real + shows "interior {..a} = {..<a}" +proof - + { + fix y + assume "a > y" + then have "y \<in> interior {..a}" + apply (simp add: mem_interior) + apply (rule_tac x="(a-y)" in exI) + apply (auto simp add: dist_norm) + done + } + moreover + { + fix y + assume "y \<in> interior {..a}" + then obtain e where e: "e > 0" "cball y e \<subseteq> {..a}" + using mem_interior_cball[of y "{..a}"] by auto + moreover from e have "y + e \<in> cball y e" + by (auto simp add: cball_def dist_norm) + ultimately have "a \<ge> y + e" by auto + then have "a > y" using e by auto + } + ultimately show ?thesis by auto +qed + +lemma interior_atLeastAtMost_real: "interior {a..b} = {a<..<b :: real}" +proof- + have "{a..b} = {a..} \<inter> {..b}" by auto + also have "interior ... = {a<..} \<inter> {..<b}" + by (simp add: interior_real_semiline interior_real_semiline') + also have "... = {a<..<b}" by auto + finally show ?thesis . +qed + +lemma frontier_real_Iic: + fixes a :: real + shows "frontier {..a} = {a}" + unfolding frontier_def by (auto simp add: interior_real_semiline') + +lemma rel_interior_real_box: + fixes a b :: real + assumes "a < b" + shows "rel_interior {a .. b} = {a <..< b}" +proof - + have "box a b \<noteq> {}" + using assms + unfolding set_eq_iff + by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) + then show ?thesis + using interior_rel_interior_gen[of "cbox a b", symmetric] + by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) +qed + +lemma rel_interior_real_semiline: + fixes a :: real + shows "rel_interior {a..} = {a<..}" +proof - + have *: "{a<..} \<noteq> {}" + unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) + then show ?thesis using interior_real_semiline interior_rel_interior_gen[of "{a..}"] + by (auto split: if_split_asm) +qed + +subsubsection \<open>Relative open sets\<close> + +definition "rel_open S \<longleftrightarrow> rel_interior S = S" + +lemma rel_open: "rel_open S \<longleftrightarrow> openin (subtopology euclidean (affine hull S)) S" + unfolding rel_open_def rel_interior_def + apply auto + using openin_subopen[of "subtopology euclidean (affine hull S)" S] + apply auto + done + +lemma openin_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)" + apply (simp add: rel_interior_def) + apply (subst openin_subopen) + apply blast + done + +lemma openin_set_rel_interior: + "openin (subtopology euclidean S) (rel_interior S)" +by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) + +lemma affine_rel_open: + fixes S :: "'n::euclidean_space set" + assumes "affine S" + shows "rel_open S" + unfolding rel_open_def + using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] + by metis + +lemma affine_closed: + fixes S :: "'n::euclidean_space set" + assumes "affine S" + shows "closed S" +proof - + { + assume "S \<noteq> {}" + then obtain L where L: "subspace L" "affine_parallel S L" + using assms affine_parallel_subspace[of S] by auto + then obtain a where a: "S = (op + a ` L)" + using affine_parallel_def[of L S] affine_parallel_commut by auto + from L have "closed L" using closed_subspace by auto + then have "closed S" + using closed_translation a by auto + } + then show ?thesis by auto +qed + +lemma closure_affine_hull: + fixes S :: "'n::euclidean_space set" + shows "closure S \<subseteq> affine hull S" + by (intro closure_minimal hull_subset affine_closed affine_affine_hull) + +lemma closure_same_affine_hull [simp]: + fixes S :: "'n::euclidean_space set" + shows "affine hull (closure S) = affine hull S" +proof - + have "affine hull (closure S) \<subseteq> affine hull S" + using hull_mono[of "closure S" "affine hull S" "affine"] + closure_affine_hull[of S] hull_hull[of "affine" S] + by auto + moreover have "affine hull (closure S) \<supseteq> affine hull S" + using hull_mono[of "S" "closure S" "affine"] closure_subset by auto + ultimately show ?thesis by auto +qed + +lemma closure_aff_dim [simp]: + fixes S :: "'n::euclidean_space set" + shows "aff_dim (closure S) = aff_dim S" +proof - + have "aff_dim S \<le> aff_dim (closure S)" + using aff_dim_subset closure_subset by auto + moreover have "aff_dim (closure S) \<le> aff_dim (affine hull S)" + using aff_dim_subset closure_affine_hull by blast + moreover have "aff_dim (affine hull S) = aff_dim S" + using aff_dim_affine_hull by auto + ultimately show ?thesis by auto +qed + +lemma rel_interior_closure_convex_shrink: + fixes S :: "_::euclidean_space set" + assumes "convex S" + and "c \<in> rel_interior S" + and "x \<in> closure S" + and "e > 0" + and "e \<le> 1" + shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" +proof - + obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" + using assms(2) unfolding mem_rel_interior_ball by auto + have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d" + proof (cases "x \<in> S") + case True + then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close> + apply (rule_tac bexI[where x=x]) + apply (auto) + done + 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 \<in> S" "y \<noteq> x" "dist y x < 1" + using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto + then show ?thesis + apply (rule_tac x=y in bexI) + unfolding True + using \<open>d > 0\<close> + apply auto + done + next + case False + then have "0 < e * d / (1 - e)" and *: "1 - e > 0" + using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by (auto) + then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)" + using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto + then show ?thesis + apply (rule_tac x=y in bexI) + unfolding dist_norm + using pos_less_divide_eq[OF *] + apply auto + done + qed + qed + then obtain y where "y \<in> 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 \<open>e > 0\<close> + by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) + have zball: "z \<in> ball c d" + using mem_ball z_def dist_norm[of c] + using y and assms(4,5) + by (auto simp add:field_simps norm_minus_commute) + have "x \<in> affine hull S" + using closure_affine_hull assms by auto + moreover have "y \<in> affine hull S" + using \<open>y \<in> S\<close> hull_subset[of S] by auto + moreover have "c \<in> affine hull S" + using assms rel_interior_subset hull_subset[of S] by auto + ultimately have "z \<in> affine hull S" + using z_def affine_affine_hull[of S] + mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] + assms + by (auto simp add: field_simps) + then have "z \<in> S" using d zball by auto + obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d" + using zball open_ball[of c d] openE[of "ball c d" z] by auto + then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S" + by auto + then have "ball z d1 \<inter> affine hull S \<subseteq> S" + using d by auto + then have "z \<in> rel_interior S" + using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto + then have "y - e *\<^sub>R (y - z) \<in> rel_interior S" + using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto + then show ?thesis using * by auto +qed + +lemma rel_interior_eq: + "rel_interior s = s \<longleftrightarrow> openin(subtopology euclidean (affine hull s)) s" +using rel_open rel_open_def by blast + +lemma rel_interior_openin: + "openin(subtopology euclidean (affine hull s)) s \<Longrightarrow> rel_interior s = s" +by (simp add: rel_interior_eq) + +lemma rel_interior_affine: + fixes S :: "'n::euclidean_space set" + shows "affine S \<Longrightarrow> rel_interior S = S" +using affine_rel_open rel_open_def by auto + +lemma rel_interior_eq_closure: + fixes S :: "'n::euclidean_space set" + shows "rel_interior S = closure S \<longleftrightarrow> affine S" +proof (cases "S = {}") + case True + then show ?thesis + by auto +next + case False show ?thesis + proof + assume eq: "rel_interior S = closure S" + have "S = {} \<or> S = affine hull S" + apply (rule connected_clopen [THEN iffD1, rule_format]) + apply (simp add: affine_imp_convex convex_connected) + apply (rule conjI) + apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) + apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) + done + with False have "affine hull S = S" + by auto + then show "affine S" + by (metis affine_hull_eq) + next + assume "affine S" + then show "rel_interior S = closure S" + by (simp add: rel_interior_affine affine_closed) + qed +qed + + +subsubsection\<open>Relative interior preserves under linear transformations\<close> + +lemma rel_interior_translation_aux: + fixes a :: "'n::euclidean_space" + shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)" +proof - + { + fix x + assume x: "x \<in> rel_interior S" + then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S" + using mem_rel_interior[of x S] by auto + then have "open ((\<lambda>x. a + x) ` T)" + and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)" + and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S" + using affine_hull_translation[of a S] open_translation[of T a] x by auto + then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)" + using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto + } + then show ?thesis by auto +qed + +lemma rel_interior_translation: + fixes a :: "'n::euclidean_space" + shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S" +proof - + have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S" + using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"] + translation_assoc[of "-a" "a"] + by auto + then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)" + using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"] + by auto + then show ?thesis + using rel_interior_translation_aux[of a S] by auto +qed + + +lemma affine_hull_linear_image: + assumes "bounded_linear f" + shows "f ` (affine hull s) = affine hull f ` s" + apply rule + unfolding subset_eq ball_simps + apply (rule_tac[!] hull_induct, rule hull_inc) + prefer 3 + apply (erule imageE) + apply (rule_tac x=xa in image_eqI) + apply assumption + apply (rule hull_subset[unfolded subset_eq, rule_format]) + apply assumption +proof - + interpret f: bounded_linear f by fact + show "affine {x. f x \<in> affine hull f ` s}" + unfolding affine_def + by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) + show "affine {x. x \<in> f ` (affine hull s)}" + using affine_affine_hull[unfolded affine_def, of s] + unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric]) +qed auto + + +lemma rel_interior_injective_on_span_linear_image: + fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" + and S :: "'m::euclidean_space set" + assumes "bounded_linear f" + and "inj_on f (span S)" + shows "rel_interior (f ` S) = f ` (rel_interior S)" +proof - + { + fix z + assume z: "z \<in> rel_interior (f ` S)" + then have "z \<in> f ` S" + using rel_interior_subset[of "f ` S"] by auto + then obtain x where x: "x \<in> S" "f x = z" by auto + obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)" + using z rel_interior_cball[of "f ` S"] by auto + obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K" + using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto + define e1 where "e1 = 1 / K" + then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x" + using K pos_le_divide_eq[of e1] by auto + define e where "e = e1 * e2" + then have "e > 0" using e1 e2 by auto + { + fix y + assume y: "y \<in> cball x e \<inter> affine hull S" + then have h1: "f y \<in> affine hull (f ` S)" + using affine_hull_linear_image[of f S] assms by auto + from y have "norm (x-y) \<le> e1 * e2" + using cball_def[of x e] dist_norm[of x y] e_def by auto + moreover have "f x - f y = f (x - y)" + using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto + moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)" + using e1 by auto + ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2" + by auto + then have "f y \<in> cball z e2" + using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto + then have "f y \<in> f ` S" + using y e2 h1 by auto + then have "y \<in> S" + using assms y hull_subset[of S] affine_hull_subset_span + inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>] + by (metis Int_iff span_inc subsetCE) + } + then have "z \<in> f ` (rel_interior S)" + using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto + } + moreover + { + fix x + assume x: "x \<in> rel_interior S" + then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S" + using rel_interior_cball[of S] by auto + have "x \<in> S" using x rel_interior_subset by auto + then have *: "f x \<in> f ` S" by auto + have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0" + using assms subspace_span linear_conv_bounded_linear[of f] + linear_injective_on_subspace_0[of f "span S"] + by auto + then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)" + using assms injective_imp_isometric[of "span S" f] + subspace_span[of S] closed_subspace[of "span S"] + by auto + define e where "e = e1 * e2" + hence "e > 0" using e1 e2 by auto + { + fix y + assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)" + then have "y \<in> f ` (affine hull S)" + using affine_hull_linear_image[of f S] assms by auto + then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto + with y have "norm (f x - f xy) \<le> e1 * e2" + using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto + moreover have "f x - f xy = f (x - xy)" + using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto + moreover have *: "x - xy \<in> span S" + using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy + affine_hull_subset_span[of S] span_inc + by auto + moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))" + using e1 by auto + ultimately have "e1 * norm (x - xy) \<le> e1 * e2" + by auto + then have "xy \<in> cball x e2" + using cball_def[of x e2] dist_norm[of x xy] e1 by auto + then have "y \<in> f ` S" + using xy e2 by auto + } + then have "f x \<in> rel_interior (f ` S)" + using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto + } + ultimately show ?thesis by auto +qed + +lemma rel_interior_injective_linear_image: + fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" + assumes "bounded_linear f" + and "inj f" + shows "rel_interior (f ` S) = f ` (rel_interior S)" + using assms rel_interior_injective_on_span_linear_image[of f S] + subset_inj_on[of f "UNIV" "span S"] + by auto + + +subsection\<open>Some Properties of subset of standard basis\<close> + +lemma affine_hull_substd_basis: + assumes "d \<subseteq> Basis" + shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}" + (is "affine hull (insert 0 ?A) = ?B") +proof - + have *: "\<And>A. op + (0::'a) ` A = A" "\<And>A. op + (- (0::'a)) ` A = A" + by auto + show ?thesis + unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. +qed + +lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S" + by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset) + + +subsection \<open>Openness and compactness are preserved by convex hull operation.\<close> + +lemma open_convex_hull[intro]: + fixes s :: "'a::real_normed_vector set" + assumes "open s" + shows "open (convex hull s)" + unfolding open_contains_cball convex_hull_explicit + unfolding mem_Collect_eq ball_simps(8) +proof (rule, rule) + fix a + assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a" + then obtain t u where obt: "finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a" + by auto + + from assms[unfolded open_contains_cball] obtain b + where b: "\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s" + using bchoice[of s "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto + have "b ` t \<noteq> {}" + using obt by auto + define i where "i = b ` t" + + show "\<exists>e > 0. + cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}" + apply (rule_tac x = "Min i" in exI) + unfolding subset_eq + apply rule + defer + apply rule + unfolding mem_Collect_eq + proof - + show "0 < Min i" + unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>] + using b + apply simp + apply rule + apply (erule_tac x=x in ballE) + using \<open>t\<subseteq>s\<close> + apply auto + done + next + fix y + assume "y \<in> cball a (Min i)" + then have y: "norm (a - y) \<le> Min i" + unfolding dist_norm[symmetric] by auto + { + fix x + assume "x \<in> t" + then have "Min i \<le> b x" + unfolding i_def + apply (rule_tac Min_le) + using obt(1) + apply auto + done + then have "x + (y - a) \<in> cball x (b x)" + using y unfolding mem_cball dist_norm by auto + moreover from \<open>x\<in>t\<close> have "x \<in> s" + using obt(2) by auto + ultimately have "x + (y - a) \<in> s" + using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast + } + moreover + have *: "inj_on (\<lambda>v. v + (y - a)) t" + unfolding inj_on_def by auto + have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1" + unfolding setsum.reindex[OF *] o_def using obt(4) by auto + moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y" + unfolding setsum.reindex[OF *] o_def using obt(4,5) + by (simp add: setsum.distrib setsum_subtractf scaleR_left.setsum[symmetric] scaleR_right_distrib) + ultimately + show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y" + apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI) + apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI) + using obt(1, 3) + apply auto + done + qed +qed + +lemma compact_convex_combinations: + fixes s t :: "'a::real_normed_vector set" + assumes "compact s" "compact t" + shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t}" +proof - + let ?X = "{0..1} \<times> s \<times> t" + let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" + have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t} = ?h ` ?X" + apply (rule set_eqI) + unfolding image_iff mem_Collect_eq + apply rule + apply auto + apply (rule_tac x=u in rev_bexI) + apply simp + apply (erule rev_bexI) + apply (erule rev_bexI) + apply simp + apply auto + done + have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" + unfolding continuous_on by (rule ballI) (intro tendsto_intros) + then show ?thesis + unfolding * + apply (rule compact_continuous_image) + apply (intro compact_Times compact_Icc assms) + done +qed + +lemma finite_imp_compact_convex_hull: + fixes s :: "'a::real_normed_vector set" + assumes "finite s" + shows "compact (convex hull s)" +proof (cases "s = {}") + case True + then show ?thesis by simp +next + case False + with assms show ?thesis + proof (induct rule: finite_ne_induct) + case (singleton x) + show ?case by simp + next + case (insert x A) + let ?f = "\<lambda>(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y" + let ?T = "{0..1::real} \<times> (convex hull A)" + have "continuous_on ?T ?f" + unfolding split_def continuous_on by (intro ballI tendsto_intros) + moreover have "compact ?T" + by (intro compact_Times compact_Icc insert) + ultimately have "compact (?f ` ?T)" + by (rule compact_continuous_image) + also have "?f ` ?T = convex hull (insert x A)" + unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>] + apply safe + apply (rule_tac x=a in exI, simp) + apply (rule_tac x="1 - a" in exI, simp) + apply fast + apply (rule_tac x="(u, b)" in image_eqI, simp_all) + done + finally show "compact (convex hull (insert x A))" . + qed +qed + +lemma compact_convex_hull: + fixes s :: "'a::euclidean_space set" + assumes "compact s" + shows "compact (convex hull s)" +proof (cases "s = {}") + case True + then show ?thesis using compact_empty by simp +next + case False + then obtain w where "w \<in> s" by auto + show ?thesis + unfolding caratheodory[of s] + proof (induct ("DIM('a) + 1")) + case 0 + have *: "{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}" + using compact_empty by auto + from 0 show ?case unfolding * by simp + next + case (Suc n) + show ?case + proof (cases "n = 0") + case True + have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s" + unfolding set_eq_iff and mem_Collect_eq + proof (rule, rule) + fix x + assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" + then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" + by auto + show "x \<in> s" + proof (cases "card t = 0") + case True + then show ?thesis + using t(4) unfolding card_0_eq[OF t(1)] by simp + next + case False + then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto + then obtain a where "t = {a}" unfolding card_Suc_eq by auto + then show ?thesis using t(2,4) by simp + qed + next + fix x assume "x\<in>s" + then show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" + apply (rule_tac x="{x}" in exI) + unfolding convex_hull_singleton + apply auto + done + qed + then show ?thesis using assms by simp + next + case False + have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = + {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. + 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> {x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> x \<in> convex hull t}}" + unfolding set_eq_iff and mem_Collect_eq + proof (rule, rule) + fix x + assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> + 0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)" + then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" + "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n" "v \<in> convex hull t" + by auto + moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t" + apply (rule convexD_alt) + using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex] + using obt(7) and hull_mono[of t "insert u t"] + apply auto + done + ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" + apply (rule_tac x="insert u t" in exI) + apply (auto simp add: card_insert_if) + done + next + fix x + assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" + then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" + by auto + show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> + 0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)" + proof (cases "card t = Suc n") + case False + then have "card t \<le> n" using t(3) by auto + then show ?thesis + apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) + using \<open>w\<in>s\<close> and t + apply (auto intro!: exI[where x=t]) + done + next + case True + then obtain a u where au: "t = insert a u" "a\<notin>u" + apply (drule_tac card_eq_SucD) + apply auto + done + show ?thesis + proof (cases "u = {}") + case True + then have "x = a" using t(4)[unfolded au] by auto + show ?thesis unfolding \<open>x = a\<close> + apply (rule_tac x=a in exI) + apply (rule_tac x=a in exI) + apply (rule_tac x=1 in exI) + using t and \<open>n \<noteq> 0\<close> + unfolding au + apply (auto intro!: exI[where x="{a}"]) + done + next + case False + obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1" + "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" + using t(4)[unfolded au convex_hull_insert[OF False]] + by auto + have *: "1 - vx = ux" using obt(3) by auto + show ?thesis + apply (rule_tac x=a in exI) + apply (rule_tac x=b in exI) + apply (rule_tac x=vx in exI) + using obt and t(1-3) + unfolding au and * using card_insert_disjoint[OF _ au(2)] + apply (auto intro!: exI[where x=u]) + done + qed + qed + qed + then show ?thesis + using compact_convex_combinations[OF assms Suc] by simp + qed + qed +qed + + +subsection \<open>Extremal points of a simplex are some vertices.\<close> + +lemma dist_increases_online: + fixes a b d :: "'a::real_inner" + assumes "d \<noteq> 0" + shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b" +proof (cases "inner a d - inner b d > 0") + case True + then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" + apply (rule_tac add_pos_pos) + using assms + apply auto + done + then show ?thesis + apply (rule_tac disjI2) + unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff + apply (simp add: algebra_simps inner_commute) + done +next + case False + then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" + apply (rule_tac add_pos_nonneg) + using assms + apply auto + done + then show ?thesis + apply (rule_tac disjI1) + unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff + apply (simp add: algebra_simps inner_commute) + done +qed + +lemma norm_increases_online: + fixes d :: "'a::real_inner" + shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a" + using dist_increases_online[of d a 0] unfolding dist_norm by auto + +lemma simplex_furthest_lt: + fixes s :: "'a::real_inner set" + assumes "finite s" + shows "\<forall>x \<in> convex hull s. x \<notin> s \<longrightarrow> (\<exists>y \<in> convex hull s. norm (x - a) < norm(y - a))" + using assms +proof induct + fix x s + assume as: "finite s" "x\<notin>s" "\<forall>x\<in>convex hull s. x \<notin> s \<longrightarrow> (\<exists>y\<in>convex hull s. norm (x - a) < norm (y - a))" + show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow> + (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))" + proof (rule, rule, cases "s = {}") + case False + fix y + assume y: "y \<in> convex hull insert x s" "y \<notin> insert x s" + obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "y = u *\<^sub>R x + v *\<^sub>R b" + using y(1)[unfolded convex_hull_insert[OF False]] by auto + show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)" + proof (cases "y \<in> convex hull s") + case True + then obtain z where "z \<in> convex hull s" "norm (y - a) < norm (z - a)" + using as(3)[THEN bspec[where x=y]] and y(2) by auto + then show ?thesis + apply (rule_tac x=z in bexI) + unfolding convex_hull_insert[OF False] + apply auto + done + next + case False + show ?thesis + using obt(3) + proof (cases "u = 0", case_tac[!] "v = 0") + assume "u = 0" "v \<noteq> 0" + then have "y = b" using obt by auto + then show ?thesis using False and obt(4) by auto + next + assume "u \<noteq> 0" "v = 0" + then have "y = x" using obt by auto + then show ?thesis using y(2) by auto + next + assume "u \<noteq> 0" "v \<noteq> 0" + then obtain w where w: "w>0" "w<u" "w<v" + using real_lbound_gt_zero[of u v] and obt(1,2) by auto + have "x \<noteq> b" + proof + assume "x = b" + then have "y = b" unfolding obt(5) + using obt(3) by (auto simp add: scaleR_left_distrib[symmetric]) + then show False using obt(4) and False by simp + qed + then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto + show ?thesis + using dist_increases_online[OF *, of a y] + proof (elim disjE) + assume "dist a y < dist a (y + w *\<^sub>R (x - b))" + then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" + unfolding dist_commute[of a] + unfolding dist_norm obt(5) + by (simp add: algebra_simps) + moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s" + unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq + apply (rule_tac x="u + w" in exI) + apply rule + defer + apply (rule_tac x="v - w" in exI) + using \<open>u \<ge> 0\<close> and w and obt(3,4) + apply auto + done + ultimately show ?thesis by auto + next + assume "dist a y < dist a (y - w *\<^sub>R (x - b))" + then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" + unfolding dist_commute[of a] + unfolding dist_norm obt(5) + by (simp add: algebra_simps) + moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s" + unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq + apply (rule_tac x="u - w" in exI) + apply rule + defer + apply (rule_tac x="v + w" in exI) + using \<open>u \<ge> 0\<close> and w and obt(3,4) + apply auto + done + ultimately show ?thesis by auto + qed + qed auto + qed + qed auto +qed (auto simp add: assms) + +lemma simplex_furthest_le: + fixes s :: "'a::real_inner set" + assumes "finite s" + and "s \<noteq> {}" + shows "\<exists>y\<in>s. \<forall>x\<in> convex hull s. norm (x - a) \<le> norm (y - a)" +proof - + have "convex hull s \<noteq> {}" + using hull_subset[of s convex] and assms(2) by auto + then obtain x where x: "x \<in> convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)" + using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a] + unfolding dist_commute[of a] + unfolding dist_norm + by auto + show ?thesis + proof (cases "x \<in> s") + case False + then obtain y where "y \<in> convex hull s" "norm (x - a) < norm (y - a)" + using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) + by auto + then show ?thesis + using x(2)[THEN bspec[where x=y]] by auto + next + case True + with x show ?thesis by auto + qed +qed + +lemma simplex_furthest_le_exists: + fixes s :: "('a::real_inner) set" + shows "finite s \<Longrightarrow> \<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm (x - a) \<le> norm (y - a)" + using simplex_furthest_le[of s] by (cases "s = {}") auto + +lemma simplex_extremal_le: + fixes s :: "'a::real_inner set" + assumes "finite s" + and "s \<noteq> {}" + shows "\<exists>u\<in>s. \<exists>v\<in>s. \<forall>x\<in>convex hull s. \<forall>y \<in> convex hull s. norm (x - y) \<le> norm (u - v)" +proof - + have "convex hull s \<noteq> {}" + using hull_subset[of s convex] and assms(2) by auto + then obtain u v where obt: "u \<in> convex hull s" "v \<in> convex hull s" + "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)" + using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] + by (auto simp: dist_norm) + then show ?thesis + proof (cases "u\<notin>s \<or> v\<notin>s", elim disjE) + assume "u \<notin> s" + then obtain y where "y \<in> convex hull s" "norm (u - v) < norm (y - v)" + using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) + by auto + then show ?thesis + using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) + by auto + next + assume "v \<notin> s" + then obtain y where "y \<in> convex hull s" "norm (v - u) < norm (y - u)" + using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) + by auto + then show ?thesis + using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) + by (auto simp add: norm_minus_commute) + qed auto +qed + +lemma simplex_extremal_le_exists: + fixes s :: "'a::real_inner set" + shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s \<Longrightarrow> + \<exists>u\<in>s. \<exists>v\<in>s. norm (x - y) \<le> norm (u - v)" + using convex_hull_empty simplex_extremal_le[of s] + by(cases "s = {}") auto + + +subsection \<open>Closest point of a convex set is unique, with a continuous projection.\<close> + +definition closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a" + where "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))" + +lemma closest_point_exists: + assumes "closed s" + and "s \<noteq> {}" + shows "closest_point s a \<in> s" + and "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y" + unfolding closest_point_def + apply(rule_tac[!] someI2_ex) + apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) + done + +lemma closest_point_in_set: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s a \<in> s" + by (meson closest_point_exists) + +lemma closest_point_le: "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x" + using closest_point_exists[of s] by auto + +lemma closest_point_self: + assumes "x \<in> s" + shows "closest_point s x = x" + unfolding closest_point_def + apply (rule some1_equality, rule ex1I[of _ x]) + using assms + apply auto + done + +lemma closest_point_refl: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s x = x \<longleftrightarrow> x \<in> s" + using closest_point_in_set[of s x] closest_point_self[of x s] + by auto + +lemma closer_points_lemma: + assumes "inner y z > 0" + shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y" +proof - + have z: "inner z z > 0" + unfolding inner_gt_zero_iff using assms by auto + then show ?thesis + using assms + apply (rule_tac x = "inner y z / inner z z" in exI) + apply rule + defer + proof rule+ + fix v + assume "0 < v" and "v \<le> inner y z / inner z z" + then show "norm (v *\<^sub>R z - y) < norm y" + unfolding norm_lt using z and assms + by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>]) + qed auto +qed + +lemma closer_point_lemma: + assumes "inner (y - x) (z - x) > 0" + shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y" +proof - + obtain u where "u > 0" + and u: "\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" + using closer_points_lemma[OF assms] by auto + show ?thesis + apply (rule_tac x="min u 1" in exI) + using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close> + unfolding dist_norm by (auto simp add: norm_minus_commute field_simps) +qed + +lemma any_closest_point_dot: + assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" + shows "inner (a - x) (y - x) \<le> 0" +proof (rule ccontr) + assume "\<not> ?thesis" + then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" + using closer_point_lemma[of a x y] by auto + let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" + have "?z \<in> s" + using convexD_alt[OF assms(1,3,4), of u] using u by auto + then show False + using assms(5)[THEN bspec[where x="?z"]] and u(3) + by (auto simp add: dist_commute algebra_simps) +qed + +lemma any_closest_point_unique: + fixes x :: "'a::real_inner" + assumes "convex s" "closed s" "x \<in> s" "y \<in> s" + "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z" + shows "x = y" + using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] + unfolding norm_pths(1) and norm_le_square + by (auto simp add: algebra_simps) + +lemma closest_point_unique: + assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" + shows "x = closest_point s a" + using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] + using closest_point_exists[OF assms(2)] and assms(3) by auto + +lemma closest_point_dot: + assumes "convex s" "closed s" "x \<in> s" + shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0" + apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) + using closest_point_exists[OF assms(2)] and assms(3) + apply auto + done + +lemma closest_point_lt: + assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a" + shows "dist a (closest_point s a) < dist a x" + apply (rule ccontr) + apply (rule_tac notE[OF assms(4)]) + apply (rule closest_point_unique[OF assms(1-3), of a]) + using closest_point_le[OF assms(2), of _ a] + apply fastforce + done + +lemma closest_point_lipschitz: + assumes "convex s" + and "closed s" "s \<noteq> {}" + shows "dist (closest_point s x) (closest_point s y) \<le> dist x y" +proof - + have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0" + and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0" + apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) + using closest_point_exists[OF assms(2-3)] + apply auto + done + then show ?thesis unfolding dist_norm and norm_le + using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"] + by (simp add: inner_add inner_diff inner_commute) +qed + +lemma continuous_at_closest_point: + assumes "convex s" + and "closed s" + and "s \<noteq> {}" + shows "continuous (at x) (closest_point s)" + unfolding continuous_at_eps_delta + using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto + +lemma continuous_on_closest_point: + assumes "convex s" + and "closed s" + and "s \<noteq> {}" + shows "continuous_on t (closest_point s)" + by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) + + +subsubsection \<open>Various point-to-set separating/supporting hyperplane theorems.\<close> + +lemma supporting_hyperplane_closed_point: + fixes z :: "'a::{real_inner,heine_borel}" + assumes "convex s" + and "closed s" + and "s \<noteq> {}" + and "z \<notin> s" + shows "\<exists>a b. \<exists>y\<in>s. inner a z < b \<and> inner a y = b \<and> (\<forall>x\<in>s. inner a x \<ge> b)" +proof - + obtain y where "y \<in> s" and y: "\<forall>x\<in>s. dist z y \<le> dist z x" + by (metis distance_attains_inf[OF assms(2-3)]) + show ?thesis + apply (rule_tac x="y - z" in exI) + apply (rule_tac x="inner (y - z) y" in exI) + apply (rule_tac x=y in bexI) + apply rule + defer + apply rule + defer + apply rule + apply (rule ccontr) + using \<open>y \<in> s\<close> + proof - + show "inner (y - z) z < inner (y - z) y" + apply (subst diff_gt_0_iff_gt [symmetric]) + unfolding inner_diff_right[symmetric] and inner_gt_zero_iff + using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> + apply auto + done + next + fix x + assume "x \<in> s" + have *: "\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" + using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto + assume "\<not> inner (y - z) y \<le> inner (y - z) x" + then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" + using closer_point_lemma[of z y x] by (auto simp add: inner_diff) + then show False + using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps) + qed auto +qed + +lemma separating_hyperplane_closed_point: + fixes z :: "'a::{real_inner,heine_borel}" + assumes "convex s" + and "closed s" + and "z \<notin> s" + shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)" +proof (cases "s = {}") + case True + then show ?thesis + apply (rule_tac x="-z" in exI) + apply (rule_tac x=1 in exI) + using less_le_trans[OF _ inner_ge_zero[of z]] + apply auto + done +next + case False + obtain y where "y \<in> s" and y: "\<forall>x\<in>s. dist z y \<le> dist z x" + by (metis distance_attains_inf[OF assms(2) False]) + show ?thesis + apply (rule_tac x="y - z" in exI) + apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI) + apply rule + defer + apply rule + proof - + fix x + assume "x \<in> s" + have "\<not> 0 < inner (z - y) (x - y)" + apply (rule notI) + apply (drule closer_point_lemma) + proof - + assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z" + then obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" + by auto + then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]] + using assms(1)[unfolded convex_alt, THEN bspec[where x=y]] + using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps) + qed + moreover have "0 < (norm (y - z))\<^sup>2" + using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto + then have "0 < inner (y - z) (y - z)" + unfolding power2_norm_eq_inner by simp + ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x" + unfolding power2_norm_eq_inner and not_less + by (auto simp add: field_simps inner_commute inner_diff) + qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto) +qed + +lemma separating_hyperplane_closed_0: + assumes "convex (s::('a::euclidean_space) set)" + and "closed s" + and "0 \<notin> s" + shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)" +proof (cases "s = {}") + case True + have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)" + defer + apply (subst norm_le_zero_iff[symmetric]) + apply (auto simp: SOME_Basis) + done + then show ?thesis + apply (rule_tac x="SOME i. i\<in>Basis" in exI) + apply (rule_tac x=1 in exI) + using True using DIM_positive[where 'a='a] + apply auto + done +next + case False + then show ?thesis + using False using separating_hyperplane_closed_point[OF assms] + apply (elim exE) + unfolding inner_zero_right + apply (rule_tac x=a in exI) + apply (rule_tac x=b in exI) + apply auto + done +qed + + +subsubsection \<open>Now set-to-set for closed/compact sets\<close> + +lemma separating_hyperplane_closed_compact: + fixes s :: "'a::euclidean_space set" + assumes "convex s" + and "closed s" + and "convex t" + and "compact t" + and "t \<noteq> {}" + and "s \<inter> t = {}" + shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)" +proof (cases "s = {}") + case True + obtain b where b: "b > 0" "\<forall>x\<in>t. norm x \<le> b" + using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto + obtain z :: 'a where z: "norm z = b + 1" + using vector_choose_size[of "b + 1"] and b(1) by auto + then have "z \<notin> t" using b(2)[THEN bspec[where x=z]] by auto + then obtain a b where ab: "inner a z < b" "\<forall>x\<in>t. b < inner a x" + using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] + by auto + then show ?thesis + using True by auto +next + case False + then obtain y where "y \<in> s" by auto + obtain a b where "0 < b" "\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. b < inner a x" + using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] + using closed_compact_differences[OF assms(2,4)] + using assms(6) by auto blast + then have ab: "\<forall>x\<in>s. \<forall>y\<in>t. b + inner a y < inner a x" + apply - + apply rule + apply rule + apply (erule_tac x="x - y" in ballE) + apply (auto simp add: inner_diff) + done + define k where "k = (SUP x:t. a \<bullet> x)" + show ?thesis + apply (rule_tac x="-a" in exI) + apply (rule_tac x="-(k + b / 2)" in exI) + apply (intro conjI ballI) + unfolding inner_minus_left and neg_less_iff_less + proof - + fix x assume "x \<in> t" + then have "inner a x - b / 2 < k" + unfolding k_def + proof (subst less_cSUP_iff) + show "t \<noteq> {}" by fact + show "bdd_above (op \<bullet> a ` t)" + using ab[rule_format, of y] \<open>y \<in> s\<close> + by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) + qed (auto intro!: bexI[of _ x] \<open>0<b\<close>) + then show "inner a x < k + b / 2" + by auto + next + fix x + assume "x \<in> s" + then have "k \<le> inner a x - b" + unfolding k_def + apply (rule_tac cSUP_least) + using assms(5) + using ab[THEN bspec[where x=x]] + apply auto + done + then show "k + b / 2 < inner a x" + using \<open>0 < b\<close> by auto + qed +qed + +lemma separating_hyperplane_compact_closed: + fixes s :: "'a::euclidean_space set" + assumes "convex s" + and "compact s" + and "s \<noteq> {}" + and "convex t" + and "closed t" + and "s \<inter> t = {}" + shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)" +proof - + obtain a b where "(\<forall>x\<in>t. inner a x < b) \<and> (\<forall>x\<in>s. b < inner a x)" + using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) + by auto + then show ?thesis + apply (rule_tac x="-a" in exI) + apply (rule_tac x="-b" in exI) + apply auto + done +qed + + +subsubsection \<open>General case without assuming closure and getting non-strict separation\<close> + +lemma separating_hyperplane_set_0: + assumes "convex s" "(0::'a::euclidean_space) \<notin> s" + shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)" +proof - + let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}" + have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` s" "finite f" for f + proof - + obtain c where c: "f = ?k ` c" "c \<subseteq> s" "finite c" + using finite_subset_image[OF as(2,1)] by auto + then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x" + using separating_hyperplane_closed_0[OF convex_convex_hull, of c] + using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) + using subset_hull[of convex, OF assms(1), symmetric, of c] + by force + then have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" + apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) + using hull_subset[of c convex] + unfolding subset_eq and inner_scaleR + by (auto simp add: inner_commute del: ballE elim!: ballE) + then show "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" + unfolding c(1) frontier_cball sphere_def dist_norm by auto + qed + have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` s)) \<noteq> {}" + apply (rule compact_imp_fip) + apply (rule compact_frontier[OF compact_cball]) + using * closed_halfspace_ge + by auto + then obtain x where "norm x = 1" "\<forall>y\<in>s. x\<in>?k y" + unfolding frontier_cball dist_norm sphere_def by auto + then show ?thesis + by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) +qed + +lemma separating_hyperplane_sets: + fixes s t :: "'a::euclidean_space set" + assumes "convex s" + and "convex t" + and "s \<noteq> {}" + and "t \<noteq> {}" + and "s \<inter> t = {}" + shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>s. inner a x \<le> b) \<and> (\<forall>x\<in>t. inner a x \<ge> b)" +proof - + from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] + obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> t \<and> y \<in> s}. 0 \<le> inner a x" + using assms(3-5) by fastforce + then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x" + by (force simp add: inner_diff) + then have bdd: "bdd_above ((op \<bullet> a)`s)" + using \<open>t \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *]) + show ?thesis + using \<open>a\<noteq>0\<close> + by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"]) + (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *) +qed + + +subsection \<open>More convexity generalities\<close> + +lemma convex_closure [intro,simp]: + fixes s :: "'a::real_normed_vector set" + assumes "convex s" + shows "convex (closure s)" + apply (rule convexI) + apply (unfold closure_sequential, elim exE) + apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) + apply (rule,rule) + apply (rule convexD [OF assms]) + apply (auto del: tendsto_const intro!: tendsto_intros) + done + +lemma convex_interior [intro,simp]: + fixes s :: "'a::real_normed_vector set" + assumes "convex s" + shows "convex (interior s)" + unfolding convex_alt Ball_def mem_interior + apply (rule,rule,rule,rule,rule,rule) + apply (elim exE conjE) +proof - + fix x y u + assume u: "0 \<le> u" "u \<le> (1::real)" + fix e d + assume ed: "ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e" + show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" + apply (rule_tac x="min d e" in exI) + apply rule + unfolding subset_eq + defer + apply rule + proof - + fix z + assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" + then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s" + apply (rule_tac assms[unfolded convex_alt, rule_format]) + using ed(1,2) and u + unfolding subset_eq mem_ball Ball_def dist_norm + apply (auto simp add: algebra_simps) + done + then show "z \<in> s" + using u by (auto simp add: algebra_simps) + qed(insert u ed(3-4), auto) +qed + +lemma convex_hull_eq_empty[simp]: "convex hull s = {} \<longleftrightarrow> s = {}" + using hull_subset[of s convex] convex_hull_empty by auto + + +subsection \<open>Moving and scaling convex hulls.\<close> + +lemma convex_hull_set_plus: + "convex hull (s + t) = convex hull s + convex hull t" + unfolding set_plus_image + apply (subst convex_hull_linear_image [symmetric]) + apply (simp add: linear_iff scaleR_right_distrib) + apply (simp add: convex_hull_Times) + done + +lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` t = {a} + t" + unfolding set_plus_def by auto + +lemma convex_hull_translation: + "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)" + unfolding translation_eq_singleton_plus + by (simp only: convex_hull_set_plus convex_hull_singleton) + +lemma convex_hull_scaling: + "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)" + using linear_scaleR by (rule convex_hull_linear_image [symmetric]) + +lemma convex_hull_affinity: + "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)" + by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation) + + +subsection \<open>Convexity of cone hulls\<close> + +lemma convex_cone_hull: + assumes "convex S" + shows "convex (cone hull S)" +proof (rule convexI) + fix x y + assume xy: "x \<in> cone hull S" "y \<in> cone hull S" + then have "S \<noteq> {}" + using cone_hull_empty_iff[of S] by auto + fix u v :: real + assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1" + then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S" + using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto + from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" + using cone_hull_expl[of S] by auto + from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S" + using cone_hull_expl[of S] by auto + { + assume "cx + cy \<le> 0" + then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0" + using x y by auto + then have "u *\<^sub>R x + v *\<^sub>R y = 0" + by auto + then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" + using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto + } + moreover + { + assume "cx + cy > 0" + then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S" + using assms mem_convex_alt[of S xx yy cx cy] x y by auto + then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S" + using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close> + by (auto simp add: scaleR_right_distrib) + then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" + using x y by auto + } + moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto + ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast +qed + +lemma cone_convex_hull: + assumes "cone S" + shows "cone (convex hull S)" +proof (cases "S = {}") + case True + then show ?thesis by auto +next + case False + then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" + using cone_iff[of S] assms by auto + { + fix c :: real + assume "c > 0" + then have "op *\<^sub>R c ` (convex hull S) = convex hull (op *\<^sub>R c ` S)" + using convex_hull_scaling[of _ S] by auto + also have "\<dots> = convex hull S" + using * \<open>c > 0\<close> by auto + finally have "op *\<^sub>R c ` (convex hull S) = convex hull S" + by auto + } + then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> (op *\<^sub>R c ` (convex hull S)) = (convex hull S)" + using * hull_subset[of S convex] by auto + then show ?thesis + using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto +qed + +subsection \<open>Convex set as intersection of halfspaces\<close> + +lemma convex_halfspace_intersection: + fixes s :: "('a::euclidean_space) set" + assumes "closed s" "convex s" + shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}" + apply (rule set_eqI) + apply rule + unfolding Inter_iff Ball_def mem_Collect_eq + apply (rule,rule,erule conjE) +proof - + fix x + assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa" + then have "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}" + by blast + then show "x \<in> s" + apply (rule_tac ccontr) + apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) + apply (erule exE)+ + apply (erule_tac x="-a" in allE) + apply (erule_tac x="-b" in allE) + apply auto + done +qed auto + + +subsection \<open>Radon's theorem (from Lars Schewe)\<close> + +lemma radon_ex_lemma: + assumes "finite c" "affine_dependent c" + shows "\<exists>u. setsum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) c = 0" +proof - + from assms(2)[unfolded affine_dependent_explicit] + obtain s u where + "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" + by blast + then show ?thesis + apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI) + unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms(1), symmetric] + apply (auto simp add: Int_absorb1) + done +qed + +lemma radon_s_lemma: + assumes "finite s" + and "setsum f s = (0::real)" + shows "setsum f {x\<in>s. 0 < f x} = - setsum f {x\<in>s. f x < 0}" +proof - + have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" + by auto + show ?thesis + unfolding add_eq_0_iff[symmetric] and setsum.inter_filter[OF assms(1)] + and setsum.distrib[symmetric] and * + using assms(2) + by assumption +qed + +lemma radon_v_lemma: + assumes "finite s" + and "setsum f s = 0" + and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)" + shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}" +proof - + have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" + using assms(3) by auto + show ?thesis + unfolding eq_neg_iff_add_eq_0 and setsum.inter_filter[OF assms(1)] + and setsum.distrib[symmetric] and * + using assms(2) + apply assumption + done +qed + +lemma radon_partition: + assumes "finite c" "affine_dependent c" + shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}" +proof - + obtain u v where uv: "setsum u c = 0" "v\<in>c" "u v \<noteq> 0" "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0" + using radon_ex_lemma[OF assms] by auto + have fin: "finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}" + using assms(1) by auto + define z where "z = inverse (setsum u {x\<in>c. u x > 0}) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}" + have "setsum u {x \<in> c. 0 < u x} \<noteq> 0" + proof (cases "u v \<ge> 0") + case False + then have "u v < 0" by auto + then show ?thesis + proof (cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0") + case True + then show ?thesis + using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto + next + case False + then have "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c" + apply (rule_tac setsum_mono) + apply auto + done + then show ?thesis + unfolding setsum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto + qed + qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) + + then have *: "setsum u {x\<in>c. u x > 0} > 0" + unfolding less_le + apply (rule_tac conjI) + apply (rule_tac setsum_nonneg) + apply auto + done + moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c" + "(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)" + using assms(1) + apply (rule_tac[!] setsum.mono_neutral_left) + apply auto + done + then have "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}" + "(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)" + unfolding eq_neg_iff_add_eq_0 + using uv(1,4) + by (auto simp add: setsum.union_inter_neutral[OF fin, symmetric]) + moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x" + apply rule + apply (rule mult_nonneg_nonneg) + using * + apply auto + done + ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}" + unfolding convex_hull_explicit mem_Collect_eq + apply (rule_tac x="{v \<in> c. u v < 0}" in exI) + apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI) + using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def + apply (auto simp add: setsum_negf setsum_right_distrib[symmetric]) + done + moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x" + apply rule + apply (rule mult_nonneg_nonneg) + using * + apply auto + done + then have "z \<in> convex hull {v \<in> c. u v > 0}" + unfolding convex_hull_explicit mem_Collect_eq + apply (rule_tac x="{v \<in> c. 0 < u v}" in exI) + apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI) + using assms(1) + unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def + using * + apply (auto simp add: setsum_negf setsum_right_distrib[symmetric]) + done + ultimately show ?thesis + apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI) + apply (rule_tac x="{v\<in>c. u v > 0}" in exI) + apply auto + done +qed + +lemma radon: + assumes "affine_dependent c" + obtains m p where "m \<subseteq> c" "p \<subseteq> c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}" +proof - + from assms[unfolded affine_dependent_explicit] + obtain s u where + "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" + by blast + then have *: "finite s" "affine_dependent s" and s: "s \<subseteq> c" + unfolding affine_dependent_explicit by auto + from radon_partition[OF *] + obtain m p where "m \<inter> p = {}" "m \<union> p = s" "convex hull m \<inter> convex hull p \<noteq> {}" + by blast + then show ?thesis + apply (rule_tac that[of p m]) + using s + apply auto + done +qed + + +subsection \<open>Helly's theorem\<close> + +lemma helly_induct: + fixes f :: "'a::euclidean_space set set" + assumes "card f = n" + and "n \<ge> DIM('a) + 1" + and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}" + shows "\<Inter>f \<noteq> {}" + using assms +proof (induct n arbitrary: f) + case 0 + then show ?case by auto +next + case (Suc n) + have "finite f" + using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite) + show "\<Inter>f \<noteq> {}" + apply (cases "n = DIM('a)") + apply (rule Suc(5)[rule_format]) + unfolding \<open>card f = Suc n\<close> + proof - + assume ng: "n \<noteq> DIM('a)" + then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" + apply (rule_tac bchoice) + unfolding ex_in_conv + apply (rule, rule Suc(1)[rule_format]) + unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close> + defer + defer + apply (rule Suc(4)[rule_format]) + defer + apply (rule Suc(5)[rule_format]) + using Suc(3) \<open>finite f\<close> + apply auto + done + then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto + show ?thesis + proof (cases "inj_on X f") + case False + then obtain s t where st: "s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t" + unfolding inj_on_def by auto + then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto + show ?thesis + unfolding * + unfolding ex_in_conv[symmetric] + apply (rule_tac x="X s" in exI) + apply rule + apply (rule X[rule_format]) + using X st + apply auto + done + next + case True + then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}" + using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] + unfolding card_image[OF True] and \<open>card f = Suc n\<close> + using Suc(3) \<open>finite f\<close> and ng + by auto + have "m \<subseteq> X ` f" "p \<subseteq> X ` f" + using mp(2) by auto + then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" + unfolding subset_image_iff by auto + then have "f \<union> (g \<union> h) = f" by auto + then have f: "f = g \<union> h" + using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True + unfolding mp(2)[unfolded image_Un[symmetric] gh] + by auto + have *: "g \<inter> h = {}" + using mp(1) + unfolding gh + using inj_on_image_Int[OF True gh(3,4)] + by auto + have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h" + apply (rule_tac [!] hull_minimal) + using Suc gh(3-4) + unfolding subset_eq + apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) + apply rule + prefer 3 + apply rule + proof - + fix x + assume "x \<in> X ` g" + then obtain y where "y \<in> g" "x = X y" + unfolding image_iff .. + then show "x \<in> \<Inter>h" + using X[THEN bspec[where x=y]] using * f by auto + next + fix x + assume "x \<in> X ` h" + then obtain y where "y \<in> h" "x = X y" + unfolding image_iff .. + then show "x \<in> \<Inter>g" + using X[THEN bspec[where x=y]] using * f by auto + qed auto + then show ?thesis + unfolding f using mp(3)[unfolded gh] by blast + qed + qed auto +qed + +lemma helly: + fixes f :: "'a::euclidean_space set set" + assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s" + and "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}" + shows "\<Inter>f \<noteq> {}" + apply (rule helly_induct) + using assms + apply auto + done + + +subsection \<open>Epigraphs of convex functions\<close> + +definition "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}" + +lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" + unfolding epigraph_def by auto + +lemma convex_epigraph: "convex (epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s" + unfolding convex_def convex_on_def + unfolding Ball_def split_paired_All epigraph_def + unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric] + apply safe + defer + apply (erule_tac x=x in allE) + apply (erule_tac x="f x" in allE) + apply safe + apply (erule_tac x=xa in allE) + apply (erule_tac x="f xa" in allE) + prefer 3 + apply (rule_tac y="u * f a + v * f aa" in order_trans) + defer + apply (auto intro!:mult_left_mono add_mono) + done + +lemma convex_epigraphI: "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex (epigraph s f)" + unfolding convex_epigraph by auto + +lemma convex_epigraph_convex: "convex s \<Longrightarrow> convex_on s f \<longleftrightarrow> convex(epigraph s f)" + by (simp add: convex_epigraph) + + +subsubsection \<open>Use this to derive general bound property of convex function\<close> + +lemma convex_on: + assumes "convex s" + shows "convex_on s f \<longleftrightarrow> + (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow> + f (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> setsum (\<lambda>i. u i * f(x i)) {1..k})" + unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq + unfolding fst_setsum snd_setsum fst_scaleR snd_scaleR + apply safe + apply (drule_tac x=k in spec) + apply (drule_tac x=u in spec) + apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec) + apply simp + using assms[unfolded convex] + apply simp + apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans) + defer + apply (rule setsum_mono) + apply (erule_tac x=i in allE) + unfolding real_scaleR_def + apply (rule mult_left_mono) + using assms[unfolded convex] + apply auto + done + + +subsection \<open>Convexity of general and special intervals\<close> + +lemma is_interval_convex: + fixes s :: "'a::euclidean_space set" + assumes "is_interval s" + shows "convex s" +proof (rule convexI) + fix x y and u v :: real + assume as: "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = 1" + then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0" + by auto + { + fix a b + assume "\<not> b \<le> u * a + v * b" + then have "u * a < (1 - v) * b" + unfolding not_le using as(4) by (auto simp add: field_simps) + then have "a < b" + unfolding * using as(4) *(2) + apply (rule_tac mult_left_less_imp_less[of "1 - v"]) + apply (auto simp add: field_simps) + done + then have "a \<le> u * a + v * b" + unfolding * using as(4) + by (auto simp add: field_simps intro!:mult_right_mono) + } + moreover + { + fix a b + assume "\<not> u * a + v * b \<le> a" + then have "v * b > (1 - u) * a" + unfolding not_le using as(4) by (auto simp add: field_simps) + then have "a < b" + unfolding * using as(4) + apply (rule_tac mult_left_less_imp_less) + apply (auto simp add: field_simps) + done + then have "u * a + v * b \<le> b" + unfolding ** + using **(2) as(3) + by (auto simp add: field_simps intro!:mult_right_mono) + } + ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" + apply - + apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) + using as(3-) DIM_positive[where 'a='a] + apply (auto simp: inner_simps) + done +qed + +lemma is_interval_connected: + fixes s :: "'a::euclidean_space set" + shows "is_interval s \<Longrightarrow> connected s" + using is_interval_convex convex_connected by auto + +lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" + apply (rule_tac[!] is_interval_convex)+ + using is_interval_box is_interval_cbox + apply auto + done + +subsection \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent.\<close> + +lemma is_interval_1: + "is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)" + unfolding is_interval_def by auto + +lemma is_interval_connected_1: + fixes s :: "real set" + shows "is_interval s \<longleftrightarrow> connected s" + apply rule + apply (rule is_interval_connected, assumption) + unfolding is_interval_1 + apply rule + apply rule + apply rule + apply rule + apply (erule conjE) + apply (rule ccontr) +proof - + fix a b x + assume as: "connected s" "a \<in> s" "b \<in> s" "a \<le> x" "x \<le> b" "x \<notin> s" + then have *: "a < x" "x < b" + unfolding not_le [symmetric] by auto + let ?halfl = "{..<x} " + let ?halfr = "{x<..}" + { + fix y + assume "y \<in> s" + with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto + then have "y \<in> ?halfr \<union> ?halfl" by auto + } + moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto + then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}" + using as(2-3) by auto + ultimately show False + apply (rule_tac notE[OF as(1)[unfolded connected_def]]) + apply (rule_tac x = ?halfl in exI) + apply (rule_tac x = ?halfr in exI) + apply rule + apply (rule open_lessThan) + apply rule + apply (rule open_greaterThan) + apply auto + done +qed + +lemma is_interval_convex_1: + fixes s :: "real set" + shows "is_interval s \<longleftrightarrow> convex s" + by (metis is_interval_convex convex_connected is_interval_connected_1) + +lemma connected_convex_1: + fixes s :: "real set" + shows "connected s \<longleftrightarrow> convex s" + by (metis is_interval_convex convex_connected is_interval_connected_1) + +lemma connected_convex_1_gen: + fixes s :: "'a :: euclidean_space set" + assumes "DIM('a) = 1" + shows "connected s \<longleftrightarrow> convex s" +proof - + obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f" + using subspace_isomorphism [where 'a = 'a and 'b = real] + by (metis DIM_real dim_UNIV subspace_UNIV assms) + then have "f -` (f ` s) = s" + by (simp add: inj_vimage_image_eq) + then show ?thesis + by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image) +qed + +subsection \<open>Another intermediate value theorem formulation\<close> + +lemma ivt_increasing_component_on_1: + fixes f :: "real \<Rightarrow> 'a::euclidean_space" + assumes "a \<le> b" + and "continuous_on {a..b} f" + and "(f a)\<bullet>k \<le> y" "y \<le> (f b)\<bullet>k" + shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y" +proof - + have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b" + apply (rule_tac[!] imageI) + using assms(1) + apply auto + done + then show ?thesis + using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] + by (simp add: Topology_Euclidean_Space.connected_continuous_image assms) +qed + +lemma ivt_increasing_component_1: + fixes f :: "real \<Rightarrow> 'a::euclidean_space" + shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow> + f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y" + by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on) + +lemma ivt_decreasing_component_on_1: + fixes f :: "real \<Rightarrow> 'a::euclidean_space" + assumes "a \<le> b" + and "continuous_on {a..b} f" + and "(f b)\<bullet>k \<le> y" + and "y \<le> (f a)\<bullet>k" + shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y" + apply (subst neg_equal_iff_equal[symmetric]) + using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"] + using assms using continuous_on_minus + apply auto + done + +lemma ivt_decreasing_component_1: + fixes f :: "real \<Rightarrow> 'a::euclidean_space" + shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow> + f b\<bullet>k \<le> y \<Longrightarrow> y \<le> f a\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y" + by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) + + +subsection \<open>A bound within a convex hull, and so an interval\<close> + +lemma convex_on_convex_hull_bound: + assumes "convex_on (convex hull s) f" + and "\<forall>x\<in>s. f x \<le> b" + shows "\<forall>x\<in> convex hull s. f x \<le> b" +proof + fix x + assume "x \<in> convex hull s" + then obtain k u v where + obt: "\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x" + unfolding convex_hull_indexed mem_Collect_eq by auto + have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b" + using setsum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"] + unfolding setsum_left_distrib[symmetric] obt(2) mult_1 + apply (drule_tac meta_mp) + apply (rule mult_left_mono) + using assms(2) obt(1) + apply auto + done + then show "f x \<le> b" + using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] + unfolding obt(2-3) + using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] + by auto +qed + +lemma inner_setsum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1" + by (simp add: inner_setsum_left setsum.If_cases inner_Basis) + +lemma convex_set_plus: + assumes "convex s" and "convex t" shows "convex (s + t)" +proof - + have "convex {x + y |x y. x \<in> s \<and> y \<in> t}" + using assms by (rule convex_sums) + moreover have "{x + y |x y. x \<in> s \<and> y \<in> t} = s + t" + unfolding set_plus_def by auto + finally show "convex (s + t)" . +qed + +lemma convex_set_setsum: + assumes "\<And>i. i \<in> A \<Longrightarrow> convex (B i)" + shows "convex (\<Sum>i\<in>A. B i)" +proof (cases "finite A") + case True then show ?thesis using assms + by induct (auto simp: convex_set_plus) +qed auto + +lemma finite_set_setsum: + assumes "finite A" and "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)" + using assms by (induct set: finite, simp, simp add: finite_set_plus) + +lemma set_setsum_eq: + "finite A \<Longrightarrow> (\<Sum>i\<in>A. B i) = {\<Sum>i\<in>A. f i |f. \<forall>i\<in>A. f i \<in> B i}" + apply (induct set: finite) + apply simp + apply simp + apply (safe elim!: set_plus_elim) + apply (rule_tac x="fun_upd f x a" in exI) + apply simp + apply (rule_tac f="\<lambda>x. a + x" in arg_cong) + apply (rule setsum.cong [OF refl]) + apply clarsimp + apply fast + done + +lemma box_eq_set_setsum_Basis: + shows "{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. image (\<lambda>x. x *\<^sub>R i) (B i))" + apply (subst set_setsum_eq [OF finite_Basis]) + apply safe + apply (fast intro: euclidean_representation [symmetric]) + apply (subst inner_setsum_left) + apply (subgoal_tac "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i") + apply (drule (1) bspec) + apply clarsimp + apply (frule setsum.remove [OF finite_Basis]) + apply (erule trans) + apply simp + apply (rule setsum.neutral) + apply clarsimp + apply (frule_tac x=i in bspec, assumption) + apply (drule_tac x=x in bspec, assumption) + apply clarsimp + apply (cut_tac u=x and v=i in inner_Basis, assumption+) + apply (rule ccontr) + apply simp + done + +lemma convex_hull_set_setsum: + "convex hull (\<Sum>i\<in>A. B i) = (\<Sum>i\<in>A. convex hull (B i))" +proof (cases "finite A") + assume "finite A" then show ?thesis + by (induct set: finite, simp, simp add: convex_hull_set_plus) +qed simp + +lemma convex_hull_eq_real_cbox: + fixes x y :: real assumes "x \<le> y" + shows "convex hull {x, y} = cbox x y" +proof (rule hull_unique) + show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto + show "convex (cbox x y)" + by (rule convex_box) +next + fix s assume "{x, y} \<subseteq> s" and "convex s" + then show "cbox x y \<subseteq> s" + unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def + by - (clarify, simp (no_asm_use), fast) +qed + +lemma unit_interval_convex_hull: + "cbox (0::'a::euclidean_space) One = convex hull {x. \<forall>i\<in>Basis. (x\<bullet>i = 0) \<or> (x\<bullet>i = 1)}" + (is "?int = convex hull ?points") +proof - + have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1" + by (simp add: inner_setsum_left setsum.If_cases inner_Basis) + have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}" + by (auto simp: cbox_def) + also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)" + by (simp only: box_eq_set_setsum_Basis) + also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))" + by (simp only: convex_hull_eq_real_cbox zero_le_one) + also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))" + by (simp only: convex_hull_linear_image linear_scaleR_left) + also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})" + by (simp only: convex_hull_set_setsum) + also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}" + by (simp only: box_eq_set_setsum_Basis) + also have "convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}} = convex hull ?points" + by simp + finally show ?thesis . +qed + +text \<open>And this is a finite set of vertices.\<close> + +lemma unit_cube_convex_hull: + obtains s :: "'a::euclidean_space set" + where "finite s" and "cbox 0 (\<Sum>Basis) = convex hull s" + apply (rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"]) + apply (rule finite_subset[of _ "(\<lambda>s. (\<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"]) + prefer 3 + apply (rule unit_interval_convex_hull) + apply rule + unfolding mem_Collect_eq +proof - + fix x :: 'a + assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1" + show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis" + apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"]) + using as + apply (subst euclidean_eq_iff) + apply auto + done +qed auto + +text \<open>Hence any cube (could do any nonempty interval).\<close> + +lemma cube_convex_hull: + assumes "d > 0" + obtains s :: "'a::euclidean_space set" where + "finite s" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull s" +proof - + let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a" + have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)" + apply (rule set_eqI, rule) + unfolding image_iff + defer + apply (erule bexE) + proof - + fix y + assume as: "y\<in>cbox (x - ?d) (x + ?d)" + then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)" + using assms by (simp add: mem_box field_simps inner_simps) + with \<open>0 < d\<close> show "\<exists>z\<in>cbox 0 (\<Sum>Basis). y = x - ?d + (2 * d) *\<^sub>R z" + by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto + next + fix y z + assume as: "z\<in>cbox 0 (\<Sum>Basis)" "y = x - ?d + (2*d) *\<^sub>R z" + have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d" + using assms as(1)[unfolded mem_box] + apply (erule_tac x=i in ballE) + apply rule + prefer 2 + apply (rule mult_right_le_one_le) + using assms + apply auto + done + then show "y \<in> cbox (x - ?d) (x + ?d)" + unfolding as(2) mem_box + apply - + apply rule + using as(1)[unfolded mem_box] + apply (erule_tac x=i in ballE) + using assms + apply (auto simp: inner_simps) + done + qed + obtain s where "finite s" "cbox 0 (\<Sum>Basis::'a) = convex hull s" + using unit_cube_convex_hull by auto + then show ?thesis + apply (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) + unfolding * and convex_hull_affinity + apply auto + done +qed + +subsubsection\<open>Representation of any interval as a finite convex hull\<close> + +lemma image_stretch_interval: + "(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = + (if (cbox a b) = {} then {} else + cbox (\<Sum>k\<in>Basis. (min (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k::'a) + (\<Sum>k\<in>Basis. (max (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k))" +proof cases + assume *: "cbox a b \<noteq> {}" + show ?thesis + unfolding box_ne_empty if_not_P[OF *] + apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) + apply (subst choice_Basis_iff[symmetric]) + proof (intro allI ball_cong refl) + fix x i :: 'a assume "i \<in> Basis" + with * have a_le_b: "a \<bullet> i \<le> b \<bullet> i" + unfolding box_ne_empty by auto + show "(\<exists>xa. x \<bullet> i = m i * xa \<and> a \<bullet> i \<le> xa \<and> xa \<le> b \<bullet> i) \<longleftrightarrow> + min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) \<le> x \<bullet> i \<and> x \<bullet> i \<le> max (m i * (a \<bullet> i)) (m i * (b \<bullet> i))" + proof (cases "m i = 0") + case True + with a_le_b show ?thesis by auto + next + case False + then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i" + by (auto simp add: field_simps) + from False have + "min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (a \<bullet> i) else m i * (b \<bullet> i))" + "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (b \<bullet> i) else m i * (a \<bullet> i))" + using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) + with False show ?thesis using a_le_b + unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps) + qed + qed +qed simp + +lemma interval_image_stretch_interval: + "\<exists>u v. (\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" + unfolding image_stretch_interval by auto + +lemma cbox_translation: "cbox (c + a) (c + b) = image (\<lambda>x. c + x) (cbox a b)" + using image_affinity_cbox [of 1 c a b] + using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] + by (auto simp add: inner_left_distrib add.commute) + +lemma cbox_image_unit_interval: + fixes a :: "'a::euclidean_space" + assumes "cbox a b \<noteq> {}" + shows "cbox a b = + op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One" +using assms +apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) +apply (simp add: min_def max_def algebra_simps setsum_subtractf euclidean_representation) +done + +lemma closed_interval_as_convex_hull: + fixes a :: "'a::euclidean_space" + obtains s where "finite s" "cbox a b = convex hull s" +proof (cases "cbox a b = {}") + case True with convex_hull_empty that show ?thesis + by blast +next + case False + obtain s::"'a set" where "finite s" and eq: "cbox 0 One = convex hull s" + by (blast intro: unit_cube_convex_hull) + have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)" + by (rule linear_compose_setsum) (auto simp: algebra_simps linearI) + have "finite (op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` s)" + by (rule finite_imageI \<open>finite s\<close>)+ + then show ?thesis + apply (rule that) + apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric]) + apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False]) + done +qed + + +subsection \<open>Bounded convex function on open set is continuous\<close> + +lemma convex_on_bounded_continuous: + fixes s :: "('a::real_normed_vector) set" + assumes "open s" + and "convex_on s f" + and "\<forall>x\<in>s. \<bar>f x\<bar> \<le> b" + shows "continuous_on s f" + apply (rule continuous_at_imp_continuous_on) + unfolding continuous_at_real_range +proof (rule,rule,rule) + fix x and e :: real + assume "x \<in> s" "e > 0" + define B where "B = \<bar>b\<bar> + 1" + have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B" + unfolding B_def + defer + apply (drule assms(3)[rule_format]) + apply auto + done + obtain k where "k > 0" and k: "cball x k \<subseteq> s" + using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] + using \<open>x\<in>s\<close> by auto + show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e" + apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) + apply rule + defer + proof (rule, rule) + fix y + assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" + show "\<bar>f y - f x\<bar> < e" + proof (cases "y = x") + case False + define t where "t = k / norm (y - x)" + have "2 < t" "0<t" + unfolding t_def using as False and \<open>k>0\<close> + by (auto simp add:field_simps) + have "y \<in> s" + apply (rule k[unfolded subset_eq,rule_format]) + unfolding mem_cball dist_norm + apply (rule order_trans[of _ "2 * norm (x - y)"]) + using as + by (auto simp add: field_simps norm_minus_commute) + { + define w where "w = x + t *\<^sub>R (y - x)" + have "w \<in> s" + unfolding w_def + apply (rule k[unfolded subset_eq,rule_format]) + unfolding mem_cball dist_norm + unfolding t_def + using \<open>k>0\<close> + apply auto + done + have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" + by (auto simp add: algebra_simps) + also have "\<dots> = 0" + using \<open>t > 0\<close> by (auto simp add:field_simps) + finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" + unfolding w_def using False and \<open>t > 0\<close> + by (auto simp add: algebra_simps) + have "2 * B < e * t" + unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False + by (auto simp add:field_simps) + then have "(f w - f x) / t < e" + using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>] + using \<open>t > 0\<close> by (auto simp add:field_simps) + then have th1: "f y - f x < e" + apply - + apply (rule le_less_trans) + defer + apply assumption + using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] + using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close> + by (auto simp add:field_simps) + } + moreover + { + define w where "w = x - t *\<^sub>R (y - x)" + have "w \<in> s" + unfolding w_def + apply (rule k[unfolded subset_eq,rule_format]) + unfolding mem_cball dist_norm + unfolding t_def + using \<open>k > 0\<close> + apply auto + done + have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" + by (auto simp add: algebra_simps) + also have "\<dots> = x" + using \<open>t > 0\<close> by (auto simp add:field_simps) + finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" + unfolding w_def using False and \<open>t > 0\<close> + by (auto simp add: algebra_simps) + have "2 * B < e * t" + unfolding t_def + using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False + by (auto simp add:field_simps) + then have *: "(f w - f y) / t < e" + using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>] + using \<open>t > 0\<close> + by (auto simp add:field_simps) + have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" + using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] + using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close> + by (auto simp add:field_simps) + also have "\<dots> = (f w + t * f y) / (1 + t)" + using \<open>t > 0\<close> by (auto simp add: divide_simps) + also have "\<dots> < e + f y" + using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps) + finally have "f x - f y < e" by auto + } + ultimately show ?thesis by auto + qed (insert \<open>0<e\<close>, auto) + qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps) +qed + + +subsection \<open>Upper bound on a ball implies upper and lower bounds\<close> + +lemma convex_bounds_lemma: + fixes x :: "'a::real_normed_vector" + assumes "convex_on (cball x e) f" + and "\<forall>y \<in> cball x e. f y \<le> b" + shows "\<forall>y \<in> cball x e. \<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" + apply rule +proof (cases "0 \<le> e") + case True + fix y + assume y: "y \<in> cball x e" + define z where "z = 2 *\<^sub>R x - y" + have *: "x - (2 *\<^sub>R x - y) = y - x" + by (simp add: scaleR_2) + have z: "z \<in> cball x e" + using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute) + have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" + unfolding z_def by (auto simp add: algebra_simps) + then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" + using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] + using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] + by (auto simp add:field_simps) +next + case False + fix y + assume "y \<in> cball x e" + then have "dist x y < 0" + using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) + then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" + using zero_le_dist[of x y] by auto +qed + + +subsubsection \<open>Hence a convex function on an open set is continuous\<close> + +lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n" + by auto + +lemma convex_on_continuous: + assumes "open (s::('a::euclidean_space) set)" "convex_on s f" + shows "continuous_on s f" + unfolding continuous_on_eq_continuous_at[OF assms(1)] +proof + note dimge1 = DIM_positive[where 'a='a] + fix x + assume "x \<in> s" + then obtain e where e: "cball x e \<subseteq> s" "e > 0" + using assms(1) unfolding open_contains_cball by auto + define d where "d = e / real DIM('a)" + have "0 < d" + unfolding d_def using \<open>e > 0\<close> dimge1 by auto + let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a" + obtain c + where c: "finite c" + and c1: "convex hull c \<subseteq> cball x e" + and c2: "cball x d \<subseteq> convex hull c" + proof + define c where "c = (\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d})" + show "finite c" + unfolding c_def by (simp add: finite_set_setsum) + have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}" + unfolding box_eq_set_setsum_Basis + unfolding c_def convex_hull_set_setsum + apply (subst convex_hull_linear_image [symmetric]) + apply (simp add: linear_iff scaleR_add_left) + apply (rule setsum.cong [OF refl]) + apply (rule image_cong [OF _ refl]) + apply (rule convex_hull_eq_real_cbox) + apply (cut_tac \<open>0 < d\<close>, simp) + done + then have 2: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cball (x \<bullet> i) d}" + by (simp add: dist_norm abs_le_iff algebra_simps) + show "cball x d \<subseteq> convex hull c" + unfolding 2 + apply clarsimp + apply (simp only: dist_norm) + apply (subst inner_diff_left [symmetric]) + apply simp + apply (erule (1) order_trans [OF Basis_le_norm]) + done + have e': "e = (\<Sum>(i::'a)\<in>Basis. d)" + by (simp add: d_def DIM_positive) + show "convex hull c \<subseteq> cball x e" + unfolding 2 + apply clarsimp + apply (subst euclidean_dist_l2) + apply (rule order_trans [OF setL2_le_setsum]) + apply (rule zero_le_dist) + unfolding e' + apply (rule setsum_mono) + apply simp + done + qed + define k where "k = Max (f ` c)" + have "convex_on (convex hull c) f" + apply(rule convex_on_subset[OF assms(2)]) + apply(rule subset_trans[OF _ e(1)]) + apply(rule c1) + done + then have k: "\<forall>y\<in>convex hull c. f y \<le> k" + apply (rule_tac convex_on_convex_hull_bound) + apply assumption + unfolding k_def + apply (rule, rule Max_ge) + using c(1) + apply auto + done + have "d \<le> e" + unfolding d_def + apply (rule mult_imp_div_pos_le) + using \<open>e > 0\<close> + unfolding mult_le_cancel_left1 + apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive) + done + then have dsube: "cball x d \<subseteq> cball x e" + by (rule subset_cball) + have conv: "convex_on (cball x d) f" + apply (rule convex_on_subset) + apply (rule convex_on_subset[OF assms(2)]) + apply (rule e(1)) + apply (rule dsube) + done + then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>" + apply (rule convex_bounds_lemma) + apply (rule ballI) + apply (rule k [rule_format]) + apply (erule rev_subsetD) + apply (rule c2) + done + then have "continuous_on (ball x d) f" + apply (rule_tac convex_on_bounded_continuous) + apply (rule open_ball, rule convex_on_subset[OF conv]) + apply (rule ball_subset_cball) + apply force + done + then show "continuous (at x) f" + unfolding continuous_on_eq_continuous_at[OF open_ball] + using \<open>d > 0\<close> by auto +qed + + +subsection \<open>Line segments, Starlike Sets, etc.\<close> + +definition midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a" + where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" + +definition closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" + where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}" + +definition open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where + "open_segment a b \<equiv> closed_segment a b - {a,b}" + +lemmas segment = open_segment_def closed_segment_def + +lemma in_segment: + "x \<in> closed_segment a b \<longleftrightarrow> (\<exists>u. 0 \<le> u \<and> u \<le> 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" + "x \<in> open_segment a b \<longleftrightarrow> a \<noteq> b \<and> (\<exists>u. 0 < u \<and> u < 1 \<and> 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: + "linear f \<Longrightarrow> closed_segment (f a) (f b) = f ` (closed_segment a b)" + by (force simp add: in_segment linear_add_cmul) + +lemma open_segment_linear_image: + "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> 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 (\<lambda>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 (\<lambda>x. c + x) (open_segment a b)" +by (simp add: open_segment_def closed_segment_translation translation_diff) + +lemma open_segment_PairD: + "(x, x') \<in> open_segment (a, a') (b, b') + \<Longrightarrow> (x \<in> open_segment a b \<or> a = b) \<and> (x' \<in> open_segment a' b' \<or> a' = b')" + by (auto simp: in_segment) + +lemma closed_segment_PairD: + "(x, x') \<in> closed_segment (a, a') (b, b') \<Longrightarrow> x \<in> closed_segment a b \<and> x' \<in> closed_segment a' b'" + by (auto simp: closed_segment_def) + +lemma closed_segment_translation_eq [simp]: + "d + x \<in> closed_segment (d + a) (d + b) \<longleftrightarrow> x \<in> closed_segment a b" +proof - + have *: "\<And>d x a b. x \<in> closed_segment a b \<Longrightarrow> d + x \<in> 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 \<in> open_segment (d + a) (d + b) \<longleftrightarrow> x \<in> open_segment a b" + by (simp add: open_segment_def) + +definition "between = (\<lambda>(a,b) x. x \<in> closed_segment a b)" + +definition "starlike s \<longleftrightarrow> (\<exists>a\<in>s. \<forall>x\<in>s. closed_segment a x \<subseteq> s)" + +lemma starlike_UNIV [simp]: "starlike UNIV" + by (simp add: starlike_def) + +lemma midpoint_refl: "midpoint x x = x" + unfolding midpoint_def + unfolding scaleR_right_distrib + unfolding scaleR_left_distrib[symmetric] + by auto + +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 \<longleftrightarrow> a + b = c + c" +proof - + have "midpoint a b = c \<longleftrightarrow> scaleR 2 (midpoint a b) = scaleR 2 c" + by simp + then show ?thesis + unfolding midpoint_def scaleR_2 [symmetric] by simp +qed + +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 *: "\<And>x y::'a. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2" + unfolding equation_minus_iff by auto + have **: "\<And>x y::'a. 2 *\<^sub>R x = y \<Longrightarrow> 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 \<longleftrightarrow> a = b" + "midpoint a b = b \<longleftrightarrow> a = b" + unfolding midpoint_eq_iff by auto + +lemma convex_contains_segment: + "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. closed_segment a b \<subseteq> s)" + unfolding convex_alt closed_segment_def by auto + +lemma closed_segment_subset: "\<lbrakk>x \<in> s; y \<in> s; convex s\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> s" + by (simp add: convex_contains_segment) + +lemma closed_segment_subset_convex_hull: + "\<lbrakk>x \<in> convex hull s; y \<in> convex hull s\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> convex hull s" + using convex_contains_segment by blast + +lemma convex_imp_starlike: + "convex s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> starlike s" + unfolding convex_contains_segment starlike_def by auto + +lemma segment_convex_hull: + "closed_segment a b = convex hull {a,b}" +proof - + have *: "\<And>x. {x} \<noteq> {}" 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 \<in> open_segment w z \<Longrightarrow> u \<in> closed_segment w z" + by (auto simp add: closed_segment_def open_segment_def) + +lemma segment_open_subset_closed: + "open_segment a b \<subseteq> closed_segment a b" + by (auto simp: closed_segment_def open_segment_def) + +lemma bounded_closed_segment: + fixes a :: "'a::euclidean_space" shows "bounded (closed_segment a b)" + by (simp add: segment_convex_hull compact_convex_hull compact_imp_bounded) + +lemma bounded_open_segment: + fixes a :: "'a::euclidean_space" 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 \<in> closed_segment a b" "b \<in> closed_segment a b" + unfolding segment_convex_hull + by (auto intro!: hull_subset[unfolded subset_eq, rule_format]) + +lemma segment_furthest_le: + fixes a b x y :: "'a::euclidean_space" + assumes "x \<in> closed_segment a b" + shows "norm (y - x) \<le> norm (y - a) \<or> norm (y - x) \<le> norm (y - b)" +proof - + obtain z where "z \<in> {a, b}" "norm (x - y) \<le> norm (z - y)" + using simplex_furthest_le[of "{a, b}" y] + using assms[unfolded segment_convex_hull] + by auto + then show ?thesis + by (auto simp add:norm_minus_commute) +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 \<in> closed_segment a b" + shows "norm (x - a) \<le> norm (b - a)" +proof - + obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" + using assms by (auto simp add: closed_segment_def) + then show "norm (x - a) \<le> 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 \<in> closed_segment a b" + shows "norm (x - a) \<le> norm (b - a)" "norm (x - b) \<le> norm (b - a)" +apply (simp add: assms segment_bound1) +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 \<union> {a,b}" + using open_segment_def by auto + +lemma convex_contains_open_segment: + "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. open_segment a b \<subseteq> s)" + by (simp add: convex_contains_segment closed_segment_eq_open) + +lemma closed_segment_eq_real_ivl: + fixes a b::real + shows "closed_segment a b = (if a \<le> b then {a .. b} else {b .. a})" +proof - + have "b \<le> a \<Longrightarrow> closed_segment b a = {b .. a}" + and "a \<le> b \<Longrightarrow> closed_segment a b = {a .. b}" + by (auto simp: convex_hull_eq_real_cbox segment_convex_hull) + thus ?thesis + by (auto simp: closed_segment_commute) +qed + +lemma closed_segment_real_eq: + fixes u::real shows "closed_segment u v = (\<lambda>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 dist_in_closed_segment: + fixes a :: "'a :: euclidean_space" + assumes "x \<in> closed_segment a b" + shows "dist x a \<le> dist a b \<and> dist x b \<le> dist a b" +proof (intro conjI) + obtain u where u: "0 \<le> u" "u \<le> 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 \<open>0 \<le> u\<close> abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib) + also have "... \<le> dist a b" + by (simp add: mult_left_le_one_le u) + finally show "dist x a \<le> 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 \<open>u \<le> 1\<close> by force + then show ?thesis + by (simp add: dist_norm real_vector.scale_right_diff_distrib) + qed + also have "... \<le> dist a b" + by (simp add: mult_left_le_one_le u) + finally show "dist x b \<le> dist a b" . +qed + +lemma dist_in_open_segment: + fixes a :: "'a :: euclidean_space" + assumes "x \<in> open_segment a b" + shows "dist x a < dist a b \<and> 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 \<open>0 < u\<close>) + 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 \<open>u < 1\<close>) + finally show "dist x a < dist a b" . + have ab_ne0: "dist a b \<noteq> 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 \<open>u < 1\<close> 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 \<open>0 < u\<close> by simp + finally show "dist x b < dist a b" . +qed + +lemma dist_decreases_open_segment_0: + fixes x :: "'a :: euclidean_space" + assumes "x \<in> open_segment 0 b" + shows "dist c x < dist c 0 \<or> dist c x < dist c b" +proof (rule ccontr, clarsimp simp: not_less) + obtain u where u: "0 \<noteq> b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b" + using assms by (auto simp: in_segment) + have xb: "x \<bullet> b < b \<bullet> b" + using u x by auto + assume "norm c \<le> dist c x" + then have "c \<bullet> c \<le> (c - x) \<bullet> (c - x)" + by (simp add: dist_norm norm_le) + moreover have "0 < x \<bullet> b" + using u x by auto + ultimately have less: "c \<bullet> b < x \<bullet> b" + by (simp add: x algebra_simps inner_commute u) + assume "dist c b \<le> dist c x" + then have "(c - b) \<bullet> (c - b) \<le> (c - x) \<bullet> (c - x)" + by (simp add: dist_norm norm_le) + then have "(b \<bullet> b) * (1 - u*u) \<le> 2 * (b \<bullet> c) * (1-u)" + by (simp add: x algebra_simps inner_commute) + then have "(1+u) * (b \<bullet> b) * (1-u) \<le> 2 * (b \<bullet> c) * (1-u)" + by (simp add: algebra_simps) + then have "(1+u) * (b \<bullet> b) \<le> 2 * (b \<bullet> c)" + using \<open>u < 1\<close> by auto + with xb have "c \<bullet> b \<ge> x \<bullet> 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 \<in> open_segment a b" + shows "dist c x < dist c a \<or> dist c x < dist c b" +proof - + have *: "x - a \<in> 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 + +lemma dist_decreases_closed_segment: + fixes a :: "'a :: euclidean_space" + assumes "x \<in> closed_segment a b" + shows "dist c x \<le> dist c a \<or> dist c x \<le> dist c b" +apply (cases "x \<in> 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) + +lemma convex_intermediate_ball: + fixes a :: "'a :: euclidean_space" + shows "\<lbrakk>ball a r \<subseteq> T; T \<subseteq> cball a r\<rbrakk> \<Longrightarrow> 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 \<subseteq> 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 real_sum_of_halves scaleR_left.add) + +lemma notin_segment_midpoint: + fixes a :: "'a :: euclidean_space" + shows "a \<noteq> b \<Longrightarrow> a \<notin> closed_segment (midpoint a b) b" +by (auto simp: dist_midpoint dest!: dist_in_closed_segment) + +subsubsection\<open>More lemmas, especially for working with the underlying formula\<close> + +lemma segment_eq_compose: + fixes a :: "'a :: real_vector" + shows "(\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\<lambda>x. a + x) o (\<lambda>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 \<longleftrightarrow> a=b \<or> 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 \<or> 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 \<longleftrightarrow> a=b \<or> u=0" + using segment_degen_1 [of "1-u" b a] + by (auto simp: algebra_simps) + +lemma closed_segment_image_interval: + "closed_segment a b = (\<lambda>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 (\<lambda>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 open_segment_bound1: + assumes "x \<in> 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 \<noteq> 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 \<in> 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]: + fixes a :: "'a::euclidean_space" + shows "closure(open_segment a b) = (if a = b then {} else closed_segment a b)" +proof - + have "closure ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\<lambda>u. u *\<^sub>R (b - a)) ` closure {0<..<1}" if "a \<noteq> b" + apply (rule closure_injective_linear_image [symmetric]) + apply (simp add:) + using that by (simp add: inj_on_def) + then show ?thesis + by (simp add: segment_image_interval segment_eq_compose closure_greaterThanLessThan [symmetric] + closure_translation image_comp [symmetric] del: closure_greaterThanLessThan) +qed + +lemma closed_open_segment_iff [simp]: + fixes a :: "'a::euclidean_space" shows "closed(open_segment a b) \<longleftrightarrow> 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) \<longleftrightarrow> 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 ((\<lambda>u. u *\<^sub>R (b-a)) ` {0<..<1})" + by (rule convex_linear_image) auto + then show ?thesis + apply (simp add: open_segment_image_interval segment_eq_compose) + by (metis image_comp convex_translation) +qed + +lemmas convex_segment = convex_closed_segment convex_open_segment + +lemma connected_segment [iff]: + fixes x :: "'a :: real_normed_vector" + shows "connected (closed_segment x y)" + by (simp add: convex_connected) + +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 \<in> rel_interior S" "b \<in> closure S" + shows "open_segment a b \<subseteq> 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 \<in> open_segment a b" + then show "x \<in> rel_interior S" + unfolding closed_segment_def open_segment_def using assms + by (auto intro: rel_interior_closure_convex_shrink) +qed + +subsection\<open>More results about segments\<close> + +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} \<inter> 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 *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> + dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) = + (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1)" for x + proof - + have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> + dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) = + ((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and> + dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a))" + unfolding eq_diff_eq [symmetric] by simp + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + norm ((a+b) - (2 *\<^sub>R x)) \<le> norm (b - a))" + by (simp add: dist_half_times2) (simp add: dist_norm) + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) \<le> norm (b - a))" + by auto + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + norm ((1 - u * 2) *\<^sub>R (b - a)) \<le> norm (b - a))" + by (simp add: algebra_simps scaleR_2) + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + \<bar>1 - u * 2\<bar> * norm (b - a) \<le> norm (b - a))" + by simp + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> \<le> 1)" + by (simp add: mult_le_cancel_right2 False) + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 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} \<inter> 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 *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> + dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = + (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> u < 1)" for x + proof - + have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> + dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = + ((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and> + dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a))" + unfolding eq_diff_eq [symmetric] by simp + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + norm ((a+b) - (2 *\<^sub>R x)) < norm (b - a))" + by (simp add: dist_half_times2) (simp add: dist_norm) + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) < norm (b - a))" + by auto + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + norm ((1 - u * 2) *\<^sub>R (b - a)) < norm (b - a))" + by (simp add: algebra_simps scaleR_2) + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> + \<bar>1 - u * 2\<bar> * norm (b - a) < norm (b - a))" + by simp + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> < 1)" + by (simp add: mult_le_cancel_right2 False) + also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> 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 closed_segment_neq_empty [simp]: "closed_segment a b \<noteq> {}" + by auto + +lemma open_segment_eq_empty [simp]: "open_segment a b = {} \<longleftrightarrow> a = b" +proof - + { assume a1: "open_segment a b = {}" + have "{} \<noteq> {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 \<longleftrightarrow> 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 \<noteq> b" + shows "inj_on (\<lambda>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) \<longleftrightarrow> 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) \<longleftrightarrow> 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} \<longleftrightarrow> a = c \<and> b = c" + by auto + +lemma open_segment_eq_sing: "open_segment a b \<noteq> {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 subset_closed_segment: + "closed_segment a b \<subseteq> closed_segment c d \<longleftrightarrow> + a \<in> closed_segment c d \<and> b \<in> closed_segment c d" + by auto (meson contra_subsetD convex_closed_segment convex_contains_segment) + +lemma subset_co_segment: + "closed_segment a b \<subseteq> open_segment c d \<longleftrightarrow> + a \<in> open_segment c d \<and> b \<in> open_segment c d" +using closed_segment_subset by blast + +lemma subset_open_segment: + fixes a :: "'a::euclidean_space" + shows "open_segment a b \<subseteq> open_segment c d \<longleftrightarrow> + a = b \<or> a \<in> closed_segment c d \<and> b \<in> 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 \<open>a \<noteq> b\<close> have "c \<noteq> d" + using closed_segment_idem singleton_iff by auto + have "\<exists>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 \<and> 0 < uc \<and> uc < 1" + if neq: "(1 - ua) *\<^sub>R c + ua *\<^sub>R d \<noteq> (1 - ub) *\<^sub>R c + ub *\<^sub>R d" "c \<noteq> 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 \<le> ua" "ua \<le> 1" "0 \<le> ub" "ub \<le> 1" + for u ua ub + proof - + have "ua \<noteq> ub" + using neq by auto + moreover have "(u - 1) * ua \<le> 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 "\<not> p \<le> 0" "\<not> q \<le> 0" + using p q not_less by blast+ + then show ?thesis + by (metis \<open>ua \<noteq> ub\<close> 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 \<open>ua \<noteq> ub\<close> + 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 \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close> show ?lhs + unfolding open_segment_image_interval closed_segment_def + by (fastforce simp add:) + next + assume lhs: ?lhs + with \<open>a \<noteq> b\<close> have "c \<noteq> d" + by (meson finite_open_segment rev_finite_subset) + have "closure (open_segment a b) \<subseteq> closure (open_segment c d)" + using lhs closure_mono by blast + then have "closed_segment a b \<subseteq> closed_segment c d" + by (simp add: \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close>) + then show ?rhs + by (force simp: \<open>a \<noteq> b\<close>) + qed +qed + +lemma subset_oc_segment: + fixes a :: "'a::euclidean_space" + shows "open_segment a b \<subseteq> closed_segment c d \<longleftrightarrow> + a = b \<or> a \<in> closed_segment c d \<and> b \<in> 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 + + +subsection\<open>Betweenness\<close> + +lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b" + unfolding between_def by auto + +lemma between: "between (a, b) (x::'a::euclidean_space) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)" +proof (cases "a = b") + case True + then show ?thesis + unfolding between_def split_conv + by (auto simp add: dist_commute) +next + case False + then have Fal: "norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0" + by auto + have *: "\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" + by (auto simp add: algebra_simps) + show ?thesis + unfolding between_def split_conv closed_segment_def mem_Collect_eq + apply rule + apply (elim exE conjE) + apply (subst dist_triangle_eq) + proof - + fix u + assume as: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" + then have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" + unfolding as(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 as(2,3) + by (auto simp add: field_simps) + next + assume as: "dist a b = dist a x + dist x b" + have "norm (a - x) / norm (a - b) \<le> 1" + using Fal2 unfolding as[unfolded dist_norm] norm_ge_zero by auto + then show "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" + apply (rule_tac x="dist a x / dist a b" in exI) + unfolding dist_norm + apply (subst euclidean_eq_iff) + apply rule + defer + apply rule + prefer 3 + apply rule + proof - + fix i :: 'a + assume i: "i \<in> Basis" + have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i = + ((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)" + using Fal by (auto simp add: field_simps inner_simps) + also have "\<dots> = x\<bullet>i" + apply (rule divide_eq_imp[OF Fal]) + unfolding as[unfolded dist_norm] + using as[unfolded dist_triangle_eq] + apply - + apply (subst (asm) euclidean_eq_iff) + using i + apply (erule_tac x=i in ballE) + apply (auto simp add: field_simps inner_simps) + done + finally show "x \<bullet> i = + ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i" + by auto + qed (insert Fal2, auto) + qed +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 *: "\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y" + by auto + show ?t1 ?t2 + unfolding between midpoint_def dist_norm + apply(rule_tac[!] *) + unfolding euclidean_eq_iff[where 'a='a] + apply (auto simp add: field_simps inner_simps) + done +qed + +lemma between_mem_convex_hull: + "between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}" + unfolding between_mem_segment segment_convex_hull .. + + +subsection \<open>Shrinking towards the interior of a convex set\<close> + +lemma mem_interior_convex_shrink: + fixes s :: "'a::euclidean_space set" + assumes "convex s" + and "c \<in> interior s" + and "x \<in> s" + and "0 < e" + and "e \<le> 1" + shows "x - e *\<^sub>R (x - c) \<in> interior s" +proof - + obtain d where "d > 0" and d: "ball c d \<subseteq> s" + using assms(2) unfolding mem_interior by auto + show ?thesis + unfolding mem_interior + apply (rule_tac x="e*d" in exI) + apply rule + defer + unfolding subset_eq Ball_def mem_ball + proof (rule, rule) + fix y + assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" + have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" + using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) + have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \<bar>1/e\<bar> * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" + unfolding dist_norm + unfolding norm_scaleR[symmetric] + apply (rule arg_cong[where f=norm]) + using \<open>e > 0\<close> + by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) + also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" + by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) + also have "\<dots> < d" + using as[unfolded dist_norm] and \<open>e > 0\<close> + by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) + finally show "y \<in> s" + apply (subst *) + apply (rule assms(1)[unfolded convex_alt,rule_format]) + apply (rule d[unfolded subset_eq,rule_format]) + unfolding mem_ball + using assms(3-5) + apply auto + done + qed (insert \<open>e>0\<close> \<open>d>0\<close>, auto) +qed + +lemma mem_interior_closure_convex_shrink: + fixes s :: "'a::euclidean_space set" + assumes "convex s" + and "c \<in> interior s" + and "x \<in> closure s" + and "0 < e" + and "e \<le> 1" + shows "x - e *\<^sub>R (x - c) \<in> interior s" +proof - + obtain d where "d > 0" and d: "ball c d \<subseteq> s" + using assms(2) unfolding mem_interior by auto + have "\<exists>y\<in>s. norm (y - x) * (1 - e) < e * d" + proof (cases "x \<in> s") + case True + then show ?thesis + using \<open>e > 0\<close> \<open>d > 0\<close> + apply (rule_tac bexI[where x=x]) + apply (auto) + done + 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 \<in> s" "y \<noteq> x" "dist y x < 1" + using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto + then show ?thesis + apply (rule_tac x=y in bexI) + unfolding True + using \<open>d > 0\<close> + apply auto + done + next + case False + then have "0 < e * d / (1 - e)" and *: "1 - e > 0" + using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto + then obtain y where "y \<in> s" "y \<noteq> x" "dist y x < e * d / (1 - e)" + using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto + then show ?thesis + apply (rule_tac x=y in bexI) + unfolding dist_norm + using pos_less_divide_eq[OF *] + apply auto + done + qed + qed + then obtain y where "y \<in> 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 \<open>e > 0\<close> + by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) + have "z \<in> interior s" + apply (rule interior_mono[OF d,unfolded subset_eq,rule_format]) + unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5) + apply (auto simp add:field_simps norm_minus_commute) + done + then show ?thesis + unfolding * + apply - + apply (rule mem_interior_convex_shrink) + using assms(1,4-5) \<open>y\<in>s\<close> + apply auto + done +qed + +lemma in_interior_closure_convex_segment: + fixes S :: "'a::euclidean_space set" + assumes "convex S" and a: "a \<in> interior S" and b: "b \<in> closure S" + shows "open_segment a b \<subseteq> 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 "... \<in> interior S" using mem_interior_closure_convex_shrink [OF assms] u + by simp + finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \<in> interior S" . +qed + + +subsection \<open>Some obvious but surprisingly hard simplex lemmas\<close> + +lemma simplex: + assumes "finite s" + and "0 \<notin> s" + shows "convex hull (insert 0 s) = + {y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s \<le> 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y)}" + unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]] + apply (rule set_eqI, rule) + unfolding mem_Collect_eq + apply (erule_tac[!] exE) + apply (erule_tac[!] conjE)+ + unfolding setsum_clauses(2)[OF \<open>finite s\<close>] + apply (rule_tac x=u in exI) + defer + apply (rule_tac x="\<lambda>x. if x = 0 then 1 - setsum u s else u x" in exI) + using assms(2) + unfolding if_smult and setsum_delta_notmem[OF assms(2)] + apply auto + done + +lemma substd_simplex: + assumes d: "d \<subseteq> Basis" + shows "convex hull (insert 0 d) = + {x. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}" + (is "convex hull (insert 0 ?p) = ?s") +proof - + let ?D = d + have "0 \<notin> ?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 \<open>finite d\<close> \<open>0 \<notin> ?p\<close>] + apply (rule set_eqI) + unfolding mem_Collect_eq + apply rule + apply (elim exE conjE) + apply (erule_tac[2] conjE)+ + proof - + fix x :: "'a::euclidean_space" + fix u + assume as: "\<forall>x\<in>?D. 0 \<le> u x" "setsum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x" + have *: "\<forall>i\<in>Basis. i:d \<longrightarrow> u i = x\<bullet>i" + and "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)" + using as(3) + unfolding substdbasis_expansion_unique[OF assms] + by auto + then have **: "setsum u ?D = setsum (op \<bullet> x) ?D" + apply - + apply (rule setsum.cong) + using assms + apply auto + done + have "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1" + proof (rule,rule) + fix i :: 'a + assume i: "i \<in> Basis" + have "i \<in> d \<Longrightarrow> 0 \<le> x\<bullet>i" + unfolding *[rule_format,OF i,symmetric] + apply (rule_tac as(1)[rule_format]) + apply auto + done + moreover have "i \<notin> d \<Longrightarrow> 0 \<le> x\<bullet>i" + using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close>[rule_format, OF i] by auto + ultimately show "0 \<le> x\<bullet>i" by auto + qed (insert as(2)[unfolded **], auto) + then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)" + using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close> by auto + next + fix x :: "'a::euclidean_space" + assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "setsum (op \<bullet> x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)" + show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> setsum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x" + using as d + unfolding substdbasis_expansion_unique[OF assms] + apply (rule_tac x="inner x" in exI) + apply auto + done + qed +qed + +lemma std_simplex: + "convex hull (insert 0 Basis) = + {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis \<le> 1}" + using substd_simplex[of Basis] by auto + +lemma interior_std_simplex: + "interior (convex hull (insert 0 Basis)) = + {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis < 1}" + apply (rule set_eqI) + unfolding mem_interior std_simplex + unfolding subset_eq mem_Collect_eq Ball_def mem_ball + unfolding Ball_def[symmetric] + apply rule + apply (elim exE conjE) + defer + apply (erule conjE) +proof - + fix x :: 'a + fix e + assume "e > 0" and as: "\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> setsum (op \<bullet> xa) Basis \<le> 1" + show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> setsum (op \<bullet> x) Basis < 1" + apply safe + proof - + fix i :: 'a + assume i: "i \<in> Basis" + then show "0 < x \<bullet> i" + using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close> + unfolding dist_norm + by (auto elim!: ballE[where x=i] simp: inner_simps) + next + have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using \<open>e > 0\<close> + unfolding dist_norm + by (auto intro!: mult_strict_left_mono simp: SOME_Basis) + have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) \<bullet> i = + x\<bullet>i + (if i = (SOME i. i\<in>Basis) then e/2 else 0)" + by (auto simp: SOME_Basis inner_Basis inner_simps) + then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis = + setsum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis" + apply (rule_tac setsum.cong) + apply auto + done + have "setsum (op \<bullet> x) Basis < setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis" + unfolding * setsum.distrib + using \<open>e > 0\<close> DIM_positive[where 'a='a] + apply (subst setsum.delta') + apply (auto simp: SOME_Basis) + done + also have "\<dots> \<le> 1" + using ** + apply (drule_tac as[rule_format]) + apply auto + done + finally show "setsum (op \<bullet> x) Basis < 1" by auto + qed +next + fix x :: 'a + assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "setsum (op \<bullet> x) Basis < 1" + obtain a :: 'b where "a \<in> UNIV" using UNIV_witness .. + let ?d = "(1 - setsum (op \<bullet> x) Basis) / real (DIM('a))" + have "Min ((op \<bullet> x) ` Basis) > 0" + apply (rule Min_grI) + using as(1) + apply auto + done + moreover have "?d > 0" + using as(2) by (auto simp: Suc_le_eq DIM_positive) + ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1" + apply (rule_tac x="min (Min ((op \<bullet> x) ` Basis)) D" for D in exI) + apply rule + defer + apply (rule, rule) + proof - + fix y + assume y: "dist x y < min (Min (op \<bullet> x ` Basis)) ?d" + have "setsum (op \<bullet> y) Basis \<le> setsum (\<lambda>i. x\<bullet>i + ?d) Basis" + proof (rule setsum_mono) + fix i :: 'a + assume i: "i \<in> Basis" + then have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" + apply - + apply (rule le_less_trans) + using Basis_le_norm[OF i, of "y - x"] + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] + apply (auto simp add: norm_minus_commute inner_diff_left) + done + then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto + qed + also have "\<dots> \<le> 1" + unfolding setsum.distrib setsum_constant + by (auto simp add: Suc_le_eq) + finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1" + proof safe + fix i :: 'a + assume i: "i \<in> Basis" + have "norm (x - y) < x\<bullet>i" + apply (rule less_le_trans) + apply (rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) + using i + apply auto + done + then show "0 \<le> y\<bullet>i" + using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] + by (auto simp: inner_simps) + qed + qed auto +qed + +lemma interior_std_simplex_nonempty: + obtains a :: "'a::euclidean_space" where + "a \<in> interior(convex hull (insert 0 Basis))" +proof - + let ?D = "Basis :: 'a set" + let ?a = "setsum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" + { + fix i :: 'a + assume i: "i \<in> Basis" + have "?a \<bullet> i = inverse (2 * real DIM('a))" + by (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) + (simp_all add: setsum.If_cases i) } + note ** = this + show ?thesis + apply (rule that[of ?a]) + unfolding interior_std_simplex mem_Collect_eq + proof safe + fix i :: 'a + assume i: "i \<in> Basis" + show "0 < ?a \<bullet> i" + unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive) + next + have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D" + apply (rule setsum.cong) + apply rule + apply auto + done + also have "\<dots> < 1" + unfolding setsum_constant divide_inverse[symmetric] + by (auto simp add: field_simps) + finally show "setsum (op \<bullet> ?a) ?D < 1" by auto + qed +qed + +lemma rel_interior_substd_simplex: + assumes d: "d \<subseteq> Basis" + shows "rel_interior (convex hull (insert 0 d)) = + {x::'a::euclidean_space. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}" + (is "rel_interior (convex hull (insert 0 ?p)) = ?s") +proof - + have "finite d" + apply (rule finite_subset) + using assms + apply auto + done + 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 ?p)) = + {x::'a::euclidean_space. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}" + using affine_hull_convex_hull affine_hull_substd_basis assms by auto + have aux: "\<And>x::'a. \<forall>i\<in>Basis. (\<forall>i\<in>d. 0 \<le> x\<bullet>i) \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i" + by auto + { + fix x :: "'a::euclidean_space" + assume x: "x \<in> rel_interior (convex hull (insert 0 ?p))" + then obtain e where e0: "e > 0" and + "ball x e \<inter> {xa. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0)} \<subseteq> convex hull (insert 0 ?p)" + using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto + then have as: "\<forall>xa. dist x xa < e \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0) \<longrightarrow> + (\<forall>i\<in>d. 0 \<le> xa \<bullet> i) \<and> setsum (op \<bullet> xa) d \<le> 1" + unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto + have x0: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)" + using x rel_interior_subset substd_simplex[OF assms] by auto + have "(\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)" + apply rule + apply rule + proof - + fix i :: 'a + assume "i \<in> d" + then have "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> ia" + apply - + apply (rule as[rule_format,THEN conjunct1]) + unfolding dist_norm + using d \<open>e > 0\<close> x0 + apply (auto simp: inner_simps inner_Basis) + done + then show "0 < x \<bullet> i" + apply (erule_tac x=i in ballE) + using \<open>e > 0\<close> \<open>i \<in> d\<close> d + apply (auto simp: inner_simps inner_Basis) + done + next + obtain a where a: "a \<in> d" + using \<open>d \<noteq> {}\<close> by auto + then have **: "dist x (x + (e / 2) *\<^sub>R a) < e" + using \<open>e > 0\<close> norm_Basis[of a] d + unfolding dist_norm + by auto + have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)" + using a d by (auto simp: inner_simps inner_Basis) + then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d = + setsum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d" + using d by (intro setsum.cong) auto + have "a \<in> Basis" + using \<open>a \<in> d\<close> d by auto + then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)" + using x0 d \<open>a\<in>d\<close> by (auto simp add: inner_add_left inner_Basis) + have "setsum (op \<bullet> x) d < setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d" + unfolding * setsum.distrib + using \<open>e > 0\<close> \<open>a \<in> d\<close> + using \<open>finite d\<close> + by (auto simp add: setsum.delta') + also have "\<dots> \<le> 1" + using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] + by auto + finally show "setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)" + using x0 by auto + qed + } + moreover + { + fix x :: "'a::euclidean_space" + assume as: "x \<in> ?s" + have "\<forall>i. 0 < x\<bullet>i \<or> 0 = x\<bullet>i \<longrightarrow> 0 \<le> x\<bullet>i" + by auto + moreover have "\<forall>i. i \<in> d \<or> i \<notin> d" by auto + ultimately + have "\<forall>i. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<forall>i. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i" + by metis + then have h2: "x \<in> convex hull (insert 0 ?p)" + using as assms + unfolding substd_simplex[OF assms] by fastforce + obtain a where a: "a \<in> d" + using \<open>d \<noteq> {}\<close> by auto + let ?d = "(1 - setsum (op \<bullet> x) d) / real (card d)" + have "0 < card d" using \<open>d \<noteq> {}\<close> \<open>finite d\<close> + by (simp add: card_gt_0_iff) + have "Min ((op \<bullet> x) ` d) > 0" + using as \<open>d \<noteq> {}\<close> \<open>finite d\<close> by (simp add: Min_grI) + moreover have "?d > 0" using as using \<open>0 < card d\<close> by auto + ultimately have h3: "min (Min ((op \<bullet> x) ` d)) ?d > 0" + by auto + + have "x \<in> rel_interior (convex hull (insert 0 ?p))" + unfolding rel_interior_ball mem_Collect_eq h0 + apply (rule,rule h2) + unfolding substd_simplex[OF assms] + apply (rule_tac x="min (Min ((op \<bullet> x) ` d)) ?d" in exI) + apply (rule, rule h3) + apply safe + unfolding mem_ball + proof - + fix y :: 'a + assume y: "dist x y < min (Min (op \<bullet> x ` d)) ?d" + assume y2: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> y\<bullet>i = 0" + have "setsum (op \<bullet> y) d \<le> setsum (\<lambda>i. x\<bullet>i + ?d) d" + proof (rule setsum_mono) + fix i + assume "i \<in> d" + with d have i: "i \<in> Basis" + by auto + have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" + apply (rule le_less_trans) + using Basis_le_norm[OF i, of "y - x"] + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] + apply (auto simp add: norm_minus_commute inner_simps) + done + then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto + qed + also have "\<dots> \<le> 1" + unfolding setsum.distrib setsum_constant using \<open>0 < card d\<close> + by auto + finally show "setsum (op \<bullet> y) d \<le> 1" . + + fix i :: 'a + assume i: "i \<in> Basis" + then show "0 \<le> y\<bullet>i" + proof (cases "i\<in>d") + case True + have "norm (x - y) < x\<bullet>i" + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] + using Min_gr_iff[of "op \<bullet> x ` d" "norm (x - y)"] \<open>0 < card d\<close> \<open>i:d\<close> + by (simp add: card_gt_0_iff) + then show "0 \<le> y\<bullet>i" + using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] + by (auto simp: inner_simps) + qed (insert y2, auto) + qed + } + ultimately have + "\<And>x. x \<in> rel_interior (convex hull insert 0 d) \<longleftrightarrow> + x \<in> {x. (\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)}" + by blast + then show ?thesis by (rule set_eqI) + qed +qed + +lemma rel_interior_substd_simplex_nonempty: + assumes "d \<noteq> {}" + and "d \<subseteq> Basis" + obtains a :: "'a::euclidean_space" + where "a \<in> rel_interior (convex hull (insert 0 d))" +proof - + let ?D = d + let ?a = "setsum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D" + have "finite d" + apply (rule finite_subset) + using assms(2) + apply auto + done + then have d1: "0 < real (card d)" + using \<open>d \<noteq> {}\<close> by auto + { + fix i + assume "i \<in> d" + have "?a \<bullet> i = inverse (2 * real (card d))" + apply (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"]) + unfolding inner_setsum_left + apply (rule setsum.cong) + using \<open>i \<in> d\<close> \<open>finite d\<close> setsum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"] + d1 assms(2) + by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)]) + } + note ** = this + show ?thesis + apply (rule that[of ?a]) + unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq + proof safe + fix i + assume "i \<in> d" + have "0 < inverse (2 * real (card d))" + using d1 by auto + also have "\<dots> = ?a \<bullet> i" using **[of i] \<open>i \<in> d\<close> + by auto + finally show "0 < ?a \<bullet> i" by auto + next + have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D" + by (rule setsum.cong) (rule refl, rule **) + also have "\<dots> < 1" + unfolding setsum_constant divide_real_def[symmetric] + by (auto simp add: field_simps) + finally show "setsum (op \<bullet> ?a) ?D < 1" by auto + next + fix i + assume "i \<in> Basis" and "i \<notin> d" + have "?a \<in> span d" + proof (rule span_setsum[of d "(\<lambda>b. b /\<^sub>R (2 * real (card d)))" d]) + { + fix x :: "'a::euclidean_space" + assume "x \<in> d" + then have "x \<in> span d" + using span_superset[of _ "d"] by auto + then have "x /\<^sub>R (2 * real (card d)) \<in> span d" + using span_mul[of x "d" "(inverse (real (card d)) / 2)"] by auto + } + then show "\<And>x. x\<in>d \<Longrightarrow> x /\<^sub>R (2 * real (card d)) \<in> span d" + by auto + qed + then show "?a \<bullet> i = 0 " + using \<open>i \<notin> d\<close> unfolding span_substd_basis[OF assms(2)] using \<open>i \<in> Basis\<close> by auto + qed +qed + + +subsection \<open>Relative interior of convex set\<close> + +lemma rel_interior_convex_nonempty_aux: + fixes S :: "'n::euclidean_space set" + assumes "convex S" + and "0 \<in> S" + shows "rel_interior S \<noteq> {}" +proof (cases "S = {0}") + case True + then show ?thesis using rel_interior_sing by auto +next + case False + obtain B where B: "independent B \<and> B \<le> S \<and> S \<le> span B \<and> card B = dim S" + using basis_exists[of S] by auto + then have "B \<noteq> {}" + using B assms \<open>S \<noteq> {0}\<close> span_empty by auto + have "insert 0 B \<le> span B" + using subspace_span[of B] subspace_0[of "span B"] span_inc by auto + then have "span (insert 0 B) \<le> span B" + using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast + then have "convex hull insert 0 B \<le> span B" + using convex_hull_subset_span[of "insert 0 B"] by auto + then have "span (convex hull insert 0 B) \<le> 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 \<in> 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 \<Rightarrow> 'n" where + fd: "card d = card B" "linear f" "f ` B = d" + "f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = (0::real)} \<and> inj_on f (span B)" + and d: "d \<subseteq> 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 \<noteq> {}" + using fd B \<open>B \<noteq> {}\<close> 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)"] \<open>linear f\<close> + by auto + moreover have "rel_interior (f ` (convex hull insert 0 B)) = + f ` rel_interior (convex hull insert 0 B)" + apply (rule rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"]) + using \<open>bounded_linear f\<close> fd * + apply auto + done + ultimately have "rel_interior (convex hull insert 0 B) \<noteq> {}" + using rel_interior_substd_simplex_nonempty[OF \<open>d \<noteq> {}\<close> d] + apply auto + apply blast + done + moreover have "convex hull (insert 0 B) \<subseteq> 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 = {} \<longleftrightarrow> S = {}" +proof - + { + assume "S \<noteq> {}" + then obtain a where "a \<in> S" by auto + then have "0 \<in> op + (-a) ` S" + using assms exI[of "(\<lambda>x. x \<in> S \<and> - a + x = 0)" a] by auto + then have "rel_interior (op + (-a) ` S) \<noteq> {}" + using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"] + convex_translation[of S "-a"] assms + by auto + then have "rel_interior S \<noteq> {}" + using rel_interior_translation by auto + } + then show ?thesis + using rel_interior_empty 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 \<in> rel_interior S" "y \<in> rel_interior S" "0 \<le> u" "u \<le> 1" + then have "x \<in> S" + using rel_interior_subset by auto + have "x - u *\<^sub>R (x-y) \<in> 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 \<open>x \<in> S\<close> by auto + next + case True + then show ?thesis using assm by auto + qed + then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> 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) \<le> 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 \<in> rel_interior S" + using rel_interior_eq_empty assms by auto + { fix x + assume x: "x \<in> closure S" + { + assume "x = a" + then have "x \<in> closure (rel_interior S)" + using a unfolding closure_def by auto + } + moreover + { + assume "x \<noteq> a" + { + fix e :: real + assume "e > 0" + define e1 where "e1 = min 1 (e/norm (x - a))" + then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (x - a) \<le> e" + using \<open>x \<noteq> a\<close> \<open>e > 0\<close> 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 "\<exists>y. y \<in> rel_interior S \<and> y \<noteq> x \<and> dist y x \<le> e" + apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI) + using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \<open>x \<noteq> a\<close> + apply simp + done + } + then have "x islimpt rel_interior S" + unfolding islimpt_approachable_le by auto + then have "x \<in> closure(rel_interior S)" + unfolding closure_def by auto + } + ultimately have "x \<in> closure(rel_interior S)" by auto + } + then show ?thesis using h1 by auto + next + case True + then have "rel_interior S = {}" + using rel_interior_empty by auto + then have "closure (rel_interior S) = {}" + using closure_empty by auto + with True show ?thesis by auto + qed +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 (subtopology euclidean (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 \<le> 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)" + apply auto + using assms + apply simp + done + also have "\<dots> = (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 "\<dots> = y - e *\<^sub>R (y-x)" + using e_def + apply (simp add: algebra_simps) + using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"] + apply auto + done + finally have "z = y - e *\<^sub>R (y-x)" + by auto + moreover have "e > 0" using e_def assms by auto + moreover have "e \<le> 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) \<supseteq> 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 \<in> rel_interior (closure S)" + obtain x where x: "x \<in> rel_interior S" + using \<open>S \<noteq> {}\<close> assms rel_interior_eq_empty by auto + have "z \<in> 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 \<inter> affine hull closure S \<le> 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 \<in> cball z e" + using mem_cball y_def dist_norm[of z y] e by auto + have "x \<in> affine hull closure S" + using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast + moreover have "z \<in> affine hull closure S" + using z rel_interior_subset hull_subset[of "closure S"] by blast + ultimately have "y \<in> 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 \<in> 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 \<le> 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 \<open>y \<in> closure S\<close> + by auto + 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 \<longleftrightarrow> 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 \<longleftrightarrow> rel_interior S1 \<le> S2 \<and> S2 \<subseteq> closure S1" + (is "?A \<longleftrightarrow> ?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 \<subseteq> closure S2" + by (metis assms(1) convex_closure_rel_interior closure_mono) + moreover from \<open>?B\<close> have "closure S1 \<supseteq> 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 \<inter> closure S = {} \<longleftrightarrow> A \<inter> 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 show ?thesis + apply (simp add: rel_interior_eq openin_open) + apply (rule_tac x="ball (inverse 2 *\<^sub>R (a + b)) (norm(b - a) / 2)" in exI) + apply (simp add: open_segment_as_ball) + done +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 + +lemma starlike_convex_tweak_boundary_points: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "S \<noteq> {}" and ST: "rel_interior S \<subseteq> T" and TS: "T \<subseteq> closure S" + shows "starlike T" +proof - + have "rel_interior S \<noteq> {}" + by (simp add: assms rel_interior_eq_empty) + then obtain a where a: "a \<in> rel_interior S" by blast + with ST have "a \<in> T" by blast + have *: "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S" + apply (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) + using assms by blast + show ?thesis + unfolding starlike_def + apply (rule bexI [OF _ \<open>a \<in> T\<close>]) + apply (simp add: closed_segment_eq_open) + apply (intro conjI ballI a \<open>a \<in> T\<close> rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) + apply (simp add: order_trans [OF * ST]) + done +qed + +subsection\<open>The relative frontier of a set\<close> + +definition "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 = {} \<longleftrightarrow> affine S" + apply (simp add: rel_frontier_def) + apply (simp add: rel_interior_eq_closure [symmetric]) + using rel_interior_subset_closure by blast + +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 \<subseteq> 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 + apply simp + by (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((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (rel_frontier S)" +by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) + +lemma closed_affine_hull [iff]: + fixes S :: "'n::euclidean_space set" + shows "closed (affine hull S)" + by (metis affine_affine_hull affine_closed) + +lemma rel_frontier_nonempty_interior: + fixes S :: "'n::euclidean_space set" + shows "interior S \<noteq> {} \<Longrightarrow> 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 \<Longrightarrow> rel_frontier S = frontier S" +by (simp add: frontier_def rel_frontier_def rel_interior_interior) + +lemma closed_rel_frontier [iff]: + fixes S :: "'n::euclidean_space set" + shows "closed (rel_frontier S)" +proof - + have *: "closedin (subtopology euclidean (affine hull S)) (closure S - rel_interior S)" + apply (rule closedin_diff[of "subtopology euclidean (affine hull S)""closure S" "rel_interior S"]) + using closed_closedin_trans[of "affine hull S" "closure S"] closed_affine_hull[of S] + closure_affine_hull[of S] openin_rel_interior[of S] + apply auto + done + show ?thesis + apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) + unfolding rel_frontier_def + using * closed_affine_hull + apply auto + done +qed + +lemma closed_rel_boundary: + fixes S :: "'n::euclidean_space set" + shows "closed S \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> 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 \<Longrightarrow> 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 "\<lbrakk>convex S; convex T\<rbrakk> + \<Longrightarrow> rel_interior S = rel_interior T \<longleftrightarrow> 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 "\<lbrakk>convex S; convex T\<rbrakk> + \<Longrightarrow> rel_interior S = rel_interior T \<longleftrightarrow> + rel_interior S \<subseteq> T \<and> T \<subseteq> 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 \<noteq> {}" + and "S1 \<le> rel_frontier S2" + shows "aff_dim S1 < aff_dim S2" +proof - + have "S1 \<subseteq> closure S2" + using assms unfolding rel_frontier_def by auto + then have *: "affine hull S1 \<subseteq> affine hull S2" + using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast + then have "aff_dim S1 \<le> 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 \<noteq> {}" + using aff_dim_empty[of S1] aff_dim_empty[of S2] \<open>S2 \<noteq> {}\<close> by auto + have **: "affine hull S1 = affine hull S2" + apply (rule affine_dim_equal) + using * affine_affine_hull + apply auto + using \<open>S1 \<noteq> {}\<close> hull_subset[of S1] + apply auto + using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] + apply auto + done + obtain a where a: "a \<in> rel_interior S1" + using \<open>S1 \<noteq> {}\<close> rel_interior_eq_empty assms by auto + obtain T where T: "open T" "a \<in> T \<inter> S1" "T \<inter> affine hull S1 \<subseteq> S1" + using mem_rel_interior[of a S1] a by auto + then have "a \<in> T \<inter> closure S2" + using a assms unfolding rel_frontier_def by auto + then obtain b where b: "b \<in> T \<inter> rel_interior S2" + using open_inter_closure_rel_interior[of S2 T] assms T by auto + then have "b \<in> affine hull S1" + using rel_interior_subset hull_subset[of S2] ** by auto + then have "b \<in> 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 \<in> rel_interior S" + shows "\<forall>x\<in>affine hull S. \<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" +proof - + obtain e1 where e1: "e1 > 0 \<and> cball z e1 \<inter> affine hull S \<subseteq> S" + using mem_rel_interior_cball[of z S] assms by auto + { + fix x + assume x: "x \<in> affine hull S" + { + assume "x \<noteq> z" + define m where "m = 1 + e1/norm(x-z)" + hence "m > 1" using e1 \<open>x \<noteq> z\<close> by auto + { + fix e + assume e: "e > 1 \<and> e \<le> m" + have "z \<in> 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 \<in> 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 "\<dots> = (e - 1) * norm (x-z)" + using norm_scaleR e by auto + also have "\<dots> \<le> (m - 1) * norm (x - z)" + using e mult_right_mono[of _ _ "norm(x-z)"] by auto + also have "\<dots> = (e1 / norm (x - z)) * norm (x - z)" + using m_def by auto + also have "\<dots> = e1" + using \<open>x \<noteq> z\<close> e1 by simp + finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \<le> e1" + by auto + have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \<in> 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 \<in> S" + using e * e1 by auto + } + then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )" + using \<open>m> 1 \<close> 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 \<and> e \<le> m" + then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" + using e1 x \<open>x = z\<close> by (auto simp add: algebra_simps) + then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" + using e by auto + } + then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" + using \<open>m > 1\<close> by auto + } + ultimately have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> 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 \<in> rel_interior S" + shows "\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> 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 \<noteq> {}" + assumes "\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" + shows "z \<in> rel_interior S" +proof - + obtain x where x: "x \<in> rel_interior S" + using rel_interior_eq_empty assms by auto + then have "x \<in> S" + using rel_interior_subset by auto + then obtain e where e: "e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" + using assms by auto + define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" + then have "y \<in> S" using e by auto + define e1 where "e1 = 1/e" + then have "0 < e1 \<and> 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"] \<open>0 < e1 \<and> e1 < 1\<close> \<open>y \<in> S\<close> x assms + by auto +qed + +lemma convex_rel_interior_iff: + fixes S :: "'n::euclidean_space set" + assumes "convex S" + and "S \<noteq> {}" + shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> 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 \<noteq> {}" + shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> 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 \<in> interior S \<longleftrightarrow> (\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S)" +proof (cases "aff_dim S = int DIM('n)") + case False + { + assume "z \<in> interior S" + then have False + using False interior_rel_interior_gen[of S] by auto + } + moreover + { + assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S" + { + fix x + obtain e1 where e1: "e1 > 0 \<and> z + e1 *\<^sub>R (x - z) \<in> S" + using r by auto + obtain e2 where e2: "e2 > 0 \<and> z + e2 *\<^sub>R (z - x) \<in> S" + using r by auto + define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" + then have x1: "x1 \<in> 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 \<in> 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" + using x1_def x2_def + apply (auto simp add: algebra_simps) + using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] + apply auto + done + then have z: "z \<in> 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 \<in> 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 add: aff_dim_UNIV) + then have False + using False by auto + } + ultimately show ?thesis by auto +next + case True + then have "S \<noteq> {}" + using aff_dim_empty[of S] by auto + have *: "affine hull S = UNIV" + using True affine_hull_UNIV by auto + { + assume "z \<in> interior S" + then have "z \<in> rel_interior S" + using True interior_rel_interior_gen[of S] by auto + then have **: "\<forall>x. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" + using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> * by auto + fix x + obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \<in> 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 \<in> S" + using e1 e_def by auto + then have "\<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S" + by auto + } + moreover + { + assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S" + { + fix x + obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \<in> 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 \<in> S" + using e1 e_def by auto + then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" by auto + } + then have "z \<in> rel_interior S" + using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> by auto + then have "z \<in> interior S" + using True interior_rel_interior_gen[of S] by auto + } + ultimately show ?thesis by auto +qed + + +subsubsection \<open>Relative interior and closure under common operations\<close> + +lemma rel_interior_inter_aux: "\<Inter>{rel_interior S |S. S : I} \<subseteq> \<Inter>I" +proof - + { + fix y + assume "y \<in> \<Inter>{rel_interior S |S. S : I}" + then have y: "\<forall>S \<in> I. y \<in> rel_interior S" + by auto + { + fix S + assume "S \<in> I" + then have "y \<in> S" + using rel_interior_subset y by auto + } + then have "y \<in> \<Inter>I" by auto + } + then show ?thesis by auto +qed + +lemma closure_Int: "closure (\<Inter>I) \<le> \<Inter>{closure S |S. S \<in> I}" +proof - + { + fix y + assume "y \<in> \<Inter>I" + then have y: "\<forall>S \<in> I. y \<in> S" by auto + { + fix S + assume "S \<in> I" + then have "y \<in> closure S" + using closure_subset y by auto + } + then have "y \<in> \<Inter>{closure S |S. S \<in> I}" + by auto + } + then have "\<Inter>I \<subseteq> \<Inter>{closure S |S. S \<in> I}" + by auto + moreover have "closed (\<Inter>{closure S |S. S \<in> I})" + unfolding closed_Inter closed_closure by auto + ultimately show ?thesis using closure_hull[of "\<Inter>I"] + hull_minimal[of "\<Inter>I" "\<Inter>{closure S |S. S \<in> I}" "closed"] by auto +qed + +lemma convex_closure_rel_interior_inter: + assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" + and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}" + shows "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})" +proof - + obtain x where x: "\<forall>S\<in>I. x \<in> rel_interior S" + using assms by auto + { + fix y + assume "y \<in> \<Inter>{closure S |S. S \<in> I}" + then have y: "\<forall>S \<in> I. y \<in> closure S" + by auto + { + assume "y = x" + then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})" + using x closure_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto + } + moreover + { + assume "y \<noteq> x" + { fix e :: real + assume e: "e > 0" + define e1 where "e1 = min 1 (e/norm (y - x))" + then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (y - x) \<le> e" + using \<open>y \<noteq> x\<close> \<open>e > 0\<close> 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 \<in> I" + then have "z \<in> 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 \<in> \<Inter>{rel_interior S |S. S \<in> I}" + by auto + have "\<exists>z. z \<in> \<Inter>{rel_interior S |S. S \<in> I} \<and> z \<noteq> y \<and> dist z y \<le> e" + apply (rule_tac x="z" in exI) + using \<open>y \<noteq> x\<close> z_def * e1 e dist_norm[of z y] + apply simp + done + } + then have "y islimpt \<Inter>{rel_interior S |S. S \<in> I}" + unfolding islimpt_approachable_le by blast + then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})" + unfolding closure_def by auto + } + ultimately have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})" + by auto + } + then show ?thesis by auto +qed + +lemma convex_closure_inter: + assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" + and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}" + shows "closure (\<Inter>I) = \<Inter>{closure S |S. S \<in> I}" +proof - + have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})" + using convex_closure_rel_interior_inter assms by auto + moreover + have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)" + using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"] + by auto + ultimately show ?thesis + using closure_Int[of I] by auto +qed + +lemma convex_inter_rel_interior_same_closure: + assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" + and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}" + shows "closure (\<Inter>{rel_interior S |S. S \<in> I}) = closure (\<Inter>I)" +proof - + have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})" + using convex_closure_rel_interior_inter assms by auto + moreover + have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)" + using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"] + by auto + ultimately show ?thesis + using closure_Int[of I] by auto +qed + +lemma convex_rel_interior_inter: + assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" + and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}" + shows "rel_interior (\<Inter>I) \<subseteq> \<Inter>{rel_interior S |S. S \<in> I}" +proof - + have "convex (\<Inter>I)" + using assms convex_Inter by auto + moreover + have "convex (\<Inter>{rel_interior S |S. S \<in> I})" + apply (rule convex_Inter) + using assms convex_rel_interior + apply auto + done + ultimately + have "rel_interior (\<Inter>{rel_interior S |S. S \<in> I}) = rel_interior (\<Inter>I)" + using convex_inter_rel_interior_same_closure assms + closure_eq_rel_interior_eq[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"] + by blast + then show ?thesis + using rel_interior_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto +qed + +lemma convex_rel_interior_finite_inter: + assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" + and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}" + and "finite I" + shows "rel_interior (\<Inter>I) = \<Inter>{rel_interior S |S. S \<in> I}" +proof - + have "\<Inter>I \<noteq> {}" + using assms rel_interior_inter_aux[of I] by auto + have "convex (\<Inter>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 \<in> \<Inter>{rel_interior S |S. S \<in> I}" + { + fix x + assume x: "x \<in> \<Inter>I" + { + fix S + assume S: "S \<in> I" + then have "z \<in> rel_interior S" "x \<in> S" + using z x by auto + then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S)" + using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto + } + then obtain mS where + mS: "\<forall>S\<in>I. mS S > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> mS S \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" by metis + define e where "e = Min (mS ` I)" + then have "e \<in> mS ` I" using assms \<open>I \<noteq> {}\<close> by simp + then have "e > 1" using mS by auto + moreover have "\<forall>S\<in>I. e \<le> mS S" + using e_def assms by auto + ultimately have "\<exists>e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> \<Inter>I" + using mS by auto + } + then have "z \<in> rel_interior (\<Inter>I)" + using convex_rel_interior_iff[of "\<Inter>I" z] \<open>\<Inter>I \<noteq> {}\<close> \<open>convex (\<Inter>I)\<close> 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 \<inter> rel_interior T \<noteq> {}" + shows "closure (S \<inter> T) = closure S \<inter> 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 \<inter> rel_interior T \<noteq> {}" + shows "rel_interior (S \<inter> T) = rel_interior S \<inter> 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 \<inter> T \<noteq> {}" + shows "closure (S \<inter> T) = closure S \<inter> 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 \<longleftrightarrow> closed_segment a b \<subseteq> 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 \<longleftrightarrow> closed_segment a b \<subseteq> 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 \<inter> T \<noteq> {}" + shows "rel_interior (S \<inter> T) = rel_interior S \<inter> 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 \<inter> T \<noteq> {}" + shows "rel_frontier(S \<inter> T) = frontier S \<inter> T" +using assms +apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def) +by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) + +lemma subset_rel_interior_convex: + fixes S T :: "'n::euclidean_space set" + assumes "convex S" + and "convex T" + and "S \<le> closure T" + and "\<not> S \<subseteq> rel_frontier T" + shows "rel_interior S \<subseteq> rel_interior T" +proof - + have *: "S \<inter> closure T = S" + using assms by auto + have "\<not> rel_interior S \<subseteq> 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 \<inter> rel_interior (closure T) \<noteq> {}" + using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] + by auto + then have "rel_interior S \<inter> rel_interior T = rel_interior (S \<inter> 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 "\<dots> = rel_interior S" + using * by auto + finally show ?thesis + by auto +qed + +lemma rel_interior_convex_linear_image: + fixes f :: "'m::euclidean_space \<Rightarrow> '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 rel_interior_empty rel_interior_eq_empty by auto +next + case False + have *: "f ` (rel_interior S) \<subseteq> f ` S" + unfolding image_mono using rel_interior_subset by auto + have "f ` S \<subseteq> f ` (closure S)" + unfolding image_mono using closure_subset by auto + also have "\<dots> = f ` (closure (rel_interior S))" + using convex_closure_rel_interior assms by auto + also have "\<dots> \<subseteq> 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) \<subseteq> f ` rel_interior S" + using rel_interior_subset by auto + moreover + { + fix z + assume "z \<in> f ` rel_interior S" + then obtain z1 where z1: "z1 \<in> rel_interior S" "f z1 = z" by auto + { + fix x + assume "x \<in> f ` S" + then obtain x1 where x1: "x1 \<in> 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] \<open>convex S\<close> 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 \<open>linear f\<close> by (simp add: linear_add_cmul) + 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 "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S" + using e by auto + } + then have "z \<in> rel_interior (f ` S)" + using convex_rel_interior_iff[of "f ` S" z] \<open>convex S\<close> + \<open>linear f\<close> \<open>S \<noteq> {}\<close> 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 \<Rightarrow> 'n::euclidean_space" + assumes "linear f" + and "convex S" + and "f -` (rel_interior S) \<noteq> {}" + shows "rel_interior (f -` S) = f -` (rel_interior S)" +proof - + have "S \<noteq> {}" + using assms rel_interior_empty by auto + have nonemp: "f -` S \<noteq> {}" + by (metis assms(3) rel_interior_subset subset_empty vimage_mono) + then have "S \<inter> (range f) \<noteq> {}" + by auto + have conv: "convex (f -` S)" + using convex_linear_vimage assms by auto + then have "convex (S \<inter> range f)" + by (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image) + { + fix z + assume "z \<in> f -` (rel_interior S)" + then have z: "f z : rel_interior S" + by auto + { + fix x + assume "x \<in> f -` S" + then have "f x \<in> S" by auto + then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \<in> S" + using convex_rel_interior_iff[of S "f z"] z assms \<open>S \<noteq> {}\<close> 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 \<open>linear f\<close> by (simp add: linear_iff) + ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> f -` S" + using e by auto + } + then have "z \<in> rel_interior (f -` S)" + using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto + } + moreover + { + fix z + assume z: "z \<in> rel_interior (f -` S)" + { + fix x + assume "x \<in> S \<inter> range f" + then obtain y where y: "f y = x" "y \<in> f -` S" by auto + then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \<in> 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 \<open>linear f\<close> y by (simp add: linear_iff) + ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R f z \<in> S \<inter> range f" + using e by auto + } + then have "f z \<in> rel_interior (S \<inter> range f)" + using \<open>convex (S \<inter> (range f))\<close> \<open>S \<inter> range f \<noteq> {}\<close> + convex_rel_interior_iff[of "S \<inter> (range f)" "f z"] + by auto + moreover have "affine (range f)" + by (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image) + ultimately have "f z \<in> rel_interior S" + using convex_affine_rel_interior_Int[of S "range f"] assms by auto + then have "z \<in> 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 \<times> T) = rel_interior S \<times> rel_interior T" +proof - + { assume "S = {}" + then have ?thesis + by auto + } + moreover + { assume "T = {}" + then have ?thesis + by auto + } + moreover + { + assume "S \<noteq> {}" "T \<noteq> {}" + then have ri: "rel_interior S \<noteq> {}" "rel_interior T \<noteq> {}" + using rel_interior_eq_empty assms by auto + then have "fst -` rel_interior S \<noteq> {}" + using fst_vimage_eq_Times[of "rel_interior S"] by auto + then have "rel_interior ((fst :: 'n * 'm \<Rightarrow> 'n) -` S) = fst -` rel_interior S" + using fst_linear \<open>convex S\<close> rel_interior_convex_linear_preimage[of fst S] by auto + then have s: "rel_interior (S \<times> (UNIV :: 'm set)) = rel_interior S \<times> UNIV" + by (simp add: fst_vimage_eq_Times) + from ri have "snd -` rel_interior T \<noteq> {}" + using snd_vimage_eq_Times[of "rel_interior T"] by auto + then have "rel_interior ((snd :: 'n * 'm \<Rightarrow> 'm) -` T) = snd -` rel_interior T" + using snd_linear \<open>convex T\<close> rel_interior_convex_linear_preimage[of snd T] by auto + then have t: "rel_interior ((UNIV :: 'n set) \<times> T) = UNIV \<times> rel_interior T" + by (simp add: snd_vimage_eq_Times) + from s t have *: "rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T) = + rel_interior S \<times> rel_interior T" by auto + have "S \<times> T = S \<times> (UNIV :: 'm set) \<inter> (UNIV :: 'n set) \<times> T" + by auto + then have "rel_interior (S \<times> T) = rel_interior ((S \<times> (UNIV :: 'm set)) \<inter> ((UNIV :: 'n set) \<times> T))" + by auto + also have "\<dots> = rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T)" + apply (subst convex_rel_interior_inter_two[of "S \<times> (UNIV :: 'm set)" "(UNIV :: 'n set) \<times> T"]) + using * ri assms convex_Times + apply auto + done + finally have ?thesis using * by auto + } + ultimately show ?thesis by blast +qed + +lemma rel_interior_scaleR: + fixes S :: "'n::euclidean_space set" + assumes "c \<noteq> 0" + shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)" + using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S] + linear_conv_bounded_linear[of "op *\<^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 "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^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 ((op *\<^sub>R c) ` S) \<and> rel_open ((op *\<^sub>R c) ` S)" + by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) + +lemma convex_rel_open_finite_inter: + assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set) \<and> rel_open S" + and "finite I" + shows "convex (\<Inter>I) \<and> rel_open (\<Inter>I)" +proof (cases "\<Inter>{rel_interior S |S. S \<in> I} = {}") + case True + then have "\<Inter>I = {}" + using assms unfolding rel_open_def by auto + then show ?thesis + unfolding rel_open_def using rel_interior_empty by auto +next + case False + then have "rel_open (\<Inter>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 \<Rightarrow> 'n::euclidean_space" + assumes "linear f" + and "convex S" + and "rel_open S" + shows "convex (f ` S) \<and> 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 \<Rightarrow> 'n::euclidean_space" + assumes "linear f" + and "convex S" + and "rel_open S" + shows "convex (f -` S) \<and> 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 using rel_interior_empty 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 \<times> 'n::euclidean_space) set" + and f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space set" + assumes "convex S" + and "f = (\<lambda>y. {z. (y, z) \<in> S})" + shows "(y, z) \<in> rel_interior S \<longleftrightarrow> (y \<in> rel_interior {y. (f y \<noteq> {})} \<and> z \<in> rel_interior (f y))" +proof - + { + fix y + assume "y \<in> {y. f y \<noteq> {}}" + then obtain z where "(y, z) \<in> S" + using assms by auto + then have "\<exists>x. x \<in> S \<and> y = fst x" + apply (rule_tac x="(y, z)" in exI) + apply auto + done + then obtain x where "x \<in> S" "y = fst x" + by blast + then have "y \<in> fst ` S" + unfolding image_def by auto + } + then have "fst ` S = {y. f y \<noteq> {}}" + unfolding fst_def using assms by auto + then have h1: "fst ` rel_interior S = rel_interior {y. f y \<noteq> {}}" + using rel_interior_convex_linear_image[of fst S] assms fst_linear by auto + { + fix y + assume "y \<in> rel_interior {y. f y \<noteq> {}}" + then have "y \<in> fst ` rel_interior S" + using h1 by auto + then have *: "rel_interior S \<inter> fst -` {y} \<noteq> {}" + by auto + moreover have aff: "affine (fst -` {y})" + unfolding affine_alt by (simp add: algebra_simps) + ultimately have **: "rel_interior (S \<inter> fst -` {y}) = rel_interior S \<inter> fst -` {y}" + using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto + have conv: "convex (S \<inter> fst -` {y})" + using convex_Int assms aff affine_imp_convex by auto + { + fix x + assume "x \<in> f y" + then have "(y, x) \<in> S \<inter> (fst -` {y})" + using assms by auto + moreover have "x = snd (y, x)" by auto + ultimately have "x \<in> snd ` (S \<inter> fst -` {y})" + by blast + } + then have "snd ` (S \<inter> fst -` {y}) = f y" + using assms by auto + then have ***: "rel_interior (f y) = snd ` rel_interior (S \<inter> fst -` {y})" + using rel_interior_convex_linear_image[of snd "S \<inter> fst -` {y}"] snd_linear conv + by auto + { + fix z + assume "z \<in> rel_interior (f y)" + then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})" + using *** by auto + moreover have "{y} = fst ` rel_interior (S \<inter> fst -` {y})" + using * ** rel_interior_subset by auto + ultimately have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})" + by force + then have "(y,z) \<in> rel_interior S" + using ** by auto + } + moreover + { + fix z + assume "(y, z) \<in> rel_interior S" + then have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})" + using ** by auto + then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})" + by (metis Range_iff snd_eq_Range) + then have "z \<in> rel_interior (f y)" + using *** by auto + } + ultimately have "\<And>z. (y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)" + by auto + } + then have h2: "\<And>y z. y \<in> rel_interior {t. f t \<noteq> {}} \<Longrightarrow> + (y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)" + by auto + { + fix y z + assume asm: "(y, z) \<in> rel_interior S" + then have "y \<in> fst ` rel_interior S" + by (metis Domain_iff fst_eq_Domain) + then have "y \<in> rel_interior {t. f t \<noteq> {}}" + using h1 by auto + then have "y \<in> rel_interior {t. f t \<noteq> {}}" 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 \<times> rel_frontier T \<subseteq> rel_frontier (S \<times> T)" + by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) + + +subsubsection \<open>Relative interior of convex cone\<close> + +lemma cone_rel_interior: + fixes S :: "'m::euclidean_space set" + assumes "cone S" + shows "cone ({0} \<union> rel_interior S)" +proof (cases "S = {}") + case True + then show ?thesis + by (simp add: rel_interior_empty cone_0) +next + case False + then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" + using cone_iff[of S] assms by auto + then have *: "0 \<in> ({0} \<union> rel_interior S)" + and "\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` ({0} \<union> rel_interior S) = ({0} \<union> rel_interior S)" + by (auto simp add: rel_interior_scaleR) + then show ?thesis + using cone_iff[of "{0} \<union> rel_interior S"] by auto +qed + +lemma rel_interior_convex_cone_aux: + fixes S :: "'m::euclidean_space set" + assumes "convex S" + shows "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) \<longleftrightarrow> + c > 0 \<and> x \<in> ((op *\<^sub>R c) ` (rel_interior S))" +proof (cases "S = {}") + case True + then show ?thesis + by (simp add: rel_interior_empty cone_hull_empty) +next + case False + then obtain s where "s \<in> S" by auto + have conv: "convex ({(1 :: real)} \<times> S)" + using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] + by auto + define f where "f y = {z. (y, z) \<in> cone hull ({1 :: real} \<times> S)}" for y + then have *: "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) = + (c \<in> rel_interior {y. f y \<noteq> {}} \<and> x \<in> rel_interior (f c))" + apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \<times> S)" f c x]) + using convex_cone_hull[of "{(1 :: real)} \<times> S"] conv + apply auto + done + { + fix y :: real + assume "y \<ge> 0" + then have "y *\<^sub>R (1,s) \<in> cone hull ({1 :: real} \<times> S)" + using cone_hull_expl[of "{(1 :: real)} \<times> S"] \<open>s \<in> S\<close> by auto + then have "f y \<noteq> {}" + using f_def by auto + } + then have "{y. f y \<noteq> {}} = {0..}" + using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto + then have **: "rel_interior {y. f y \<noteq> {}} = {0<..}" + using rel_interior_real_semiline by auto + { + fix c :: real + assume "c > 0" + then have "f c = (op *\<^sub>R c ` S)" + using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto + then have "rel_interior (f c) = op *\<^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} \<times> S)) = + {(c, c *\<^sub>R x) | c x. c > 0 \<and> x \<in> rel_interior S}" + (is "?lhs = ?rhs") +proof - + { + fix z + assume "z \<in> ?lhs" + have *: "z = (fst z, snd z)" + by auto + have "z \<in> ?rhs" + using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \<open>z \<in> ?lhs\<close> + apply auto + apply (rule_tac x = "fst z" in exI) + apply (rule_tac x = x in exI) + using * + apply auto + done + } + moreover + { + fix z + assume "z \<in> ?rhs" + then have "z \<in> ?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 "\<forall>i\<in>I. convex (S i) \<and> (S i) \<noteq> {}" + shows "convex hull (\<Union>(S ` I)) = + {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)}" + (is "?lhs = ?rhs") +proof - + have "?lhs \<supseteq> ?rhs" + proof + fix x + assume "x : ?rhs" + then obtain c s where *: "setsum (\<lambda>i. c i *\<^sub>R s i) I = x" "setsum c I = 1" + "(\<forall>i\<in>I. c i \<ge> 0) \<and> (\<forall>i\<in>I. s i \<in> S i)" by auto + then have "\<forall>i\<in>I. s i \<in> convex hull (\<Union>(S ` I))" + using hull_subset[of "\<Union>(S ` I)" convex] by auto + then show "x \<in> ?lhs" + unfolding *(1)[symmetric] + apply (subst convex_setsum[of I "convex hull \<Union>(S ` I)" c s]) + using * assms convex_convex_hull + apply auto + done + qed + + { + fix i + assume "i \<in> I" + with assms have "\<exists>p. p \<in> S i" by auto + } + then obtain p where p: "\<forall>i\<in>I. p i \<in> S i" by metis + + { + fix i + assume "i \<in> I" + { + fix x + assume "x \<in> S i" + define c where "c j = (if j = i then 1::real else 0)" for j + then have *: "setsum c I = 1" + using \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. 1::real"] + by auto + define s where "s j = (if j = i then x else p j)" for j + then have "\<forall>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 = setsum (\<lambda>i. c i *\<^sub>R s i) I" + using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. x"] + by auto + then have "x \<in> ?rhs" + apply auto + apply (rule_tac x = c in exI) + apply (rule_tac x = s in exI) + using * c_def s_def p \<open>x \<in> S i\<close> + apply auto + done + } + then have "?rhs \<supseteq> S i" by auto + } + then have *: "?rhs \<supseteq> \<Union>(S ` I)" by auto + + { + fix u v :: real + assume uv: "u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1" + fix x y + assume xy: "x \<in> ?rhs \<and> y \<in> ?rhs" + from xy obtain c s where + xc: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)" + by auto + from xy obtain d t where + yc: "y = setsum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> setsum d I = 1 \<and> (\<forall>i\<in>I. t i \<in> S i)" + by auto + define e where "e i = u * c i + v * d i" for i + have ge0: "\<forall>i\<in>I. e i \<ge> 0" + using e_def xc yc uv by simp + have "setsum (\<lambda>i. u * c i) I = u * setsum c I" + by (simp add: setsum_right_distrib) + moreover have "setsum (\<lambda>i. v * d i) I = v * setsum d I" + by (simp add: setsum_right_distrib) + ultimately have sum1: "setsum e I = 1" + using e_def xc yc uv by (simp add: setsum.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 \<in> I" + have "q i \<in> 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: "\<forall>i\<in>I. q i \<in> S i" by auto + { + fix i + assume i: "i \<in> 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) \<ge> 0 \<and> v * d i \<ge> 0" + using xc yc uv i by simp + moreover from ge have "u * c i \<le> 0 \<and> v * d i \<le> 0" + using True e_def i by simp + ultimately have "u * c i = 0 \<and> 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 *: "\<forall>i\<in>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 = setsum (\<lambda>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.setsum setsum.distrib) + also have "\<dots> = setsum (\<lambda>i. e i *\<^sub>R q i) I" + using * by auto + finally have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (e i) *\<^sub>R (q i)) I" + by auto + then have "u *\<^sub>R x + v *\<^sub>R y \<in> ?rhs" + using ge0 sum1 qs by auto + } + then have "convex ?rhs" unfolding convex_def by auto + then show ?thesis + using \<open>?lhs \<supseteq> ?rhs\<close> * hull_minimal[of "\<Union>(S ` I)" ?rhs convex] + by blast +qed + +lemma convex_hull_union_two: + fixes S T :: "'m::euclidean_space set" + assumes "convex S" + and "S \<noteq> {}" + and "convex T" + and "T \<noteq> {}" + shows "convex hull (S \<union> T) = + {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> 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 "\<Union>(s ` I) = S \<union> T" + using s_def I_def by auto + then have "convex hull (\<Union>(s ` I)) = convex hull (S \<union> T)" + by auto + moreover have "convex hull \<Union>(s ` I) = + {\<Sum> i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)}" + apply (subst convex_hull_finite_union[of I s]) + using assms s_def I_def + apply auto + done + moreover have + "{\<Sum>i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)} \<le> ?rhs" + using s_def I_def by auto + ultimately show "?lhs \<subseteq> ?rhs" by auto + { + fix x + assume "x \<in> ?rhs" + then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \<and> u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T" + by auto + then have "x \<in> convex hull {s, t}" + using convex_hull_2[of s t] by auto + then have "x \<in> convex hull (S \<union> T)" + using * hull_mono[of "{s, t}" "S \<union> T"] by auto + } + then show "?lhs \<supseteq> ?rhs" by blast +qed + + +subsection \<open>Convexity on direct sums\<close> + +lemma closure_sum: + fixes S T :: "'a::real_normed_vector set" + shows "closure S + closure T \<subseteq> 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 = (\<lambda>(x,y). x + y) ` (rel_interior S \<times> rel_interior T)" + by (simp add: set_plus_image) + also have "\<dots> = (\<lambda>(x,y). x + y) ` rel_interior (S \<times> T)" + using rel_interior_Times assms by auto + also have "\<dots> = rel_interior (S + T)" + using fst_snd_linear convex_Times assms + rel_interior_convex_linear_image[of "(\<lambda>(x,y). x + y)" "S \<times> T"] + by (auto simp add: set_plus_image) + finally show ?thesis .. +qed + +lemma rel_interior_sum_gen: + fixes S :: "'a \<Rightarrow> 'n::euclidean_space set" + assumes "\<forall>i\<in>I. convex (S i)" + shows "rel_interior (setsum S I) = setsum (\<lambda>i. rel_interior (S i)) I" + apply (subst setsum_set_cond_linear[of convex]) + using rel_interior_sum rel_interior_sing[of "0"] assms + apply (auto simp add: convex_set_plus) + done + +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 \<times> T) \<and> rel_open (S \<times> 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) \<and> 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 \<noteq> {}" + assumes "\<forall>i\<in>I. convex (S i) \<and> cone (S i) \<and> S i \<noteq> {}" + shows "convex hull (\<Union>(S ` I)) = setsum S I" + (is "?lhs = ?rhs") +proof - + { + fix x + assume "x \<in> ?lhs" + then obtain c xs where + x: "x = setsum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. xs i \<in> 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 + { + fix i + assume "i \<in> I" + then have "s i \<in> S i" + using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto + } + then have "\<forall>i\<in>I. s i \<in> S i" by auto + moreover have "x = setsum s I" using x s_def by auto + ultimately have "x \<in> ?rhs" + using set_setsum_alt[of I S] assms by auto + } + moreover + { + fix x + assume "x \<in> ?rhs" + then obtain s where x: "x = setsum s I \<and> (\<forall>i\<in>I. s i \<in> S i)" + using set_setsum_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 = setsum (\<lambda>i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" + using x assms by auto + moreover have "\<forall>i\<in>I. xs i \<in> S i" + using x xs_def assms by (simp add: cone_def) + moreover have "\<forall>i\<in>I. (1 :: real) / of_nat (card I) \<ge> 0" + by auto + moreover have "setsum (\<lambda>i. (1 :: real) / of_nat (card I)) I = 1" + using assms by auto + ultimately have "x \<in> ?lhs" + apply (subst convex_hull_finite_union[of I S]) + using assms + apply blast + using assms + apply blast + apply rule + apply (rule_tac x = "(\<lambda>i. (1 :: real) / of_nat (card I))" in exI) + apply auto + done + } + 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 \<noteq> {}" + assumes "convex T" + and "cone T" + and "T \<noteq> {}" + shows "convex hull (S \<union> 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 "\<Union>(A ` I) = S \<union> T" + using A_def I_def by auto + then have "convex hull (\<Union>(A ` I)) = convex hull (S \<union> T)" + by auto + moreover have "convex hull \<Union>(A ` I) = setsum A I" + apply (subst convex_hull_finite_union_cones[of I A]) + using assms A_def I_def + apply auto + done + moreover have "setsum A I = S + T" + using A_def I_def + unfolding set_plus_def + apply auto + unfolding set_plus_def + apply auto + done + ultimately show ?thesis by auto +qed + +lemma rel_interior_convex_hull_union: + fixes S :: "'a \<Rightarrow> 'n::euclidean_space set" + assumes "finite I" + and "\<forall>i\<in>I. convex (S i) \<and> S i \<noteq> {}" + shows "rel_interior (convex hull (\<Union>(S ` I))) = + {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and> + (\<forall>i\<in>I. s i \<in> rel_interior(S i))}" + (is "?lhs = ?rhs") +proof (cases "I = {}") + case True + then show ?thesis + using convex_hull_empty rel_interior_empty by auto +next + case False + define C0 where "C0 = convex hull (\<Union>(S ` I))" + have "\<forall>i\<in>I. C0 \<ge> S i" + unfolding C0_def using hull_subset[of "\<Union>(S ` I)"] by auto + define K0 where "K0 = cone hull ({1 :: real} \<times> C0)" + define K where "K i = cone hull ({1 :: real} \<times> S i)" for i + have "\<forall>i\<in>I. K i \<noteq> {}" + unfolding K_def using assms + by (simp add: cone_hull_empty_iff[symmetric]) + { + fix i + assume "i \<in> I" + then have "convex (K i)" + unfolding K_def + apply (subst convex_cone_hull) + apply (subst convex_Times) + using assms + apply auto + done + } + then have convK: "\<forall>i\<in>I. convex (K i)" + by auto + { + fix i + assume "i \<in> I" + then have "K0 \<supseteq> K i" + unfolding K0_def K_def + apply (subst hull_mono) + using \<open>\<forall>i\<in>I. C0 \<ge> S i\<close> + apply auto + done + } + then have "K0 \<supseteq> \<Union>(K ` I)" by auto + moreover have "convex K0" + unfolding K0_def + apply (subst convex_cone_hull) + apply (subst convex_Times) + unfolding C0_def + using convex_convex_hull + apply auto + done + ultimately have geq: "K0 \<supseteq> convex hull (\<Union>(K ` I))" + using hull_minimal[of _ "K0" "convex"] by blast + have "\<forall>i\<in>I. K i \<supseteq> {1 :: real} \<times> S i" + using K_def by (simp add: hull_subset) + then have "\<Union>(K ` I) \<supseteq> {1 :: real} \<times> \<Union>(S ` I)" + by auto + then have "convex hull \<Union>(K ` I) \<supseteq> convex hull ({1 :: real} \<times> \<Union>(S ` I))" + by (simp add: hull_mono) + then have "convex hull \<Union>(K ` I) \<supseteq> {1 :: real} \<times> C0" + unfolding C0_def + using convex_hull_Times[of "{(1 :: real)}" "\<Union>(S ` I)"] convex_hull_singleton + by auto + moreover have "cone (convex hull (\<Union>(K ` I)))" + apply (subst cone_convex_hull) + using cone_Union[of "K ` I"] + apply auto + unfolding K_def + using cone_cone_hull + apply auto + done + ultimately have "convex hull (\<Union>(K ` I)) \<supseteq> K0" + unfolding K0_def + using hull_minimal[of _ "convex hull (\<Union>(K ` I))" "cone"] + by blast + then have "K0 = convex hull (\<Union>(K ` I))" + using geq by auto + also have "\<dots> = setsum K I" + apply (subst convex_hull_finite_union_cones[of I K]) + using assms + apply blast + using False + apply blast + unfolding K_def + apply rule + apply (subst convex_cone_hull) + apply (subst convex_Times) + using assms cone_cone_hull \<open>\<forall>i\<in>I. K i \<noteq> {}\<close> K_def + apply auto + done + finally have "K0 = setsum K I" by auto + then have *: "rel_interior K0 = setsum (\<lambda>i. (rel_interior (K i))) I" + using rel_interior_sum_gen[of I K] convK by auto + { + fix x + assume "x \<in> ?lhs" + then have "(1::real, x) \<in> 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) = setsum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))" + using \<open>finite I\<close> * set_setsum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto + { + fix i + assume "i \<in> I" + then have "convex (S i) \<and> k i \<in> rel_interior (cone hull {1} \<times> S i)" + using k K_def assms by auto + then have "\<exists>ci si. k i = (ci, ci *\<^sub>R si) \<and> 0 < ci \<and> si \<in> rel_interior (S i)" + using rel_interior_convex_cone[of "S i"] by auto + } + then obtain c s where + cs: "\<forall>i\<in>I. k i = (c i, c i *\<^sub>R s i) \<and> 0 < c i \<and> s i \<in> rel_interior (S i)" + by metis + then have "x = (\<Sum>i\<in>I. c i *\<^sub>R s i) \<and> setsum c I = 1" + using k by (simp add: setsum_prod) + then have "x \<in> ?rhs" + using k + apply auto + apply (rule_tac x = c in exI) + apply (rule_tac x = s in exI) + using cs + apply auto + done + } + moreover + { + fix x + assume "x \<in> ?rhs" + then obtain c s where cs: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and> + (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> 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 \<in> rel_interior (K i)" + using k_def K_def assms cs rel_interior_convex_cone[of "S i"] + by auto + } + then have "(1::real, x) \<in> rel_interior K0" + using K0_def * set_setsum_alt[of I "(\<lambda>i. rel_interior (K i))"] assms k_def cs + apply auto + apply (rule_tac x = k in exI) + apply (simp add: setsum_prod) + done + then have "x \<in> ?lhs" + using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] + by (auto simp add: convex_convex_hull) + } + ultimately show ?thesis by blast +qed + + +lemma convex_le_Inf_differential: + fixes f :: "real \<Rightarrow> real" + assumes "convex_on I f" + and "x \<in> interior I" + and "y \<in> I" + shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)" + (is "_ \<ge> _ + Inf (?F x) * (y - x)") +proof (cases rule: linorder_cases) + assume "x < y" + moreover + have "open (interior I)" by auto + from openE[OF this \<open>x \<in> interior I\<close>] + obtain e where e: "0 < e" "ball x e \<subseteq> interior I" . + moreover define t where "t = min (x + e / 2) ((x + y) / 2)" + ultimately have "x < t" "t < y" "t \<in> ball x e" + by (auto simp: dist_real_def field_simps split: split_min) + with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto + + have "open (interior I)" by auto + from openE[OF this \<open>x \<in> interior I\<close>] + obtain e where "0 < e" "ball x e \<subseteq> interior I" . + moreover define K where "K = x - e / 2" + with \<open>0 < e\<close> have "K \<in> ball x e" "K < x" + by (auto simp: dist_real_def) + ultimately have "K \<in> I" "K < x" "x \<in> I" + using interior_subset[of I] \<open>x \<in> interior I\<close> by auto + + have "Inf (?F x) \<le> (f x - f y) / (x - y)" + proof (intro bdd_belowI cInf_lower2) + show "(f x - f t) / (x - t) \<in> ?F x" + using \<open>t \<in> I\<close> \<open>x < t\<close> by auto + show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" + using \<open>convex_on I f\<close> \<open>x \<in> I\<close> \<open>y \<in> I\<close> \<open>x < t\<close> \<open>t < y\<close> + by (rule convex_on_diff) + next + fix y + assume "y \<in> ?F x" + with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>K \<in> I\<close> _ \<open>K < x\<close> _]] + show "(f K - f x) / (K - x) \<le> y" by auto + qed + then show ?thesis + using \<open>x < y\<close> by (simp add: field_simps) +next + assume "y < x" + moreover + have "open (interior I)" by auto + from openE[OF this \<open>x \<in> interior I\<close>] + obtain e where e: "0 < e" "ball x e \<subseteq> interior I" . + moreover define t where "t = x + e / 2" + ultimately have "x < t" "t \<in> ball x e" + by (auto simp: dist_real_def field_simps) + with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto + + have "(f x - f y) / (x - y) \<le> Inf (?F x)" + proof (rule cInf_greatest) + have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" + using \<open>y < x\<close> by (auto simp: field_simps) + also + fix z + assume "z \<in> ?F x" + with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>y \<in> I\<close> _ \<open>y < x\<close>]] + have "(f y - f x) / (y - x) \<le> z" + by auto + finally show "(f x - f y) / (x - y) \<le> z" . + next + have "open (interior I)" by auto + from openE[OF this \<open>x \<in> interior I\<close>] + obtain e where e: "0 < e" "ball x e \<subseteq> interior I" . + then have "x + e / 2 \<in> ball x e" + by (auto simp: dist_real_def) + with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I" + by auto + then show "?F x \<noteq> {}" + by blast + qed + then show ?thesis + using \<open>y < x\<close> by (simp add: field_simps) +qed simp + +subsection\<open>Explicit formulas for interior and relative interior of convex hull\<close> + +lemma interior_atLeastAtMost [simp]: + fixes a::real shows "interior {a..b} = {a<..<b}" + using interior_cbox [of a b] by auto + +lemma interior_atLeastLessThan [simp]: + fixes a::real shows "interior {a..<b} = {a<..<b}" + by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost interior_Int interior_interior interior_real_semiline) + +lemma interior_lessThanAtMost [simp]: + fixes a::real shows "interior {a<..b} = {a<..<b}" + by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost interior_Int + interior_interior interior_real_semiline) + +lemma at_within_closed_interval: + fixes x::real + shows "a < x \<Longrightarrow> x < b \<Longrightarrow> (at x within {a..b}) = at x" + by (metis at_within_interior greaterThanLessThan_iff interior_atLeastAtMost) + +lemma affine_independent_convex_affine_hull: + fixes s :: "'a::euclidean_space set" + assumes "~affine_dependent s" "t \<subseteq> s" + shows "convex hull t = affine hull t \<inter> convex hull s" +proof - + have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto + { fix u v x + assume uv: "setsum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "setsum v s = 1" + "(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>v\<in>t. u v *\<^sub>R v)" "x \<in> t" + then have s: "s = (s - t) \<union> t" \<comment>\<open>split into separate cases\<close> + using assms by auto + have [simp]: "(\<Sum>x\<in>t. v x *\<^sub>R x) + (\<Sum>x\<in>s - t. v x *\<^sub>R x) = (\<Sum>x\<in>t. u x *\<^sub>R x)" + "setsum v t + setsum v (s - t) = 1" + using uv fin s + by (auto simp: setsum.union_disjoint [symmetric] Un_commute) + have "(\<Sum>x\<in>s. if x \<in> t then v x - u x else v x) = 0" + "(\<Sum>x\<in>s. (if x \<in> t then v x - u x else v x) *\<^sub>R x) = 0" + using uv fin + by (subst s, subst setsum.union_disjoint, auto simp: algebra_simps setsum_subtractf)+ + } note [simp] = this + have "convex hull t \<subseteq> affine hull t" + using convex_hull_subset_affine_hull by blast + moreover have "convex hull t \<subseteq> convex hull s" + using assms hull_mono by blast + moreover have "affine hull t \<inter> convex hull s \<subseteq> convex hull t" + using assms + apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit) + apply (drule_tac x=s in spec) + apply (auto simp: fin) + apply (rule_tac x=u in exI) + apply (rename_tac v) + apply (drule_tac x="\<lambda>x. if x \<in> t then v x - u x else v x" in spec) + apply (force)+ + done + 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 \<notin> t" "s = insert a t" + by blast + then have fin: "finite t" using assms + by (metis finite_insert aff_independent_finite) + show ?thesis + using assms t fin + apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) + apply (rule subset_antisym) + apply force + apply (rule Fun.vimage_subsetD) + apply (metis add.commute diff_add_cancel surj_def) + apply (rule card_ge_dim_independent) + apply (auto simp: card_image inj_on_def dim_subset_UNIV) + done +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" + apply (rule affine_independent_span_eq [OF ind]) + apply (rule antisym) + using assms + apply auto + apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite) + done + +lemma empty_interior_affine_hull: + fixes s :: "'a::euclidean_space set" + assumes "finite s" and dim: "card s \<le> DIM ('a)" + shows "interior(affine hull s) = {}" + using assms + apply (induct s rule: finite_induct) + apply (simp_all add: affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation) + apply (rule empty_interior_lowdim) + apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) + apply (metis Suc_le_lessD not_less order_trans card_image_le finite_imageI dim_le_card) + done + +lemma empty_interior_convex_hull: + fixes s :: "'a::euclidean_space set" + assumes "finite s" and dim: "card s \<le> 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 + \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y} + \<subseteq> rel_interior (convex hull s)" + by (force simp add: rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified]) + +lemma explicit_subset_rel_interior_convex_hull_minimal: + fixes s :: "'a::euclidean_space set" + shows "finite s + \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y} + \<subseteq> rel_interior (convex hull s)" + by (force simp add: rel_interior_convex_hull_union [where S="\<lambda>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. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}" + (is "?lhs = ?rhs") +proof + show "?rhs \<le> ?lhs" + by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) +next + show "?lhs \<le> ?rhs" + proof (cases "\<exists>a. s = {a}") + case True then show "?lhs \<le> ?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 \<in> s" "b \<in> s" "a \<noteq> b" + then have s: "s = (s - {a,b}) \<union> {a,b}" \<comment>\<open>split into separate cases\<close> + by auto + have "(\<Sum>x\<in>s. if x = a then - d else if x = b then d else 0) = 0" + "(\<Sum>x\<in>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 setsum.union_disjoint, auto)+ + } note [simp] = this + { fix y + assume y: "y \<in> convex hull s" "y \<notin> ?rhs" + { fix u T a + assume ua: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "\<not> 0 < u a" "a \<in> s" + and yT: "y = (\<Sum>x\<in>s. u x *\<^sub>R x)" "y \<in> T" "open T" + and sb: "T \<inter> affine hull s \<subseteq> {w. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = w}" + have ua0: "u a = 0" + using ua by auto + obtain b where b: "b\<in>s" "a \<noteq> b" + using ua False by auto + obtain e where e: "0 < e" "ball (\<Sum>x\<in>s. u x *\<^sub>R x) e \<subseteq> 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 "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> affine hull s" + using yT y by (metis affine_hull_convex_hull hull_redundant_eq) + then have "(\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \<in> affine hull s" + using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) + then have "y - d *\<^sub>R (a - b) \<in> T \<inter> affine hull s" + using d e yT by auto + then obtain v where "\<forall>x\<in>s. 0 \<le> v x" + "setsum v s = 1" + "(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b)" + using subsetD [OF sb] yT + by auto + then have False + using assms + apply (simp add: affine_dependent_explicit_finite fs) + apply (drule_tac x="\<lambda>x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec) + using ua b d + apply (auto simp: algebra_simps setsum_subtractf setsum.distrib) + done + } note * = this + have "y \<notin> rel_interior (convex hull s)" + using y + apply (simp add: mem_rel_interior affine_hull_convex_hull) + 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 \<le> ?rhs" + by blast + qed +qed + +lemma interior_convex_hull_explicit_minimal: + fixes s :: "'a::euclidean_space set" + shows + "~ affine_dependent s + ==> interior(convex hull s) = + (if card(s) \<le> DIM('a) then {} + else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})" + apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) + apply (rule trans [of _ "rel_interior(convex hull s)"]) + apply (simp add: affine_hull_convex_hull affine_independent_span_gt rel_interior_interior) + by (simp add: rel_interior_convex_hull_explicit) + +lemma interior_convex_hull_explicit: + fixes s :: "'a::euclidean_space set" + assumes "~ affine_dependent s" + shows + "interior(convex hull s) = + (if card(s) \<le> DIM('a) then {} + else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})" +proof - + { fix u :: "'a \<Rightarrow> real" and a + assume "card Basis < card s" and u: "\<And>x. x\<in>s \<Longrightarrow> 0 < u x" "setsum u s = 1" and a: "a \<in> s" + then have cs: "Suc 0 < card s" + by (metis DIM_positive less_trans_Suc) + obtain b where b: "b \<in> s" "a \<noteq> b" + proof (cases "s \<le> {a}") + case True + then show thesis + using cs subset_singletonD by fastforce + next + case False + then show thesis + by (blast intro: that) + qed + have "u a + u b \<le> setsum u {a,b}" + using a b by simp + also have "... \<le> setsum u s" + apply (rule Groups_Big.setsum_mono2) + using a b u + apply (auto simp: less_imp_le aff_independent_finite assms) + done + finally have "u a < 1" + using \<open>b \<in> s\<close> u by fastforce + } note [simp] = this + show ?thesis + using assms + apply (auto simp: interior_convex_hull_explicit_minimal) + apply (rule_tac x=u in exI) + apply (auto simp: not_le) + done +qed + +lemma interior_closed_segment_ge2: + fixes a :: "'a::euclidean_space" + assumes "2 \<le> DIM('a)" + shows "interior(closed_segment a b) = {}" +using assms unfolding segment_convex_hull +proof - + have "card {a, b} \<le> 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 \<le> DIM('a) then {} else open_segment a b)" +proof (simp add: not_le, intro conjI impI) + assume "2 \<le> DIM('a)" + then show "interior (open_segment a b) = {}" + apply (simp add: segment_convex_hull open_segment_def) + apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2) + done +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" + apply simp + apply (rule affine_independent_span_gt) + apply (simp_all add: affine_dependent_def insert_Diff_if) + done + 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 \<le> 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 \<longleftrightarrow> {a,b} = {c,d}" +proof + assume abcd: "closed_segment a b = closed_segment c d" + show "{a,b} = {c,d}" + proof (cases "a=b \<or> c=d") + case True with abcd show ?thesis by force + next + case False + then have neq: "a \<noteq> b \<and> c \<noteq> 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 \<in> {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 \<in> {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 \<noteq> 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 \<noteq> 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 \<longleftrightarrow> a = b \<and> c = d \<or> {a,b} = {c,d}" + (is "?lhs = ?rhs") +proof + assume abcd: ?lhs + show ?rhs + proof (cases "a=b \<or> c=d") + case True with abcd show ?thesis + using finite_open_segment by fastforce + next + case False + then have a2: "a \<noteq> b \<and> c \<noteq> 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\<open>Similar results for closure and (relative or absolute) frontier.\<close> + +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. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (\<exists>x \<in> s. u x = 0) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}" +proof - + have fs: "finite s" + using assms by (simp add: aff_independent_finite) + show ?thesis + apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs) + apply (auto simp: convex_hull_finite fs) + apply (drule_tac x=u in spec) + apply (rule_tac x=u in exI) + apply force + apply (rename_tac v) + apply (rule notE [OF assms]) + apply (simp add: affine_dependent_explicit) + apply (rule_tac x=s in exI) + apply (auto simp: fs) + apply (rule_tac x = "\<lambda>x. u x - v x" in exI) + apply (force simp: setsum_subtractf scaleR_diff_left) + done +qed + +lemma frontier_convex_hull_explicit: + fixes s :: "'a::euclidean_space set" + assumes "~ affine_dependent s" + shows "frontier(convex hull s) = + {y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (DIM ('a) < card s \<longrightarrow> (\<exists>x \<in> s. u x = 0)) \<and> + setsum u s = 1 \<and> setsum (\<lambda>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 \<le> 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) = \<Union>{convex hull (s - {x}) |x. x \<in> s}" +proof - + have fs: "finite s" + using assms by (simp add: aff_independent_finite) + { fix u a + have "\<forall>x\<in>s. 0 \<le> u x \<Longrightarrow> a \<in> s \<Longrightarrow> u a = 0 \<Longrightarrow> setsum u s = 1 \<Longrightarrow> + \<exists>x v. x \<in> s \<and> + (\<forall>x\<in>s - {x}. 0 \<le> v x) \<and> + setsum v (s - {x}) = 1 \<and> (\<Sum>x\<in>s - {x}. v x *\<^sub>R x) = (\<Sum>x\<in>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.setsum_diff1 fs) + done } + moreover + { fix a u + have "a \<in> s \<Longrightarrow> \<forall>x\<in>s - {a}. 0 \<le> u x \<Longrightarrow> setsum u (s - {a}) = 1 \<Longrightarrow> + \<exists>v. (\<forall>x\<in>s. 0 \<le> v x) \<and> + (\<exists>x\<in>s. v x = 0) \<and> setsum v s = 1 \<and> (\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s - {a}. u x *\<^sub>R x)" + apply (rule_tac x="\<lambda>x. if x = a then 0 else u x" in exI) + apply (auto simp: setsum.If_cases Diff_eq if_smult fs) + done } + ultimately show ?thesis + using assms + apply (simp add: rel_frontier_convex_hull_explicit) + apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto) + 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 \<le> 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 \<le> DIM ('a) then convex hull s else \<Union>{convex hull (s - {x}) |x. x \<in> 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 \<le> Suc (DIM ('a))" "x \<in> s" + shows "x \<in> frontier(convex hull s)" +proof (cases "affine_dependent s") + case True + with assms show ?thesis + apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc) + by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty) +next + case False + { assume "card s = Suc (card Basis)" + then have cs: "Suc 0 < card s" + by (simp add: DIM_positive) + with subset_singletonD have "\<exists>y \<in> s. y \<noteq> x" + by (cases "s \<le> {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 \<le> Suc (DIM ('a))" "x \<in> s" + shows "x \<notin> 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) = {} \<longleftrightarrow> affine_dependent s" +proof - + { fix a b + assume ab: "a \<in> interior (convex hull s)" "b \<in> s" "b \<in> 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 have False using ab + by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq) + } then + show ?thesis + using assms + apply auto + apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull) + apply (auto simp: affine_dependent_def) + done +qed + + +subsection \<open>Coplanarity, and collinearity in terms of affine hull\<close> + +definition coplanar where + "coplanar s \<equiv> \<exists>u v w. s \<subseteq> affine hull {u,v,w}" + +lemma collinear_affine_hull: + "collinear s \<longleftrightarrow> (\<exists>u v. s \<subseteq> affine hull {u,v})" +proof (cases "s={}") + case True then show ?thesis + by simp +next + case False + then obtain x where x: "x \<in> s" by auto + { fix u + assume *: "\<And>x y. \<lbrakk>x\<in>s; y\<in>s\<rbrakk> \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R u" + have "\<exists>u v. s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" + apply (rule_tac x=x in exI) + apply (rule_tac x="x+u" in exI, clarify) + apply (erule exE [OF * [OF x]]) + apply (rename_tac c) + apply (rule_tac x="1+c" in exI) + apply (rule_tac x="-c" in exI) + apply (simp add: algebra_simps) + done + } moreover + { fix u v x y + assume *: "s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" + have "x\<in>s \<Longrightarrow> y\<in>s \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R (v-u)" + apply (drule subsetD [OF *])+ + apply simp + apply clarify + apply (rename_tac r1 r2) + apply (rule_tac x="r1-r2" in exI) + apply (simp add: algebra_simps) + apply (metis scaleR_left.add) + done + } 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 subset_continuous_image_segment_1: + fixes f :: "'a::euclidean_space \<Rightarrow> real" + assumes "continuous_on (closed_segment a b) f" + shows "closed_segment (f a) (f b) \<subseteq> 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 collinear_imp_coplanar: + "collinear s ==> coplanar s" +by (metis collinear_affine_hull coplanar_def insert_absorb2) + +lemma collinear_small: + assumes "finite s" "card s \<le> 2" + shows "collinear s" +proof - + have "card s = 0 \<or> card s = 1 \<or> card s = 2" + using assms by linarith + then show ?thesis using assms + using card_eq_SucD + by auto (metis collinear_2 numeral_2_eq_2) +qed + +lemma coplanar_small: + assumes "finite s" "card s \<le> 3" + shows "coplanar s" +proof - + have "card s \<le> 2 \<or> card s = Suc (Suc (Suc 0))" + using assms by linarith + then show ?thesis using assms + apply safe + apply (simp add: collinear_small collinear_imp_coplanar) + apply (safe dest!: card_eq_SucD) + apply (auto simp: coplanar_def) + apply (metis hull_subset insert_subset) + done +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) \<longleftrightarrow> 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) \<longleftrightarrow> 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 \<Rightarrow> 'b::real_normed_vector" + assumes "coplanar s" "linear f" shows "coplanar(f ` s)" +proof - + { fix u v w + assume "s \<subseteq> affine hull {u, v, w}" + then have "f ` s \<subseteq> f ` (affine hull {u, v, w})" + by (simp add: image_mono) + then have "f ` s \<subseteq> 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: "coplanar s \<Longrightarrow> coplanar ((\<lambda>x. a + x) ` s)" + unfolding coplanar_def + apply clarify + apply (rule_tac x="u+a" in exI) + apply (rule_tac x="v+a" in exI) + apply (rule_tac x="w+a" in exI) + using affine_hull_translation [of a "{u,v,w}" for u v w] + apply (force simp: add.commute) + done + +lemma coplanar_translation_eq: "coplanar((\<lambda>x. a + x) ` s) \<longleftrightarrow> coplanar s" + by (metis (no_types) coplanar_translation_imp translation_galois) + +lemma coplanar_linear_image_eq: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s" +proof + assume "coplanar s" + then show "coplanar (f ` s)" + unfolding coplanar_def + using affine_hull_linear_image [of f "{u,v,w}" for u v w] assms + by (meson coplanar_def coplanar_linear_image) +next + obtain g where g: "linear g" "g \<circ> f = id" + using linear_injective_left_inverse [OF assms] + by blast + assume "coplanar (f ` s)" + then obtain u v w where "f ` s \<subseteq> affine hull {u, v, w}" + by (auto simp: coplanar_def) + then have "g ` f ` s \<subseteq> g ` (affine hull {u, v, w})" + by blast + then have "s \<subseteq> g ` (affine hull {u, v, w})" + using g by (simp add: Fun.image_comp) + then show "coplanar s" + unfolding coplanar_def + using affine_hull_linear_image [of g "{u,v,w}" for u v w] \<open>linear g\<close> linear_conv_bounded_linear + by fastforce +qed +(*The HOL Light proof is simply + MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));; +*) + +lemma coplanar_subset: "\<lbrakk>coplanar t; s \<subseteq> t\<rbrakk> \<Longrightarrow> coplanar s" + by (meson coplanar_def order_trans) + +lemma affine_hull_3_imp_collinear: "c \<in> affine hull {a,b} \<Longrightarrow> collinear {a,b,c}" + by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) + +lemma collinear_3_imp_in_affine_hull: "\<lbrakk>collinear {a,b,c}; a \<noteq> b\<rbrakk> \<Longrightarrow> c \<in> affine hull {a,b}" + unfolding collinear_def + apply clarify + apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE) + apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE) + apply (rename_tac y x) + apply (simp add: affine_hull_2) + apply (rule_tac x="1 - x/y" in exI) + apply (simp add: algebra_simps) + done + +lemma collinear_3_affine_hull: + assumes "a \<noteq> b" + shows "collinear {a,b,c} \<longleftrightarrow> c \<in> 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} \<longleftrightarrow> a = b \<or> a = c \<or> b = c \<or> affine_dependent {a,b,c}" +apply (case_tac "a=b", simp) +apply (case_tac "a=c") +apply (simp add: insert_commute) +apply (case_tac "b=c") +apply (simp add: insert_commute) +apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) +apply (metis collinear_3_affine_hull insert_commute)+ +done + +lemma affine_dependent_imp_collinear_3: + "affine_dependent {a,b,c} \<Longrightarrow> collinear{a,b,c}" +by (simp add: collinear_3_eq_affine_dependent) + +lemma collinear_3: "NO_MATCH 0 x \<Longrightarrow> collinear {x,y,z} \<longleftrightarrow> collinear {0, x-y, z-y}" + by (auto simp add: collinear_def) + + +thm affine_hull_nonempty +corollary affine_hull_eq_empty [simp]: "affine hull S = {} \<longleftrightarrow> S = {}" + using affine_hull_nonempty by blast + +lemma affine_hull_2_alt: + fixes a b :: "'a::real_vector" + shows "affine hull {a,b} = range (\<lambda>u. a + u *\<^sub>R (b - a))" +apply (simp add: affine_hull_2, safe) +apply (rule_tac x=v in image_eqI) +apply (simp add: algebra_simps) +apply (metis scaleR_add_left scaleR_one, simp) +apply (rule_tac x="1-u" in exI) +apply (simp add: algebra_simps) +done + +lemma interior_convex_hull_3_minimal: + fixes a :: "'a::euclidean_space" + shows "\<lbrakk>~ collinear{a,b,c}; DIM('a) = 2\<rbrakk> + \<Longrightarrow> interior(convex hull {a,b,c}) = + {v. \<exists>x y z. 0 < x \<and> 0 < y \<and> 0 < z \<and> x + y + z = 1 \<and> + x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" +apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe) +apply (rule_tac x="u a" in exI, simp) +apply (rule_tac x="u b" in exI, simp) +apply (rule_tac x="u c" in exI, simp) +apply (rename_tac uu x y z) +apply (rule_tac x="\<lambda>r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) +apply simp +done + +subsection\<open>The infimum of the distance between two sets\<close> + +definition setdist :: "'a::metric_space set \<Rightarrow> 'a set \<Rightarrow> real" where + "setdist s t \<equiv> + (if s = {} \<or> t = {} then 0 + else Inf {dist x y| x y. x \<in> s \<and> y \<in> 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 \<le> setdist s t" + by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) + +lemma le_setdistI: + assumes "s \<noteq> {}" "t \<noteq> {}" "\<And>x y. \<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> d \<le> dist x y" + shows "d \<le> setdist s t" + using assms + by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) + +lemma setdist_le_dist: "\<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> setdist s t \<le> dist x y" + unfolding setdist_def + by (auto intro!: bdd_belowI [where m=0] cInf_lower) + +lemma le_setdist_iff: + "d \<le> setdist s t \<longleftrightarrow> + (\<forall>x \<in> s. \<forall>y \<in> t. d \<le> dist x y) \<and> (s = {} \<or> t = {} \<longrightarrow> d \<le> 0)" + apply (cases "s = {} \<or> 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 \<noteq> {}" "t \<noteq> {}" + obtains x y where "x \<in> s" "y \<in> 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 \<le> setdist s {a} + setdist {a} t" +proof (cases "s = {} \<or> t = {}") + case True then show ?thesis + using setdist_pos_le by fastforce +next + case False + have "\<And>x. x \<in> s \<Longrightarrow> setdist s t - dist x a \<le> setdist {a} t" + apply (rule le_setdistI, blast) + using False apply (fastforce intro: le_setdistI) + apply (simp add: algebra_simps) + apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist]) + done + then have "setdist s t - setdist {a} t \<le> 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: "\<bar>setdist {x} s - setdist {y} s\<bar> \<le> 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) (\<lambda>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 (\<lambda>y. (setdist {y} s))" + by (metis continuous_at_setdist continuous_at_imp_continuous_on) + +lemma uniformly_continuous_on_setdist: "uniformly_continuous_on t (\<lambda>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: "\<lbrakk>t \<noteq> {}; t \<subseteq> u\<rbrakk> \<Longrightarrow> setdist s u \<le> setdist s t" + apply (cases "s = {} \<or> u = {}", force) + apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono) + done + +lemma setdist_subset_left: "\<lbrakk>s \<noteq> {}; s \<subseteq> t\<rbrakk> \<Longrightarrow> setdist t u \<le> 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 = {} \<or> t = {}") + case True then show ?thesis by force +next + case False + { fix y + assume "y \<in> t" + have "continuous_on (closure s) (\<lambda>a. dist a y)" + by (auto simp: continuous_intros dist_norm) + then have *: "\<And>x. x \<in> closure s \<Longrightarrow> setdist s t \<le> dist x y" + apply (rule continuous_ge_on_closure) + apply assumption + apply (blast intro: setdist_le_dist \<open>y \<in> t\<close> ) + done + } note * = this + show ?thesis + apply (rule antisym) + using False closure_subset apply (blast intro: setdist_subset_left) + using False * + apply (force simp add: closure_eq_empty 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_compact_closed: + fixes s :: "'a::euclidean_space set" + assumes s: "compact s" and t: "closed t" + and "s \<noteq> {}" "t \<noteq> {}" + shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t" +proof - + have "{x - y |x y. x \<in> s \<and> y \<in> t} \<noteq> {}" + using assms by blast + then + have "\<exists>x \<in> s. \<exists>y \<in> t. dist x y \<le> setdist s t" + apply (rule distance_attains_inf [where a=0, OF compact_closed_differences [OF s t]]) + apply (simp add: dist_norm le_setdist_iff) + apply blast + done + then show ?thesis + by (blast intro!: antisym [OF _ setdist_le_dist] ) +qed + +lemma setdist_closed_compact: + fixes s :: "'a::euclidean_space set" + assumes s: "closed s" and t: "compact t" + and "s \<noteq> {}" "t \<noteq> {}" + shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t" + using setdist_compact_closed [OF t s \<open>t \<noteq> {}\<close> \<open>s \<noteq> {}\<close>] + by (metis dist_commute setdist_sym) + +lemma setdist_eq_0I: "\<lbrakk>x \<in> s; x \<in> t\<rbrakk> \<Longrightarrow> setdist s t = 0" + by (metis antisym dist_self setdist_le_dist setdist_pos_le) + +lemma setdist_eq_0_compact_closed: + fixes s :: "'a::euclidean_space set" + assumes s: "compact s" and t: "closed t" + shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}" + apply (cases "s = {} \<or> t = {}", force) + using setdist_compact_closed [OF s t] + apply (force intro: setdist_eq_0I ) + done + +corollary setdist_gt_0_compact_closed: + fixes s :: "'a::euclidean_space set" + assumes s: "compact s" and t: "closed t" + shows "setdist s t > 0 \<longleftrightarrow> (s \<noteq> {} \<and> t \<noteq> {} \<and> s \<inter> t = {})" + using setdist_pos_le [of s t] setdist_eq_0_compact_closed [OF assms] + by linarith + +lemma setdist_eq_0_closed_compact: + fixes s :: "'a::euclidean_space set" + assumes s: "closed s" and t: "compact t" + shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}" + using setdist_eq_0_compact_closed [OF t s] + by (metis Int_commute setdist_sym) + +lemma setdist_eq_0_bounded: + fixes s :: "'a::euclidean_space set" + assumes "bounded s \<or> bounded t" + shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> closure s \<inter> closure t \<noteq> {}" + apply (cases "s = {} \<or> t = {}", force) + using setdist_eq_0_compact_closed [of "closure s" "closure t"] + setdist_eq_0_closed_compact [of "closure s" "closure t"] assms + apply (force simp add: bounded_closure compact_eq_bounded_closed) + done + +lemma setdist_unique: + "\<lbrakk>a \<in> s; b \<in> t; \<And>x y. x \<in> s \<and> y \<in> t ==> dist a b \<le> dist x y\<rbrakk> + \<Longrightarrow> setdist s t = dist a b" + by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) + +lemma setdist_closest_point: + "\<lbrakk>closed s; s \<noteq> {}\<rbrakk> \<Longrightarrow> setdist {a} s = dist a (closest_point s a)" + apply (rule setdist_unique) + using closest_point_le + apply (auto simp: closest_point_in_set) + done + +lemma setdist_eq_0_sing_1: + fixes s :: "'a::euclidean_space set" + shows "setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s" + by (auto simp: setdist_eq_0_bounded) + +lemma setdist_eq_0_sing_2: + fixes s :: "'a::euclidean_space set" + shows "setdist s {x} = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s" + by (auto simp: setdist_eq_0_bounded) + +lemma setdist_neq_0_sing_1: + fixes s :: "'a::euclidean_space set" + shows "\<lbrakk>setdist {x} s = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> s \<noteq> {} \<and> x \<notin> closure s" + by (auto simp: setdist_eq_0_sing_1) + +lemma setdist_neq_0_sing_2: + fixes s :: "'a::euclidean_space set" + shows "\<lbrakk>setdist s {x} = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> s \<noteq> {} \<and> x \<notin> closure s" + by (auto simp: setdist_eq_0_sing_2) + +lemma setdist_sing_in_set: + fixes s :: "'a::euclidean_space set" + shows "x \<in> s \<Longrightarrow> setdist {x} s = 0" + using closure_subset by (auto simp: setdist_eq_0_sing_1) + +lemma setdist_le_sing: "x \<in> s ==> setdist s t \<le> setdist {x} t" + using setdist_subset_left by auto + +lemma setdist_eq_0_closed: + fixes s :: "'a::euclidean_space set" + shows "closed s \<Longrightarrow> (setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> s)" +by (simp add: setdist_eq_0_sing_1) + +lemma setdist_eq_0_closedin: + fixes s :: "'a::euclidean_space set" + shows "\<lbrakk>closedin (subtopology euclidean u) s; x \<in> u\<rbrakk> + \<Longrightarrow> (setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> s)" + by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) + +subsection\<open>Basic lemmas about hyperplanes and halfspaces\<close> + +lemma hyperplane_eq_Ex: + assumes "a \<noteq> 0" obtains x where "a \<bullet> x = b" + by (rule_tac x = "(b / (a \<bullet> a)) *\<^sub>R a" in that) (simp add: assms) + +lemma hyperplane_eq_empty: + "{x. a \<bullet> x = b} = {} \<longleftrightarrow> a = 0 \<and> b \<noteq> 0" + using hyperplane_eq_Ex apply auto[1] + using inner_zero_right by blast + +lemma hyperplane_eq_UNIV: + "{x. a \<bullet> x = b} = UNIV \<longleftrightarrow> a = 0 \<and> b = 0" +proof - + have "UNIV \<subseteq> {x. a \<bullet> x = b} \<Longrightarrow> a = 0 \<and> b = 0" + apply (drule_tac c = "((b+1) / (a \<bullet> a)) *\<^sub>R a" in subsetD) + apply simp_all + by (metis add_cancel_right_right divide_1 zero_neq_one) + then show ?thesis by force +qed + +lemma halfspace_eq_empty_lt: + "{x. a \<bullet> x < b} = {} \<longleftrightarrow> a = 0 \<and> b \<le> 0" +proof - + have "{x. a \<bullet> x < b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b \<le> 0" + apply (rule ccontr) + apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD) + apply force+ + done + then show ?thesis by force +qed + +lemma halfspace_eq_empty_gt: + "{x. a \<bullet> x > b} = {} \<longleftrightarrow> a = 0 \<and> b \<ge> 0" +using halfspace_eq_empty_lt [of "-a" "-b"] +by simp + +lemma halfspace_eq_empty_le: + "{x. a \<bullet> x \<le> b} = {} \<longleftrightarrow> a = 0 \<and> b < 0" +proof - + have "{x. a \<bullet> x \<le> b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b < 0" + apply (rule ccontr) + apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD) + apply force+ + done + then show ?thesis by force +qed + +lemma halfspace_eq_empty_ge: + "{x. a \<bullet> x \<ge> b} = {} \<longleftrightarrow> a = 0 \<and> b > 0" +using halfspace_eq_empty_le [of "-a" "-b"] +by simp + +subsection\<open>Use set distance for an easy proof of separation properties\<close> + +proposition separation_closures: + fixes S :: "'a::euclidean_space set" + assumes "S \<inter> closure T = {}" "T \<inter> closure S = {}" + obtains U V where "U \<inter> V = {}" "open U" "open V" "S \<subseteq> U" "T \<subseteq> V" +proof (cases "S = {} \<or> T = {}") + case True with that show ?thesis by auto +next + case False + define f where "f \<equiv> \<lambda>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} \<inter> {x. f x < 0} = {}" + by auto + show "open {x. 0 < f x}" + by (simp add: open_Collect_less contf continuous_on_const) + show "open {x. f x < 0}" + by (simp add: open_Collect_less contf continuous_on_const) + show "S \<subseteq> {x. 0 < f x}" + apply (clarsimp simp add: f_def setdist_sing_in_set) + using assms + by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) + show "T \<subseteq> {x. f x < 0}" + apply (clarsimp simp add: f_def setdist_sing_in_set) + using assms + by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) + qed +qed + +lemma separation_normal: + fixes S :: "'a::euclidean_space set" + assumes "closed S" "closed T" "S \<inter> T = {}" + obtains U V where "open U" "open V" "S \<subseteq> U" "T \<subseteq> V" "U \<inter> 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 (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \<inter> T = {}" + obtains S' T' where "openin (subtopology euclidean U) S'" + "openin (subtopology euclidean U) T'" + "S \<subseteq> S'" "T \<subseteq> T'" "S' \<inter> T' = {}" +proof (cases "S = {} \<or> T = {}") + case True with that show ?thesis + apply safe + using UT closedin_subset apply blast + using US closedin_subset apply blast + done +next + case False + define f where "f \<equiv> \<lambda>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' = "{x\<in>U. f x > 0}" and T' = "{x\<in>U. f x < 0}" in that) + show "{x \<in> U. 0 < f x} \<inter> {x \<in> U. f x < 0} = {}" + by auto + have "openin (subtopology euclidean U) {x \<in> U. f x \<in> {0<..}}" + apply (rule continuous_openin_preimage [where T=UNIV]) + apply (simp_all add: contf) + using open_greaterThan open_openin by blast + then show "openin (subtopology euclidean U) {x \<in> U. 0 < f x}" by simp + next + have "openin (subtopology euclidean U) {x \<in> U. f x \<in> {..<0}}" + apply (rule continuous_openin_preimage [where T=UNIV]) + apply (simp_all add: contf) + using open_lessThan open_openin by blast + then show "openin (subtopology euclidean U) {x \<in> U. f x < 0}" by simp + next + show "S \<subseteq> {x \<in> U. 0 < f x}" + apply (clarsimp simp add: f_def setdist_sing_in_set) + using assms + by (metis False Int_insert_right closedin_imp_subset empty_not_insert insert_absorb insert_subset linorder_neqE_linordered_idom not_le setdist_eq_0_closedin setdist_pos_le) + show "T \<subseteq> {x \<in> U. f x < 0}" + apply (clarsimp simp add: f_def setdist_sing_in_set) + using assms + by (metis False closedin_subset disjoint_iff_not_equal insert_absorb insert_subset linorder_neqE_linordered_idom not_less setdist_eq_0_closedin setdist_pos_le topspace_euclidean_subtopology) + qed +qed + +lemma separation_normal_compact: + fixes S :: "'a::euclidean_space set" + assumes "compact S" "closed T" "S \<inter> T = {}" + obtains U V where "open U" "compact(closure U)" "open V" "S \<subseteq> U" "T \<subseteq> V" "U \<inter> 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 \<subseteq> ball 0 r" + by (auto dest!: bounded_subset_ballD) + have **: "closed (T \<union> - ball 0 r)" "S \<inter> (T \<union> - ball 0 r) = {}" + using assms r by blast+ + then show ?thesis + apply (rule separation_normal [OF \<open>closed S\<close>]) + apply (rule_tac U=U and V=V in that) + by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl) +qed + +subsection\<open>Proper maps, including projections out of compact sets\<close> + +lemma finite_indexed_bound: + assumes A: "finite A" "\<And>x. x \<in> A \<Longrightarrow> \<exists>n::'a::linorder. P x n" + shows "\<exists>m. \<forall>x \<in> A. \<exists>k\<le>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 "\<forall>x \<in> A. \<exists>k\<le>m. P x k" "P a n" + by force + then show ?case + apply (rule_tac x="max m n" in exI, safe) + using max.cobounded2 apply blast + by (meson le_max_iff_disj) +qed + +proposition proper_map: + fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel" + assumes "closedin (subtopology euclidean S) K" + and com: "\<And>U. \<lbrakk>U \<subseteq> T; compact U\<rbrakk> \<Longrightarrow> compact {x \<in> S. f x \<in> U}" + and "f ` S \<subseteq> T" + shows "closedin (subtopology euclidean T) (f ` K)" +proof - + have "K \<subseteq> S" + using assms closedin_imp_subset by metis + obtain C where "closed C" and Keq: "K = S \<inter> C" + using assms by (auto simp: closedin_closed) + have *: "y \<in> f ` K" if "y \<in> T" and y: "y islimpt f ` K" for y + proof - + obtain h where "\<forall>n. (\<exists>x\<in>K. h n = f x) \<and> h n \<noteq> y" "inj h" and hlim: "(h \<longlongrightarrow> y) sequentially" + using \<open>y \<in> T\<close> y by (force simp: limpt_sequential_inj) + then obtain X where X: "\<And>n. X n \<in> K \<and> h n = f (X n) \<and> h n \<noteq> y" + by metis + then have fX: "\<And>n. f (X n) = h n" + by metis + have "compact (C \<inter> {a \<in> S. f a \<in> insert y (range (\<lambda>i. f(X(n + i))))})" for n + apply (rule closed_Int_compact [OF \<open>closed C\<close>]) + apply (rule com) + using X \<open>K \<subseteq> S\<close> \<open>f ` S \<subseteq> T\<close> \<open>y \<in> T\<close> apply blast + apply (rule compact_sequence_with_limit) + apply (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) + done + then have comf: "compact {a \<in> K. f a \<in> insert y (range (\<lambda>i. f(X(n + i))))}" for n + by (simp add: Keq Int_def conj_commute) + have ne: "\<Inter>\<F> \<noteq> {}" + if "finite \<F>" + and \<F>: "\<And>t. t \<in> \<F> \<Longrightarrow> + (\<exists>n. t = {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (n + i))))})" + for \<F> + proof - + obtain m where m: "\<And>t. t \<in> \<F> \<Longrightarrow> \<exists>k\<le>m. t = {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (k + i))))}" + apply (rule exE) + apply (rule finite_indexed_bound [OF \<open>finite \<F>\<close> \<F>], assumption, force) + done + have "X m \<in> \<Inter>\<F>" + using X le_Suc_ex by (fastforce dest: m) + then show ?thesis by blast + qed + have "\<Inter>{{a. a \<in> K \<and> f a \<in> insert y (range (\<lambda>i. f(X(n + i))))} |n. n \<in> UNIV} + \<noteq> {}" + apply (rule compact_fip_heine_borel) + using comf apply force + using ne apply (simp add: subset_iff del: insert_iff) + done + then have "\<exists>x. x \<in> (\<Inter>n. {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (n + i))))})" + by blast + then show ?thesis + apply (simp add: image_iff fX) + by (metis \<open>inj h\<close> 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 \<Rightarrow> 'b::heine_borel" + assumes f: "inj_on f S" + shows "continuous_on S f \<longleftrightarrow> (\<forall>T. compact T \<and> T \<subseteq> S \<longrightarrow> 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: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" + by (metis inv_into_f_f f) + then have *: "{x \<in> S. f x \<in> U} = g ` U" if "U \<subseteq> f ` S" for U + using that by fastforce + have gfim: "g ` f ` S \<subseteq> S" using gf by auto + have **: "compact {x \<in> f ` S. g x \<in> C}" if C: "C \<subseteq> S" "compact C" for C + proof - + obtain h :: "'a set \<Rightarrow> 'a" where "h C \<in> C \<and> h C \<notin> S \<or> compact (f ` C)" + by (force simp: C RHS) + moreover have "f ` C = {b \<in> f ` S. g b \<in> C}" + using C gf by auto + ultimately show "compact {b \<in> f ` S. g b \<in> C}" + using C by auto + qed + show ?lhs + using proper_map [OF _ _ gfim] ** + by (simp add: continuous_on_closed * closedin_imp_subset) +qed + +lemma proper_map_from_compact: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" and "compact S" + "closedin (subtopology euclidean T) K" + shows "compact {x. x \<in> S \<and> f x \<in> K}" +by (rule closedin_compact [OF \<open>compact S\<close>] continuous_closedin_preimage_gen assms)+ + +lemma proper_map_fst: + assumes "compact T" "K \<subseteq> S" "compact K" + shows "compact {z \<in> S \<times> T. fst z \<in> K}" +proof - + have "{z \<in> S \<times> T. fst z \<in> K} = K \<times> 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 (subtopology euclidean (S \<times> T)) c" + shows "closedin (subtopology euclidean S) (fst ` c)" +proof - + have *: "fst ` (S \<times> T) \<subseteq> 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 \<subseteq> T" "compact K" + shows "compact {z \<in> S \<times> T. snd z \<in> K}" +proof - + have "{z \<in> S \<times> T. snd z \<in> K} = S \<times> 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 (subtopology euclidean (S \<times> T)) c" + shows "closedin (subtopology euclidean T) (snd ` c)" +proof - + have *: "snd ` (S \<times> T) \<subseteq> 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 (subtopology euclidean (S \<times> T)) U" + shows "closedin (subtopology euclidean T) {y. \<exists>x. x \<in> S \<and> (x, y) \<in> U}" +proof - + have "U \<subseteq> S \<times> T" + by (metis clo closedin_imp_subset) + then have "{y. \<exists>x. x \<in> S \<and> (x, y) \<in> U} = snd ` U" + by force + moreover have "closedin (subtopology euclidean 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. \<exists>x. x \<in> S \<and> (x, y) \<in> T}" +proof - + have *: "{y. \<exists>x. x \<in> S \<and> Pair x y \<in> T} = + {y. \<exists>x. x \<in> S \<and> Pair x y \<in> ((S \<times> UNIV) \<inter> T)}" + by auto + show ?thesis + apply (subst *) + apply (rule closedin_closed_trans [OF _ closed_UNIV]) + apply (rule closedin_compact_projection [OF \<open>compact S\<close>]) + by (simp add: clo closedin_closed_Int) +qed + +subsubsection\<open>Representing affine hull as a finite intersection of hyperplanes\<close> + +proposition affine_hull_convex_Int_nonempty_interior: + fixes S :: "'a::real_normed_vector set" + assumes "convex S" "S \<inter> interior T \<noteq> {}" + shows "affine hull (S \<inter> T) = affine hull S" +proof + show "affine hull (S \<inter> T) \<subseteq> affine hull S" + by (simp add: hull_mono) +next + obtain a where "a \<in> S" "a \<in> T" and at: "a \<in> interior T" + using assms interior_subset by blast + then obtain e where "e > 0" and e: "cball a e \<subseteq> T" + using mem_interior_cball by blast + have *: "x \<in> op + a ` span ((\<lambda>x. x - a) ` (S \<inter> T))" if "x \<in> 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 \<open>e > 0\<close> 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) \<ge> k" + using k_def by linarith + then have "a + k *\<^sub>R (x - a) \<in> cball a e" + using \<open>0 < k\<close> False by (simp add: dist_norm field_simps) + then have T: "a + k *\<^sub>R (x - a) \<in> T" + using e by blast + have S: "a + k *\<^sub>R (x - a) \<in> S" + using k \<open>a \<in> S\<close> convexD [OF \<open>convex S\<close> \<open>a \<in> S\<close> \<open>x \<in> S\<close>, of "1-k" k] + by (simp add: algebra_simps) + have "inverse k *\<^sub>R k *\<^sub>R (x-a) \<in> span ((\<lambda>x. x - a) ` (S \<inter> T))" + apply (rule span_mul) + apply (rule span_superset) + apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"]) + apply (auto simp: S T) + done + with xa image_iff show ?thesis by fastforce + qed + show "affine hull S \<subseteq> affine hull (S \<inter> T)" + apply (simp add: subset_hull) + apply (simp add: \<open>a \<in> S\<close> \<open>a \<in> T\<close> hull_inc affine_hull_span_gen [of a]) + apply (force simp: *) + done +qed + +corollary affine_hull_convex_Int_open: + fixes S :: "'a::real_normed_vector set" + assumes "convex S" "open T" "S \<inter> T \<noteq> {}" + shows "affine hull (S \<inter> 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 \<inter> interior T \<noteq> {}" + shows "affine hull (S \<inter> 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 \<inter> T \<noteq> {}" + shows "affine hull (S \<inter> 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 (subtopology euclidean (affine hull S)) T" "S \<inter> T \<noteq> {}" + shows "affine hull (S \<inter> 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_open_in: + fixes S :: "'a::real_normed_vector set" + assumes "openin (subtopology euclidean (affine hull T)) S" "S \<noteq> {}" + 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 \<noteq> {}" + 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 "\<lbrakk>convex S; S \<inter> interior T \<noteq> {}\<rbrakk> \<Longrightarrow> aff_dim(S \<inter> 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 "\<lbrakk>convex S; open T; S \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow> aff_dim(S \<inter> T) = aff_dim S" +using aff_dim_convex_Int_nonempty_interior interior_eq by blast + +lemma affine_hull_halfspace_lt: + fixes a :: "'a::euclidean_space" + shows "affine hull {x. a \<bullet> x < r} = (if a = 0 \<and> r \<le> 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 \<bullet> x \<le> r} = (if a = 0 \<and> 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 \<bullet> x < r} = UNIV" + using affine_hull_halfspace_lt closure_same_affine_hull by fastforce + moreover have "{x. a \<bullet> x < r} \<subseteq> {x. a \<bullet> x \<le> 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 \<bullet> x > r} = (if a = 0 \<and> r \<ge> 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 \<bullet> x \<ge> r} = (if a = 0 \<and> 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 \<bullet> x < r} = + (if a = 0 \<and> r \<le> 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 \<bullet> x \<le> r} = + (if a = 0 \<and> r < 0 then -1 else DIM('a))" +proof - + have "int (DIM('a)) = aff_dim (UNIV::'a set)" + by (simp add: aff_dim_UNIV) + then have "aff_dim (affine hull {x. a \<bullet> x \<le> r}) = DIM('a)" if "(a = 0 \<longrightarrow> r \<ge> 0)" + using that by (simp add: affine_hull_halfspace_le not_less) + then show ?thesis + by (force simp: aff_dim_affine_hull) +qed + +lemma aff_dim_halfspace_gt: + fixes a :: "'a::euclidean_space" + shows "aff_dim {x. a \<bullet> x > r} = + (if a = 0 \<and> r \<ge> 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 \<bullet> x \<ge> r} = + (if a = 0 \<and> r > 0 then -1 else DIM('a))" +using aff_dim_halfspace_le [of "-a" "-r"] by simp + +subsection\<open>Properties of special hyperplanes\<close> + +lemma subspace_hyperplane: "subspace {x. a \<bullet> x = 0}" + by (simp add: subspace_def inner_right_distrib) + +lemma subspace_hyperplane2: "subspace {x. x \<bullet> a = 0}" + by (simp add: inner_commute inner_right_distrib subspace_def) + +lemma special_hyperplane_span: + fixes S :: "'n::euclidean_space set" + assumes "k \<in> Basis" + shows "{x. k \<bullet> x = 0} = span (Basis - {k})" +proof - + have *: "x \<in> span (Basis - {k})" if "k \<bullet> x = 0" for x + proof - + have "x = (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b)" + by (simp add: euclidean_representation) + also have "... = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)" + by (auto simp: setsum.remove [of _ k] inner_commute assms that) + finally have "x = (\<Sum>b\<in>Basis - {k}. (x \<bullet> b) *\<^sub>R b)" . + then show ?thesis + by (simp add: Linear_Algebra.span_finite) metis + qed + show ?thesis + apply (rule span_subspace [symmetric]) + using assms + apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) + done +qed + +lemma dim_special_hyperplane: + fixes k :: "'n::euclidean_space" + shows "k \<in> Basis \<Longrightarrow> dim {x. k \<bullet> x = 0} = DIM('n) - 1" +apply (simp add: special_hyperplane_span) +apply (rule Linear_Algebra.dim_unique [OF subset_refl]) +apply (auto simp: Diff_subset independent_substdbasis) +apply (metis member_remove remove_def span_clauses(1)) +done + +proposition dim_hyperplane: + fixes a :: "'a::euclidean_space" + assumes "a \<noteq> 0" + shows "dim {x. a \<bullet> x = 0} = DIM('a) - 1" +proof - + have span0: "span {x. a \<bullet> x = 0} = {x. a \<bullet> x = 0}" + by (rule span_unique) (auto simp: subspace_hyperplane) + then obtain B where "independent B" + and Bsub: "B \<subseteq> {x. a \<bullet> x = 0}" + and subspB: "{x. a \<bullet> x = 0} \<subseteq> span B" + and card0: "(card B = dim {x. a \<bullet> x = 0})" + and ortho: "pairwise orthogonal B" + using orthogonal_basis_exists by metis + with assms have "a \<notin> span B" + by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0 span_subspace) + then have ind: "independent (insert a B)" + by (simp add: \<open>independent B\<close> independent_insert) + have "finite B" + using \<open>independent B\<close> independent_bound by blast + have "UNIV \<subseteq> span (insert a B)" + proof fix y::'a + obtain r z where z: "y = r *\<^sub>R a + z" "a \<bullet> z = 0" + apply (rule_tac r="(a \<bullet> y) / (a \<bullet> a)" and z = "y - ((a \<bullet> y) / (a \<bullet> a)) *\<^sub>R a" in that) + using assms + by (auto simp: algebra_simps) + show "y \<in> span (insert a B)" + by (metis (mono_tags, lifting) z Bsub Convex_Euclidean_Space.span_eq + add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) + qed + then have dima: "DIM('a) = dim(insert a B)" + by (metis antisym dim_UNIV dim_subset_UNIV subset_le_dim) + then show ?thesis + by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \<open>a \<notin> span B\<close> ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB) +qed + +lemma lowdim_eq_hyperplane: + fixes S :: "'a::euclidean_space set" + assumes "dim S = DIM('a) - 1" + obtains a where "a \<noteq> 0" and "span S = {x. a \<bullet> x = 0}" +proof - + have [simp]: "dim S < DIM('a)" + by (simp add: DIM_positive assms) + then obtain b where b: "b \<noteq> 0" "span S \<subseteq> {a. b \<bullet> a = 0}" + using lowdim_subset_hyperplane [of S] by fastforce + show ?thesis + using b that real_vector_class.subspace_span [of S] + by (simp add: assms dim_hyperplane subspace_dim_equal subspace_hyperplane) +qed + +lemma dim_eq_hyperplane: + fixes S :: "'n::euclidean_space set" + shows "dim S = DIM('n) - 1 \<longleftrightarrow> (\<exists>a. a \<noteq> 0 \<and> span S = {x. a \<bullet> x = 0})" +by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) + +proposition aff_dim_eq_hyperplane: + fixes S :: "'a::euclidean_space set" + shows "aff_dim S = DIM('a) - 1 \<longleftrightarrow> (\<exists>a b. a \<noteq> 0 \<and> affine hull S = {x. a \<bullet> x = b})" +proof (cases "S = {}") + case True then show ?thesis + by (auto simp: dest: hyperplane_eq_Ex) +next + case False + then obtain c where "c \<in> S" by blast + show ?thesis + proof (cases "c = 0") + case True show ?thesis + apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane + del: One_nat_def) + apply (rule ex_cong) + apply (metis (mono_tags) span_0 \<open>c = 0\<close> image_add_0 inner_zero_right mem_Collect_eq) + done + next + case False + have xc_im: "x \<in> op + c ` {y. a \<bullet> y = 0}" if "a \<bullet> x = a \<bullet> c" for a x + proof - + have "\<exists>y. a \<bullet> y = 0 \<and> c + y = x" + by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) + then show "x \<in> op + c ` {y. a \<bullet> y = 0}" + by blast + qed + have 2: "span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = 0}" + if "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = b}" for a b + proof - + have "b = a \<bullet> c" + using span_0 that by fastforce + with that have "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = a \<bullet> c}" + by simp + then have "span ((\<lambda>x. x - c) ` S) = (\<lambda>x. x - c) ` {x. a \<bullet> x = a \<bullet> c}" + by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) + also have "... = {x. a \<bullet> x = 0}" + by (force simp: inner_distrib inner_diff_right + intro: image_eqI [where x="x+c" for x]) + finally show ?thesis . + qed + show ?thesis + apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane + del: One_nat_def, safe) + apply (fastforce simp add: inner_distrib intro: xc_im) + apply (force simp: intro!: 2) + done + qed +qed + +corollary aff_dim_hyperplane [simp]: + fixes a :: "'a::euclidean_space" + shows "a \<noteq> 0 \<Longrightarrow> aff_dim {x. a \<bullet> x = r} = DIM('a) - 1" +by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) + +subsection\<open>Some stepping theorems\<close> + +lemma dim_empty [simp]: "dim ({} :: 'a::euclidean_space set) = 0" + by (force intro!: dim_unique) + +lemma dim_insert: + fixes x :: "'a::euclidean_space" + shows "dim (insert x S) = (if x \<in> span S then dim S else dim S + 1)" +proof - + show ?thesis + proof (cases "x \<in> span S") + case True then show ?thesis + by (metis dim_span span_redundant) + next + case False + obtain B where B: "B \<subseteq> span S" "independent B" "span S \<subseteq> span B" "card B = dim (span S)" + using basis_exists [of "span S"] by blast + have 1: "insert x B \<subseteq> span (insert x S)" + by (meson \<open>B \<subseteq> span S\<close> dual_order.trans insertI1 insert_subsetI span_mono span_superset subset_insertI) + have 2: "span (insert x S) \<subseteq> span (insert x B)" + by (metis \<open>B \<subseteq> span S\<close> \<open>span S \<subseteq> span B\<close> span_breakdown_eq span_subspace subsetI subspace_span) + have 3: "independent (insert x B)" + by (metis B independent_insert span_subspace subspace_span False) + have "dim (span (insert x S)) = Suc (dim S)" + apply (rule dim_unique [OF 1 2 3]) + by (metis B False card_insert_disjoint dim_span independent_imp_finite subsetCE) + then show ?thesis + by (simp add: False) + qed +qed + +lemma dim_singleton [simp]: + fixes x :: "'a::euclidean_space" + shows "dim{x} = (if x = 0 then 0 else 1)" +by (simp add: dim_insert) + +lemma dim_eq_0 [simp]: + fixes S :: "'a::euclidean_space set" + shows "dim S = 0 \<longleftrightarrow> S \<subseteq> {0}" +apply safe +apply (metis DIM_positive DIM_real card_ge_dim_independent contra_subsetD dim_empty dim_insert dim_singleton empty_subsetI independent_empty less_not_refl zero_le) +by (metis dim_singleton dim_subset le_0_eq) + +lemma aff_dim_insert: + fixes a :: "'a::euclidean_space" + shows "aff_dim (insert a S) = (if a \<in> 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 \<notin> s'" + by (meson Set.set_insert all_not_in_conv) + show ?thesis using S + apply (simp add: hull_redundant cong: aff_dim_affine_hull2) + apply (simp add: affine_hull_insert_span_gen hull_inc) + apply (simp add: insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert span_0) + apply (metis (no_types, lifting) add_minus_cancel image_iff uminus_add_conv_diff) + done +qed + +lemma subspace_bounded_eq_trivial: + fixes S :: "'a::real_normed_vector set" + assumes "subspace S" + shows "bounded S \<longleftrightarrow> S = {0}" +proof - + have "False" if "bounded S" "x \<in> S" "x \<noteq> 0" for x + proof - + obtain B where B: "\<And>y. y \<in> S \<Longrightarrow> norm y < B" "B > 0" + using \<open>bounded S\<close> by (force simp: bounded_pos_less) + have "(B / norm x) *\<^sub>R x \<in> S" + using assms subspace_mul \<open>x \<in> S\<close> 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 \<Longrightarrow> 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 \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})" +proof (cases "S = {}") + case True then show ?thesis + by simp +next + case False + then obtain b where "b \<in> S" by blast + with False assms show ?thesis + apply safe + using affine_diffs_subspace [OF assms \<open>b \<in> S\<close>] + apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation + image_empty image_insert translation_invert) + apply force + done +qed + +lemma affine_bounded_eq_lowdim: + fixes S :: "'a::euclidean_space set" + assumes "affine S" + shows "bounded S \<longleftrightarrow> aff_dim S \<le> 0" +apply safe +using affine_bounded_eq_trivial assms apply fastforce +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) + + +lemma bounded_hyperplane_eq_trivial_0: + fixes a :: "'a::euclidean_space" + assumes "a \<noteq> 0" + shows "bounded {x. a \<bullet> x = 0} \<longleftrightarrow> DIM('a) = 1" +proof + assume "bounded {x. a \<bullet> x = 0}" + then have "aff_dim {x. a \<bullet> x = 0} \<le> 0" + by (simp add: affine_bounded_eq_lowdim affine_hyperplane) + with assms show "DIM('a) = 1" + by (simp add: le_Suc_eq aff_dim_hyperplane) +next + assume "DIM('a) = 1" + then show "bounded {x. a \<bullet> x = 0}" + by (simp add: aff_dim_hyperplane affine_bounded_eq_lowdim affine_hyperplane assms) +qed + +lemma bounded_hyperplane_eq_trivial: + fixes a :: "'a::euclidean_space" + shows "bounded {x. a \<bullet> x = r} \<longleftrightarrow> (if a = 0 then r \<noteq> 0 else DIM('a) = 1)" +proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) + assume "r \<noteq> 0" "a \<noteq> 0" + have "aff_dim {x. y \<bullet> x = 0} = aff_dim {x. a \<bullet> x = r}" if "y \<noteq> 0" for y::'a + by (metis that \<open>a \<noteq> 0\<close> aff_dim_hyperplane) + then show "bounded {x. a \<bullet> x = r} = (DIM('a) = Suc 0)" + by (metis One_nat_def \<open>a \<noteq> 0\<close> affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) +qed + +subsection\<open>General case without assuming closure and getting non-strict separation\<close> + +proposition separating_hyperplane_closed_point_inset: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "closed S" "S \<noteq> {}" "z \<notin> S" + obtains a b where "a \<in> S" "(a - z) \<bullet> z < b" "\<And>x. x \<in> S \<Longrightarrow> b < (a - z) \<bullet> x" +proof - + obtain y where "y \<in> S" and y: "\<And>u. u \<in> S \<Longrightarrow> dist z y \<le> dist z u" + using distance_attains_inf [of S z] assms by auto + then have *: "(y - z) \<bullet> z < (y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2" + using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto + show ?thesis + proof (rule that [OF \<open>y \<in> S\<close> *]) + fix x + assume "x \<in> S" + have yz: "0 < (y - z) \<bullet> (y - z)" + using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto + { assume 0: "0 < ((z - y) \<bullet> (x - y))" + with any_closest_point_dot [OF \<open>convex S\<close> \<open>closed S\<close>] + have False + using y \<open>x \<in> S\<close> \<open>y \<in> S\<close> not_less by blast + } + then have "0 \<le> ((y - z) \<bullet> (x - y))" + by (force simp: not_less inner_diff_left) + with yz have "0 < 2 * ((y - z) \<bullet> (x - y)) + (y - z) \<bullet> (y - z)" + by (simp add: algebra_simps) + then show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> 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 \<noteq> {}" "0 \<notin> S" + obtains a b where "a \<in> S" "a \<noteq> 0" "0 < b" "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x > b" +using separating_hyperplane_closed_point_inset [OF assms] +by simp (metis \<open>0 \<notin> S\<close>) + + +proposition separating_hyperplane_set_0_inspan: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "S \<noteq> {}" "0 \<notin> S" + obtains a where "a \<in> span S" "a \<noteq> 0" "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> a \<bullet> x" +proof - + define k where [abs_def]: "k c = {x. 0 \<le> c \<bullet> x}" for c :: 'a + have *: "span S \<inter> frontier (cball 0 1) \<inter> \<Inter>f' \<noteq> {}" + if f': "finite f'" "f' \<subseteq> k ` S" for f' + proof - + obtain C where "C \<subseteq> S" "finite C" and C: "f' = k ` C" + using finite_subset_image [OF f'] by blast + obtain a where "a \<in> S" "a \<noteq> 0" + using \<open>S \<noteq> {}\<close> \<open>0 \<notin> S\<close> ex_in_conv by blast + then have "norm (a /\<^sub>R (norm a)) = 1" + by simp + moreover have "a /\<^sub>R (norm a) \<in> span S" + by (simp add: \<open>a \<in> S\<close> span_mul span_superset) + ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> 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 \<open>finite C\<close> compact_eq_bounded_closed finite_imp_compact_convex_hull by auto + moreover have "convex hull C \<noteq> {}" + by (simp add: False) + moreover have "0 \<notin> convex hull C" + by (metis \<open>C \<subseteq> S\<close> \<open>convex S\<close> \<open>0 \<notin> S\<close> convex_hull_subset hull_same insert_absorb insert_subset) + ultimately obtain a b + where "a \<in> convex hull C" "a \<noteq> 0" "0 < b" + and ab: "\<And>x. x \<in> convex hull C \<Longrightarrow> a \<bullet> x > b" + using separating_hyperplane_closed_0_inset by blast + then have "a \<in> S" + by (metis \<open>C \<subseteq> S\<close> assms(1) subsetCE subset_hull) + moreover have "norm (a /\<^sub>R (norm a)) = 1" + using \<open>a \<noteq> 0\<close> by simp + moreover have "a /\<^sub>R (norm a) \<in> span S" + by (simp add: \<open>a \<in> S\<close> span_mul span_superset) + ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> sphere 0 1" + by simp + have aa: "a /\<^sub>R (norm a) \<in> (\<Inter>c\<in>C. {x. 0 \<le> c \<bullet> x})" + apply (clarsimp simp add: divide_simps) + using ab \<open>0 < b\<close> + by (metis hull_inc inner_commute less_eq_real_def less_trans) + show ?thesis + apply (simp add: C k_def) + using ass aa Int_iff empty_iff by blast + qed + qed + have "(span S \<inter> frontier(cball 0 1)) \<inter> (\<Inter> (k ` S)) \<noteq> {}" + apply (rule compact_imp_fip) + apply (blast intro: compact_cball) + using closed_halfspace_ge k_def apply blast + apply (metis *) + done + 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 \<noteq> {}" and zno: "z \<notin> S" + obtains a b where "(z + a) \<in> affine hull (insert z S)" + and "a \<noteq> 0" and "a \<bullet> z \<le> b" + and "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b" +proof - +from separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"] + have "convex (op + (- z) ` S)" + by (simp add: \<open>convex S\<close>) + moreover have "op + (- z) ` S \<noteq> {}" + by (simp add: \<open>S \<noteq> {}\<close>) + moreover have "0 \<notin> op + (- z) ` S" + using zno by auto + ultimately obtain a where "a \<in> span (op + (- z) ` S)" "a \<noteq> 0" + and a: "\<And>x. x \<in> (op + (- z) ` S) \<Longrightarrow> 0 \<le> a \<bullet> x" + using separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"] + by blast + then have szx: "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> z \<le> a \<bullet> 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) + show ?thesis + apply (rule_tac a=a and b = "a \<bullet> z" in that, simp_all) + using \<open>a \<in> span (op + (- z) ` S)\<close> affine_hull_insert_span_gen apply blast + apply (simp_all add: \<open>a \<noteq> 0\<close> szx) + done +qed + +proposition supporting_hyperplane_rel_boundary: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "x \<in> S" and xno: "x \<notin> rel_interior S" + obtains a where "a \<noteq> 0" + and "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y" + and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y" +proof - + obtain a b where aff: "(x + a) \<in> affine hull (insert x (rel_interior S))" + and "a \<noteq> 0" and "a \<bullet> x \<le> b" + and ageb: "\<And>u. u \<in> (rel_interior S) \<Longrightarrow> a \<bullet> u \<ge> 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 \<bullet> x \<le> a \<bullet> y" if "y \<in> S" for y + proof - + have con: "continuous_on (closure (rel_interior S)) (op \<bullet> a)" + by (rule continuous_intros continuous_on_subset | blast)+ + have y: "y \<in> closure (rel_interior S)" + using \<open>convex S\<close> closure_def convex_closure_rel_interior \<open>y \<in> S\<close> + by fastforce + show ?thesis + using continuous_ge_on_closure [OF con y] ageb \<open>a \<bullet> x \<le> b\<close> + by fastforce + qed + have 3: "a \<bullet> x < a \<bullet> y" if "y \<in> rel_interior S" for y + proof - + obtain e where "0 < e" "y \<in> S" and e: "cball y e \<inter> affine hull S \<subseteq> S" + using \<open>y \<in> rel_interior S\<close> by (force simp: rel_interior_cball) + define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" + have "y' \<in> cball y e" + unfolding y'_def using \<open>0 < e\<close> by force + moreover have "y' \<in> affine hull S" + unfolding y'_def + by (metis \<open>x \<in> S\<close> \<open>y \<in> S\<close> \<open>convex S\<close> aff affine_affine_hull hull_redundant + rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) + ultimately have "y' \<in> S" + using e by auto + have "a \<bullet> x \<le> a \<bullet> y" + using le_ay \<open>a \<noteq> 0\<close> \<open>y \<in> S\<close> by blast + moreover have "a \<bullet> x \<noteq> a \<bullet> y" + using le_ay [OF \<open>y' \<in> S\<close>] \<open>a \<noteq> 0\<close> + apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square) + by (metis \<open>0 < e\<close> add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2) + ultimately show ?thesis by force + qed + show ?thesis + by (rule that [OF \<open>a \<noteq> 0\<close> le_ay 3]) +qed + +lemma supporting_hyperplane_relative_frontier: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "x \<in> closure S" "x \<notin> rel_interior S" + obtains a where "a \<noteq> 0" + and "\<And>y. y \<in> closure S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y" + and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y" +using supporting_hyperplane_rel_boundary [of "closure S" x] +by (metis assms convex_closure convex_rel_interior_closure) + +subsection\<open> Some results on decomposing convex hulls, e.g. simplicial subdivision\<close> + +lemma + fixes s :: "'a::euclidean_space set" + assumes "~ (affine_dependent(s \<union> t))" + shows convex_hull_Int_subset: "convex hull s \<inter> convex hull t \<subseteq> convex hull (s \<inter> t)" (is ?C) + and affine_hull_Int_subset: "affine hull s \<inter> affine hull t \<subseteq> affine hull (s \<inter> t)" (is ?A) +proof - + have [simp]: "finite s" "finite t" + using aff_independent_finite assms by blast+ + have "setsum u (s \<inter> t) = 1 \<and> + (\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)" + if [simp]: "setsum u s = 1" + "setsum v t = 1" + and eq: "(\<Sum>x\<in>t. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" for u v + proof - + define f where "f x = (if x \<in> s then u x else 0) - (if x \<in> t then v x else 0)" for x + have "setsum f (s \<union> t) = 0" + apply (simp add: f_def setsum_Un setsum_subtractf) + apply (simp add: setsum.inter_restrict [symmetric] Int_commute) + done + moreover have "(\<Sum>x\<in>(s \<union> t). f x *\<^sub>R x) = 0" + apply (simp add: f_def setsum_Un scaleR_left_diff_distrib setsum_subtractf) + apply (simp add: if_smult setsum.inter_restrict [symmetric] Int_commute eq + cong del: if_weak_cong) + done + ultimately have "\<And>v. v \<in> s \<union> t \<Longrightarrow> f v = 0" + using aff_independent_finite assms unfolding affine_dependent_explicit + by blast + then have u [simp]: "\<And>x. x \<in> s \<Longrightarrow> u x = (if x \<in> t then v x else 0)" + by (simp add: f_def) presburger + have "setsum u (s \<inter> t) = setsum u s" + by (simp add: setsum.inter_restrict) + then have "setsum u (s \<inter> t) = 1" + using that by linarith + moreover have "(\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)" + by (auto simp: if_smult setsum.inter_restrict intro: setsum.cong) + ultimately show ?thesis + by force + qed + then show ?A ?C + by (auto simp: convex_hull_finite affine_hull_finite) +qed + + +proposition affine_hull_Int: + fixes s :: "'a::euclidean_space set" + assumes "~ (affine_dependent(s \<union> t))" + shows "affine hull (s \<inter> t) = affine hull s \<inter> affine hull t" +apply (rule subset_antisym) +apply (simp add: hull_mono) +by (simp add: affine_hull_Int_subset assms) + +proposition convex_hull_Int: + fixes s :: "'a::euclidean_space set" + assumes "~ (affine_dependent(s \<union> t))" + shows "convex hull (s \<inter> t) = convex hull s \<inter> convex hull t" +apply (rule subset_antisym) +apply (simp add: hull_mono) +by (simp add: convex_hull_Int_subset assms) + +proposition + fixes s :: "'a::euclidean_space set set" + assumes "~ (affine_dependent (\<Union>s))" + shows affine_hull_Inter: "affine hull (\<Inter>s) = (\<Inter>t\<in>s. affine hull t)" (is "?A") + and convex_hull_Inter: "convex hull (\<Inter>s) = (\<Inter>t\<in>s. convex hull t)" (is "?C") +proof - + have "finite s" + using aff_independent_finite assms finite_UnionD by blast + then have "?A \<and> ?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]: "\<not> affine_dependent (t \<union> \<Inter>F)" + by (auto intro: affine_dependent_subset) + have [simp]: "\<not> affine_dependent (\<Union>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 + +lemma affine_hull_finite_intersection_hyperplanes: + fixes s :: "'a::euclidean_space set" + obtains f where + "finite f" + "of_nat (card f) + aff_dim s = DIM('a)" + "affine hull s = \<Inter>f" + "\<And>h. h \<in> f \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x = b}" +proof - + obtain b where "b \<subseteq> s" + and indb: "\<not> affine_dependent b" + and eq: "affine hull s = affine hull b" + using affine_basis_exists by blast + obtain c where indc: "\<not> affine_dependent c" and "b \<subseteq> 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 \<le> card c" + using \<open>b \<subseteq> c\<close> infinite_super by (auto simp: card_mono) + have imeq: "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b)) = ((\<lambda>a. affine hull (c - {a})) ` (c - b))" + by blast + have card1: "card ((\<lambda>a. affine hull (c - {a})) ` (c - b)) = card (c - b)" + apply (rule card_image [OF inj_onI]) + by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) + have card2: "(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: \<open>\<not> affine_dependent c\<close> \<open>b \<subseteq> c\<close> aff_dim_affine_independent aff_dim_UNIV card_Diff_subset of_nat_diff) + qed + show ?thesis + proof (cases "c = b") + case True show ?thesis + apply (rule_tac f="{}" in that) + using True affc + apply (simp_all add: eq [symmetric]) + by (metis aff_dim_UNIV aff_dim_affine_hull) + next + case False + have ind: "\<not> affine_dependent (\<Union>a\<in>c - b. c - {a})" + by (rule affine_independent_subset [OF indc]) auto + have affeq: "affine hull s = (\<Inter>x\<in>(\<lambda>a. c - {a}) ` (c - b). affine hull x)" + using \<open>b \<subseteq> c\<close> False + apply (subst affine_hull_Inter [OF ind, symmetric]) + apply (simp add: eq double_diff) + done + have *: "1 + aff_dim (c - {t}) = int (DIM('a))" + if t: "t \<in> 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 + show ?thesis + apply (rule_tac f = "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b))" in that) + using \<open>finite c\<close> apply blast + apply (simp add: imeq card1 card2) + apply (simp add: affeq, clarify) + apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *) + done + qed +qed + +subsection\<open>Misc results about span\<close> + +lemma eq_span_insert_eq: + assumes "(x - y) \<in> span S" + shows "span(insert x S) = span(insert y S)" +proof - + have *: "span(insert x S) \<subseteq> span(insert y S)" if "(x - y) \<in> span S" for x y + proof - + have 1: "(r *\<^sub>R x - r *\<^sub>R y) \<in> span S" for r + by (metis real_vector.scale_right_diff_distrib span_mul that) + have 2: "(z - k *\<^sub>R y) - k *\<^sub>R (x - y) = z - k *\<^sub>R x" for z k + by (simp add: real_vector.scale_right_diff_distrib) + show ?thesis + apply (clarsimp simp add: span_breakdown_eq) + by (metis 1 2 diff_add_cancel real_vector.scale_right_diff_distrib span_add_eq) + qed + show ?thesis + apply (intro subset_antisym * assms) + using assms subspace_neg subspace_span minus_diff_eq by force +qed + +lemma dim_psubset: + fixes S :: "'a :: euclidean_space set" + shows "span S \<subset> span T \<Longrightarrow> dim S < dim T" +by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span) + + +lemma basis_subspace_exists: + fixes S :: "'a::euclidean_space set" + shows + "subspace S + \<Longrightarrow> \<exists>b. finite b \<and> b \<subseteq> S \<and> + independent b \<and> span b = S \<and> card b = dim S" +by (metis span_subspace basis_exists independent_imp_finite) + +lemma affine_hyperplane_sums_eq_UNIV_0: + fixes S :: "'a :: euclidean_space set" + assumes "affine S" + and "0 \<in> S" and "w \<in> S" + and "a \<bullet> w \<noteq> 0" + shows "{x + y| x y. x \<in> S \<and> a \<bullet> y = 0} = UNIV" +proof - + have "subspace S" + by (simp add: assms subspace_affine) + have span1: "span {y. a \<bullet> y = 0} \<subseteq> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}" + apply (rule span_mono) + using \<open>0 \<in> S\<close> add.left_neutral by force + have "w \<notin> span {y. a \<bullet> y = 0}" + using \<open>a \<bullet> w \<noteq> 0\<close> span_induct subspace_hyperplane by auto + moreover have "w \<in> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}" + using \<open>w \<in> S\<close> + by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_superset) + ultimately have span2: "span {y. a \<bullet> y = 0} \<noteq> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}" + by blast + have "a \<noteq> 0" using assms inner_zero_left by blast + then have "DIM('a) - 1 = dim {y. a \<bullet> y = 0}" + by (simp add: dim_hyperplane) + also have "... < dim {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}" + using span1 span2 by (blast intro: dim_psubset) + finally have DIM_lt: "DIM('a) - 1 < dim {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}" . + have subs: "subspace {x + y| x y. x \<in> S \<and> a \<bullet> y = 0}" + using subspace_sums [OF \<open>subspace S\<close> subspace_hyperplane] by simp + moreover have "span {x + y| x y. x \<in> S \<and> a \<bullet> y = 0} = UNIV" + apply (rule dim_eq_full [THEN iffD1]) + apply (rule antisym [OF dim_subset_UNIV]) + using DIM_lt apply simp + done + ultimately show ?thesis + by (simp add: subs) (metis (lifting) span_eq subs) +qed + +proposition affine_hyperplane_sums_eq_UNIV: + fixes S :: "'a :: euclidean_space set" + assumes "affine S" + and "S \<inter> {v. a \<bullet> v = b} \<noteq> {}" + and "S - {v. a \<bullet> v = b} \<noteq> {}" + shows "{x + y| x y. x \<in> S \<and> a \<bullet> 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 \<in> S" and c: "a \<bullet> c = b" + using assms by force + with affine_diffs_subspace [OF \<open>affine S\<close>] + have "subspace (op + (- c) ` S)" by blast + then have aff: "affine (op + (- c) ` S)" + by (simp add: subspace_imp_affine) + have 0: "0 \<in> op + (- c) ` S" + by (simp add: \<open>c \<in> S\<close>) + obtain d where "d \<in> S" and "a \<bullet> d \<noteq> b" and dc: "d-c \<in> op + (- c) ` S" + using assms by auto + then have adc: "a \<bullet> (d - c) \<noteq> 0" + by (simp add: c inner_diff_right) + let ?U = "op + (c+c) ` {x + y |x y. x \<in> op + (- c) ` S \<and> a \<bullet> y = 0}" + have "u + v \<in> op + (c + c) ` {x + v |x v. x \<in> op + (- c) ` S \<and> a \<bullet> v = 0}" + if "u \<in> S" "b = a \<bullet> v" for u v + apply (rule_tac x="u+v-c-c" in image_eqI) + apply (simp_all add: algebra_simps) + apply (rule_tac x="u-c" in exI) + apply (rule_tac x="v-c" in exI) + apply (simp add: algebra_simps that c) + done + moreover have "\<lbrakk>a \<bullet> v = 0; u \<in> S\<rbrakk> + \<Longrightarrow> \<exists>x ya. v + (u + c) = x + ya \<and> x \<in> S \<and> a \<bullet> ya = b" for v u + by (metis add.left_commute c inner_right_distrib pth_d) + ultimately have "{x + y |x y. x \<in> S \<and> a \<bullet> y = b} = ?U" + by (fastforce simp: algebra_simps) + also have "... = op + (c+c) ` UNIV" + by (simp add: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) + also have "... = UNIV" + by (simp add: translation_UNIV) + finally show ?thesis . +qed + +proposition dim_sums_Int: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" "subspace T" + shows "dim {x + y |x y. x \<in> S \<and> y \<in> T} + dim(S \<inter> T) = dim S + dim T" +proof - + obtain B where B: "B \<subseteq> S \<inter> T" "S \<inter> T \<subseteq> span B" + and indB: "independent B" + and cardB: "card B = dim (S \<inter> T)" + using basis_exists by blast + then obtain C D where "B \<subseteq> C" "C \<subseteq> S" "independent C" "S \<subseteq> span C" + and "B \<subseteq> D" "D \<subseteq> T" "independent D" "T \<subseteq> span D" + using maximal_independent_subset_extend + by (metis Int_subset_iff \<open>B \<subseteq> S \<inter> T\<close> indB) + then have "finite B" "finite C" "finite D" + by (simp_all add: independent_imp_finite indB independent_bound) + have Beq: "B = C \<inter> D" + apply (rule sym) + apply (rule spanning_subset_independent) + using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> apply blast + apply (meson \<open>independent C\<close> independent_mono inf.cobounded1) + using B \<open>C \<subseteq> S\<close> \<open>D \<subseteq> T\<close> apply auto + done + then have Deq: "D = B \<union> (D - C)" + by blast + have CUD: "C \<union> D \<subseteq> {x + y |x y. x \<in> S \<and> y \<in> T}" + apply safe + apply (metis add.right_neutral subsetCE \<open>C \<subseteq> S\<close> \<open>subspace T\<close> set_eq_subset span_0 span_minimal) + apply (metis add.left_neutral subsetCE \<open>D \<subseteq> T\<close> \<open>subspace S\<close> set_eq_subset span_0 span_minimal) + done + have "a v = 0" if 0: "(\<Sum>v\<in>C. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v) = 0" + and v: "v \<in> C \<union> (D-C)" for a v + proof - + have eq: "(\<Sum>v\<in>D - C. a v *\<^sub>R v) = - (\<Sum>v\<in>C. a v *\<^sub>R v)" + using that add_eq_0_iff by blast + have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> S" + apply (subst eq) + apply (rule subspace_neg [OF \<open>subspace S\<close>]) + apply (rule subspace_setsum [OF \<open>subspace S\<close>]) + by (meson subsetCE subspace_mul \<open>C \<subseteq> S\<close> \<open>subspace S\<close>) + moreover have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> T" + apply (rule subspace_setsum [OF \<open>subspace T\<close>]) + by (meson DiffD1 \<open>D \<subseteq> T\<close> \<open>subspace T\<close> subset_eq subspace_def) + ultimately have "(\<Sum>v \<in> D-C. a v *\<^sub>R v) \<in> span B" + using B by blast + then obtain e where e: "(\<Sum>v\<in>B. e v *\<^sub>R v) = (\<Sum>v \<in> D-C. a v *\<^sub>R v)" + using span_finite [OF \<open>finite B\<close>] by blast + have "\<And>c v. \<lbrakk>(\<Sum>v\<in>C. c v *\<^sub>R v) = 0; v \<in> C\<rbrakk> \<Longrightarrow> c v = 0" + using independent_explicit \<open>independent C\<close> by blast + define cc where "cc x = (if x \<in> B then a x + e x else a x)" for x + have [simp]: "C \<inter> B = B" "D \<inter> B = B" "C \<inter> - B = C-D" "B \<inter> (D - C) = {}" + using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> Beq by blast+ + have f2: "(\<Sum>v\<in>C \<inter> D. e v *\<^sub>R v) = (\<Sum>v\<in>D - C. a v *\<^sub>R v)" + using Beq e by presburger + have f3: "(\<Sum>v\<in>C \<union> D. a v *\<^sub>R v) = (\<Sum>v\<in>C - D. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v) + (\<Sum>v\<in>C \<inter> D. a v *\<^sub>R v)" + using \<open>finite C\<close> \<open>finite D\<close> setsum.union_diff2 by blast + have f4: "(\<Sum>v\<in>C \<union> (D - C). a v *\<^sub>R v) = (\<Sum>v\<in>C. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v)" + by (meson Diff_disjoint \<open>finite C\<close> \<open>finite D\<close> finite_Diff setsum.union_disjoint) + have "(\<Sum>v\<in>C. cc v *\<^sub>R v) = 0" + using 0 f2 f3 f4 + apply (simp add: cc_def Beq if_smult \<open>finite C\<close> setsum.If_cases algebra_simps setsum.distrib) + apply (simp add: add.commute add.left_commute diff_eq) + done + then have "\<And>v. v \<in> C \<Longrightarrow> cc v = 0" + using independent_explicit \<open>independent C\<close> by blast + then have C0: "\<And>v. v \<in> C - B \<Longrightarrow> a v = 0" + by (simp add: cc_def Beq) meson + then have [simp]: "(\<Sum>x\<in>C - B. a x *\<^sub>R x) = 0" + by simp + have "(\<Sum>x\<in>C. a x *\<^sub>R x) = (\<Sum>x\<in>B. a x *\<^sub>R x)" + proof - + have "C - D = C - B" + using Beq by blast + then show ?thesis + using Beq \<open>(\<Sum>x\<in>C - B. a x *\<^sub>R x) = 0\<close> f3 f4 by auto + qed + with 0 have Dcc0: "(\<Sum>v\<in>D. a v *\<^sub>R v) = 0" + apply (subst Deq) + by (simp add: \<open>finite B\<close> \<open>finite D\<close> setsum_Un) + then have D0: "\<And>v. v \<in> D \<Longrightarrow> a v = 0" + using independent_explicit \<open>independent D\<close> by blast + show ?thesis + using v C0 D0 Beq by blast + qed + then have "independent (C \<union> (D - C))" + by (simp add: independent_explicit \<open>finite C\<close> \<open>finite D\<close> setsum_Un del: Un_Diff_cancel) + then have indCUD: "independent (C \<union> D)" by simp + have "dim (S \<inter> T) = card B" + by (rule dim_unique [OF B indB refl]) + moreover have "dim S = card C" + by (metis \<open>C \<subseteq> S\<close> \<open>independent C\<close> \<open>S \<subseteq> span C\<close> basis_card_eq_dim) + moreover have "dim T = card D" + by (metis \<open>D \<subseteq> T\<close> \<open>independent D\<close> \<open>T \<subseteq> span D\<close> basis_card_eq_dim) + moreover have "dim {x + y |x y. x \<in> S \<and> y \<in> T} = card(C \<union> D)" + apply (rule dim_unique [OF CUD _ indCUD refl], clarify) + apply (meson \<open>S \<subseteq> span C\<close> \<open>T \<subseteq> span D\<close> span_add span_inc span_minimal subsetCE subspace_span sup.bounded_iff) + done + ultimately show ?thesis + using \<open>B = C \<inter> D\<close> [symmetric] + by (simp add: \<open>independent C\<close> \<open>independent D\<close> card_Un_Int independent_finite) +qed + + +lemma aff_dim_sums_Int_0: + assumes "affine S" + and "affine T" + and "0 \<in> S" "0 \<in> T" + shows "aff_dim {x + y| x y. x \<in> S \<and> y \<in> T} = (aff_dim S + aff_dim T) - aff_dim(S \<inter> T)" +proof - + have "0 \<in> {x + y |x y. x \<in> S \<and> y \<in> T}" + using assms by force + then have 0: "0 \<in> affine hull {x + y |x y. x \<in> S \<and> y \<in> 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 \<inter> T \<noteq> {}" + shows "aff_dim {x + y| x y. x \<in> S \<and> y \<in> T} = (aff_dim S + aff_dim T) - aff_dim(S \<inter> T)" +proof - + obtain a where a: "a \<in> S" "a \<in> T" using assms by force + have aff: "affine (op+ (-a) ` S)" "affine (op+ (-a) ` T)" + using assms by (auto simp: affine_translation [symmetric]) + have zero: "0 \<in> (op+ (-a) ` S)" "0 \<in> (op+ (-a) ` T)" + using a assms by auto + have [simp]: "{x + y |x y. x \<in> op + (- a) ` S \<and> y \<in> op + (- a) ` T} = + op + (- 2 *\<^sub>R a) ` {x + y| x y. x \<in> S \<and> y \<in> T}" + by (force simp: algebra_simps scaleR_2) + have [simp]: "op + (- a) ` S \<inter> op + (- a) ` T = op + (- a) ` (S \<inter> T)" + by auto + show ?thesis + using aff_dim_sums_Int_0 [OF aff zero] + by (auto simp: aff_dim_translation_eq) +qed + +lemma aff_dim_affine_Int_hyperplane: + fixes a :: "'a::euclidean_space" + assumes "affine S" + shows "aff_dim(S \<inter> {x. a \<bullet> x = b}) = + (if S \<inter> {v. a \<bullet> v = b} = {} then - 1 + else if S \<subseteq> {v. a \<bullet> 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 \<inter> {x. a \<bullet> x = b}) = aff_dim S - 1" + if "x \<in> S" "a \<bullet> x \<noteq> b" and non: "S \<inter> {v. a \<bullet> v = b} \<noteq> {}" for x + proof - + have [simp]: "{x + y| x y. x \<in> S \<and> a \<bullet> 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) \<longleftrightarrow> (affine hull S \<noteq> UNIV)" +by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) + +subsection\<open> Orthogonal bases, Gram-Schmidt process, and related theorems\<close> + +lemma span_delete_0 [simp]: "span(S - {0}) = span S" +proof + show "span (S - {0}) \<subseteq> span S" + by (blast intro!: span_mono) +next + have "span S \<subseteq> span(insert 0 (S - {0}))" + by (blast intro!: span_mono) + also have "... \<subseteq> span(S - {0})" + using span_insert_0 by blast + finally show "span S \<subseteq> span (S - {0})" . +qed + +lemma span_image_scale: + assumes "finite S" and nz: "\<And>x. x \<in> S \<Longrightarrow> c x \<noteq> 0" + shows "span ((\<lambda>x. c x *\<^sub>R x) ` S) = span S" +using assms +proof (induction S arbitrary: c) + case (empty c) show ?case by simp +next + case (insert x F c) + show ?case + proof (intro set_eqI iffI) + fix y + assume "y \<in> span ((\<lambda>x. c x *\<^sub>R x) ` insert x F)" + then show "y \<in> span (insert x F)" + using insert by (force simp: span_breakdown_eq) + next + fix y + assume "y \<in> span (insert x F)" + then show "y \<in> span ((\<lambda>x. c x *\<^sub>R x) ` insert x F)" + using insert + apply (clarsimp simp: span_breakdown_eq) + apply (rule_tac x="k / c x" in exI) + by simp + qed +qed + +lemma pairwise_orthogonal_independent: + assumes "pairwise orthogonal S" and "0 \<notin> S" + shows "independent S" +proof - + have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" + using assms by (simp add: pairwise_def orthogonal_def) + have "False" if "a \<in> S" and a: "a \<in> span (S - {a})" for a + proof - + obtain T U where "T \<subseteq> S - {a}" "a = (\<Sum>v\<in>T. U v *\<^sub>R v)" + using a by (force simp: span_explicit) + then have "a \<bullet> a = a \<bullet> (\<Sum>v\<in>T. U v *\<^sub>R v)" + by simp + also have "... = 0" + apply (simp add: inner_setsum_right) + apply (rule comm_monoid_add_class.setsum.neutral) + by (metis "0" DiffE \<open>T \<subseteq> S - {a}\<close> mult_not_zero singletonI subsetCE \<open>a \<in> S\<close>) + finally show ?thesis + using \<open>0 \<notin> S\<close> \<open>a \<in> S\<close> by auto + qed + then show ?thesis + by (force simp: dependent_def) +qed + +lemma pairwise_orthogonal_imp_finite: + fixes S :: "'a::euclidean_space set" + assumes "pairwise orthogonal S" + shows "finite S" +proof - + have "independent (S - {0})" + apply (rule pairwise_orthogonal_independent) + apply (metis Diff_iff assms pairwise_def) + by blast + then show ?thesis + by (meson independent_imp_finite infinite_remove) +qed + +lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}" + by (simp add: subspace_def orthogonal_clauses) + +lemma subspace_orthogonal_to_vectors: "subspace {y. \<forall>x \<in> S. orthogonal x y}" + by (simp add: subspace_def orthogonal_clauses) + +lemma orthogonal_to_span: + assumes a: "a \<in> span S" and x: "\<And>y. y \<in> S \<Longrightarrow> orthogonal x y" + shows "orthogonal x a" +apply (rule span_induct [OF a subspace_orthogonal_to_vector]) +apply (simp add: x) +done + +proposition Gram_Schmidt_step: + fixes S :: "'a::euclidean_space set" + assumes S: "pairwise orthogonal S" and x: "x \<in> span S" + shows "orthogonal x (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b))" +proof - + have "finite S" + by (simp add: S pairwise_orthogonal_imp_finite) + have "orthogonal (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)) x" + if "x \<in> S" for x + proof - + have "a \<bullet> x = (\<Sum>y\<in>S. if y = x then y \<bullet> a else 0)" + by (simp add: \<open>finite S\<close> inner_commute setsum.delta that) + also have "... = (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))" + apply (rule setsum.cong [OF refl], simp) + by (meson S orthogonal_def pairwise_def that) + finally show ?thesis + by (simp add: orthogonal_def algebra_simps inner_setsum_left) + qed + then show ?thesis + using orthogonal_to_span orthogonal_commute x by blast +qed + + +lemma orthogonal_extension_aux: + fixes S :: "'a::euclidean_space set" + assumes "finite T" "finite S" "pairwise orthogonal S" + shows "\<exists>U. pairwise orthogonal (S \<union> U) \<and> span (S \<union> U) = span (S \<union> T)" +using assms +proof (induction arbitrary: S) + case empty then show ?case + by simp (metis sup_bot_right) +next + case (insert a T) + have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" + using insert by (simp add: pairwise_def orthogonal_def) + define a' where "a' = a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)" + obtain U where orthU: "pairwise orthogonal (S \<union> insert a' U)" + and spanU: "span (insert a' S \<union> U) = span (insert a' S \<union> T)" + apply (rule exE [OF insert.IH [of "insert a' S"]]) + apply (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) + done + have orthS: "\<And>x. x \<in> S \<Longrightarrow> a' \<bullet> x = 0" + apply (simp add: a'_def) + using Gram_Schmidt_step [OF \<open>pairwise orthogonal S\<close>] + apply (force simp: orthogonal_def inner_commute span_inc [THEN subsetD]) + done + have "span (S \<union> insert a' U) = span (insert a' (S \<union> T))" + using spanU by simp + also have "... = span (insert a (S \<union> T))" + apply (rule eq_span_insert_eq) + apply (simp add: a'_def span_neg span_setsum span_clauses(1) span_mul) + done + also have "... = span (S \<union> insert a T)" + by simp + finally show ?case + apply (rule_tac x="insert a' U" in exI) + using orthU apply auto + done +qed + + +proposition orthogonal_extension: + fixes S :: "'a::euclidean_space set" + assumes S: "pairwise orthogonal S" + obtains U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)" +proof - + obtain B where "finite B" "span B = span T" + using basis_subspace_exists [of "span T"] subspace_span by auto + with orthogonal_extension_aux [of B S] + obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> B)" + using assms pairwise_orthogonal_imp_finite by auto + show ?thesis + apply (rule_tac U=U in that) + apply (simp add: \<open>pairwise orthogonal (S \<union> U)\<close>) + by (metis \<open>span (S \<union> U) = span (S \<union> B)\<close> \<open>span B = span T\<close> span_Un) +qed + +corollary orthogonal_extension_strong: + fixes S :: "'a::euclidean_space set" + assumes S: "pairwise orthogonal S" + obtains U where "U \<inter> (insert 0 S) = {}" "pairwise orthogonal (S \<union> U)" + "span (S \<union> U) = span (S \<union> T)" +proof - + obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)" + using orthogonal_extension assms by blast + then show ?thesis + apply (rule_tac U = "U - (insert 0 S)" in that) + apply blast + apply (force simp: pairwise_def) + apply (metis (no_types, lifting) Un_Diff_cancel span_insert_0 span_Un) + done +qed + +subsection\<open>Decomposing a vector into parts in orthogonal subspaces.\<close> + +text\<open>existence of orthonormal basis for a subspace.\<close> + +lemma orthogonal_spanningset_subspace: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" + obtains B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S" +proof - + obtain B where "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" + using basis_exists by blast + with orthogonal_extension [of "{}" B] + show ?thesis + by (metis Un_empty_left assms pairwise_empty span_inc span_subspace that) +qed + +lemma orthogonal_basis_subspace: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" + obtains B where "0 \<notin> B" "B \<subseteq> S" "pairwise orthogonal B" "independent B" + "card B = dim S" "span B = S" +proof - + obtain B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S" + using assms orthogonal_spanningset_subspace by blast + then show ?thesis + apply (rule_tac B = "B - {0}" in that) + apply (auto simp: indep_card_eq_dim_span pairwise_subset Diff_subset pairwise_orthogonal_independent elim: pairwise_subset) + done +qed + +proposition orthonormal_basis_subspace: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" + obtains B where "B \<subseteq> S" "pairwise orthogonal B" + and "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" + and "independent B" "card B = dim S" "span B = S" +proof - + obtain B where "0 \<notin> B" "B \<subseteq> S" + and orth: "pairwise orthogonal B" + and "independent B" "card B = dim S" "span B = S" + by (blast intro: orthogonal_basis_subspace [OF assms]) + have 1: "(\<lambda>x. x /\<^sub>R norm x) ` B \<subseteq> S" + using \<open>span B = S\<close> span_clauses(1) span_mul by fastforce + have 2: "pairwise orthogonal ((\<lambda>x. x /\<^sub>R norm x) ` B)" + using orth by (force simp: pairwise_def orthogonal_clauses) + have 3: "\<And>x. x \<in> (\<lambda>x. x /\<^sub>R norm x) ` B \<Longrightarrow> norm x = 1" + by (metis (no_types, lifting) \<open>0 \<notin> B\<close> image_iff norm_sgn sgn_div_norm) + have 4: "independent ((\<lambda>x. x /\<^sub>R norm x) ` B)" + by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) + have "inj_on (\<lambda>x. x /\<^sub>R norm x) B" + proof + fix x y + assume "x \<in> B" "y \<in> B" "x /\<^sub>R norm x = y /\<^sub>R norm y" + moreover have "\<And>i. i \<in> B \<Longrightarrow> norm (i /\<^sub>R norm i) = 1" + using 3 by blast + ultimately show "x = y" + by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) + qed + then have 5: "card ((\<lambda>x. x /\<^sub>R norm x) ` B) = dim S" + by (metis \<open>card B = dim S\<close> card_image) + have 6: "span ((\<lambda>x. x /\<^sub>R norm x) ` B) = S" + by (metis "1" "4" "5" assms card_eq_dim independent_finite span_subspace) + show ?thesis + by (rule that [OF 1 2 3 4 5 6]) +qed + +proposition orthogonal_subspace_decomp_exists: + fixes S :: "'a :: euclidean_space set" + obtains y z where "y \<in> span S" "\<And>w. w \<in> span S \<Longrightarrow> orthogonal z w" "x = y + z" +proof - + obtain T where "0 \<notin> T" "T \<subseteq> span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" + using orthogonal_basis_subspace subspace_span by blast + let ?a = "\<Sum>b\<in>T. (b \<bullet> x / (b \<bullet> b)) *\<^sub>R b" + have orth: "orthogonal (x - ?a) w" if "w \<in> span S" for w + by (simp add: Gram_Schmidt_step \<open>pairwise orthogonal T\<close> \<open>span T = span S\<close> orthogonal_commute that) + show ?thesis + apply (rule_tac y = "?a" and z = "x - ?a" in that) + apply (meson \<open>T \<subseteq> span S\<close> span_mul span_setsum subsetCE) + apply (fact orth, simp) + done +qed + +lemma orthogonal_subspace_decomp_unique: + fixes S :: "'a :: euclidean_space set" + assumes "x + y = x' + y'" + and ST: "x \<in> span S" "x' \<in> span S" "y \<in> span T" "y' \<in> span T" + and orth: "\<And>a b. \<lbrakk>a \<in> S; b \<in> T\<rbrakk> \<Longrightarrow> orthogonal a b" + shows "x = x' \<and> y = y'" +proof - + have "x + y - y' = x'" + by (simp add: assms) + moreover have "\<And>a b. \<lbrakk>a \<in> span S; b \<in> span T\<rbrakk> \<Longrightarrow> orthogonal a b" + by (meson orth orthogonal_commute orthogonal_to_span) + ultimately have "0 = x' - x" + by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) + with assms show ?thesis by auto +qed + +proposition dim_orthogonal_sum: + fixes A :: "'a::euclidean_space set" + assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" + shows "dim(A \<union> B) = dim A + dim B" +proof - + have 1: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" + by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) + have "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" + apply (erule span_induct [OF _ subspace_hyperplane]) + using 1 by (simp add: ) + then have 0: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" + by simp + have "dim(A \<union> B) = dim (span (A \<union> B))" + by simp + also have "... = dim ((\<lambda>(a, b). a + b) ` (span A \<times> span B))" + by (simp add: span_Un) + also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B}" + by (auto intro!: arg_cong [where f=dim]) + also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B} + dim(span A \<inter> span B)" + by (auto simp: dest: 0) + also have "... = dim (span A) + dim (span B)" + by (rule dim_sums_Int) auto + also have "... = dim A + dim B" + by simp + finally show ?thesis . +qed + +lemma dim_subspace_orthogonal_to_vectors: + fixes A :: "'a::euclidean_space set" + assumes "subspace A" "subspace B" "A \<subseteq> B" + shows "dim {y \<in> B. \<forall>x \<in> A. orthogonal x y} + dim A = dim B" +proof - + have "dim (span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)) = dim (span B)" + proof (rule arg_cong [where f=dim, OF subset_antisym]) + show "span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A) \<subseteq> span B" + by (simp add: \<open>A \<subseteq> B\<close> Collect_restrict span_mono) + next + have *: "x \<in> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)" + if "x \<in> B" for x + proof - + obtain y z where "x = y + z" "y \<in> span A" and orth: "\<And>w. w \<in> span A \<Longrightarrow> orthogonal z w" + using orthogonal_subspace_decomp_exists [of A x] that by auto + have "y \<in> span B" + by (metis span_eq \<open>y \<in> span A\<close> assms subset_iff) + then have "z \<in> {a \<in> B. \<forall>x. x \<in> A \<longrightarrow> orthogonal x a}" + by simp (metis (no_types) span_eq \<open>x = y + z\<close> \<open>subspace A\<close> \<open>subspace B\<close> orth orthogonal_commute span_add_eq that) + then have z: "z \<in> span {y \<in> B. \<forall>x\<in>A. orthogonal x y}" + by (meson span_inc subset_iff) + then show ?thesis + apply (simp add: span_Un image_def) + apply (rule bexI [OF _ z]) + apply (simp add: \<open>x = y + z\<close> \<open>y \<in> span A\<close>) + done + qed + show "span B \<subseteq> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)" + by (rule span_minimal) (auto intro: * span_minimal elim: ) + qed + then show ?thesis + by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) +qed + +subsection\<open>Parallel slices, etc.\<close> + +text\<open> 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.\<close> + +proposition affine_parallel_slice: + fixes S :: "'a :: euclidean_space set" + assumes "affine S" + and "S \<inter> {x. a \<bullet> x \<le> b} \<noteq> {}" + and "~ (S \<subseteq> {x. a \<bullet> x \<le> b})" + obtains a' b' where "a' \<noteq> 0" + "S \<inter> {x. a' \<bullet> x \<le> b'} = S \<inter> {x. a \<bullet> x \<le> b}" + "S \<inter> {x. a' \<bullet> x = b'} = S \<inter> {x. a \<bullet> x = b}" + "\<And>w. w \<in> S \<Longrightarrow> (w + a') \<in> S" +proof (cases "S \<inter> {x. a \<bullet> x = b} = {}") + case True + then obtain u v where "u \<in> S" "v \<in> S" "a \<bullet> u \<le> b" "a \<bullet> v > b" + using assms by (auto simp: not_le) + define \<eta> where "\<eta> = u + ((b - a \<bullet> u) / (a \<bullet> v - a \<bullet> u)) *\<^sub>R (v - u)" + have "\<eta> \<in> S" + by (simp add: \<eta>_def \<open>u \<in> S\<close> \<open>v \<in> S\<close> \<open>affine S\<close> mem_affine_3_minus) + moreover have "a \<bullet> \<eta> = b" + using \<open>a \<bullet> u \<le> b\<close> \<open>b < a \<bullet> v\<close> + by (simp add: \<eta>_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 \<in> S" and z: "a \<bullet> z = b" + using assms by auto + with affine_diffs_subspace [OF \<open>affine S\<close>] + have sub: "subspace (op + (- z) ` S)" by blast + then have aff: "affine (op + (- z) ` S)" and span: "span (op + (- z) ` S) = (op + (- z) ` S)" + by (auto simp: subspace_imp_affine) + obtain a' a'' where a': "a' \<in> span (op + (- z) ` S)" and a: "a = a' + a''" + and "\<And>w. w \<in> span (op + (- z) ` S) \<Longrightarrow> orthogonal a'' w" + using orthogonal_subspace_decomp_exists [of "op + (- z) ` S" "a"] by metis + then have "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> (w-z) = 0" + by (simp add: imageI orthogonal_def span) + then have a'': "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> w = (a - a') \<bullet> z" + by (simp add: a inner_diff_right) + then have ba'': "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> w = b - a' \<bullet> z" + by (simp add: inner_diff_left z) + have "\<And>w. w \<in> op + (- z) ` S \<Longrightarrow> (w + a') \<in> op + (- z) ` S" + by (metis subspace_add a' span_eq sub) + then have Sclo: "\<And>w. w \<in> S \<Longrightarrow> (w + a') \<in> S" + by fastforce + 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 + apply (rule_tac a' = "a'" and b' = "a' \<bullet> z" in that) + apply (auto simp: a ba'' inner_left_distrib False Sclo) + done + qed +qed + +lemma diffs_affine_hull_span: + assumes "a \<in> S" + shows "{x - a |x. x \<in> affine hull S} = span {x - a |x. x \<in> S}" +proof - + have *: "((\<lambda>x. x - a) ` (S - {a})) = {x. x + a \<in> S} - {0}" + by (auto simp: algebra_simps) + show ?thesis + apply (simp add: affine_hull_span2 [OF assms] *) + apply (auto simp: algebra_simps) + done +qed + +lemma aff_dim_dim_affine_diffs: + fixes S :: "'a :: euclidean_space set" + assumes "affine S" "a \<in> S" + shows "aff_dim S = dim {x - a |x. x \<in> S}" +proof - + obtain B where aff: "affine hull B = affine hull S" + and ind: "\<not> affine_dependent B" + and card: "of_nat (card B) = aff_dim S + 1" + using aff_dim_basis_exists by blast + then have "B \<noteq> {}" using assms + by (metis affine_hull_eq_empty ex_in_conv) + then obtain c where "c \<in> B" by auto + then have "c \<in> S" + by (metis aff affine_hull_eq \<open>affine S\<close> hull_inc) + have xy: "x - c = y - a \<longleftrightarrow> y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a + by (auto simp: algebra_simps) + have *: "{x - c |x. x \<in> S} = {x - a |x. x \<in> S}" + apply safe + apply (simp_all only: xy) + using mem_affine_3_minus [OF \<open>affine S\<close>] \<open>a \<in> S\<close> \<open>c \<in> S\<close> apply blast+ + done + have affS: "affine hull S = S" + by (simp add: \<open>affine S\<close>) + have "aff_dim S = of_nat (card B) - 1" + using card by simp + also have "... = dim {x - c |x. x \<in> B}" + by (simp add: affine_independent_card_dim_diffs [OF ind \<open>c \<in> B\<close>]) + also have "... = dim {x - c | x. x \<in> affine hull B}" + by (simp add: diffs_affine_hull_span \<open>c \<in> B\<close>) + also have "... = dim {x - a |x. x \<in> S}" + by (simp add: affS aff *) + finally show ?thesis . +qed + +lemma aff_dim_linear_image_le: + assumes "linear f" + shows "aff_dim(f ` S) \<le> aff_dim S" +proof - + have "aff_dim (f ` T) \<le> 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 \<in> T" by auto + have 1: "((\<lambda>x. x - f a) ` f ` T) = {x - f a |x. x \<in> f ` T}" + by auto + have 2: "{x - f a| x. x \<in> f ` T} = f ` {x - a| x. x \<in> T}" + by (force simp: linear_diff [OF assms]) + have "aff_dim (f ` T) = int (dim {x - f a |x. x \<in> f ` T})" + by (simp add: \<open>a \<in> T\<close> hull_inc aff_dim_eq_dim [of "f a"] 1) + also have "... = int (dim (f ` {x - a| x. x \<in> T}))" + by (force simp: linear_diff [OF assms] 2) + also have "... \<le> int (dim {x - a| x. x \<in> T})" + by (simp add: dim_image_le [OF assms]) + also have "... \<le> aff_dim T" + by (simp add: aff_dim_dim_affine_diffs [symmetric] \<open>a \<in> T\<close> \<open>affine T\<close>) + finally show ?thesis . + qed + then + have "aff_dim (f ` (affine hull S)) \<le> 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) \<le> aff_dim S" + by (simp add: aff_dim_linear_image_le assms(1)) +next + obtain g where "linear g" "g \<circ> f = id" + using linear_injective_left_inverse assms by blast + then have "aff_dim S \<le> aff_dim(g ` f ` S)" + by (simp add: image_comp) + also have "... \<le> aff_dim (f ` S)" + by (simp add: \<open>linear g\<close> aff_dim_linear_image_le) + finally show "aff_dim S \<le> aff_dim (f ` S)" . +qed + +text\<open>Choosing a subspace of a given dimension\<close> +proposition choose_subspace_of_subspace: + fixes S :: "'n::euclidean_space set" + assumes "n \<le> dim S" + obtains T where "subspace T" "T \<subseteq> span S" "dim T = n" +proof - + have "\<exists>T. subspace T \<and> T \<subseteq> span S \<and> dim T = n" + using assms + proof (induction n) + case 0 then show ?case by force + next + case (Suc n) + then obtain T where "subspace T" "T \<subseteq> span S" "dim T = n" + by force + then show ?case + proof (cases "span S \<subseteq> span T") + case True + have "dim S = dim T" + apply (rule span_eq_dim [OF subset_antisym [OF True]]) + by (simp add: \<open>T \<subseteq> span S\<close> span_minimal subspace_span) + then show ?thesis + using Suc.prems \<open>dim T = n\<close> by linarith + next + case False + then obtain y where y: "y \<in> S" "y \<notin> T" + by (meson span_mono subsetI) + then have "span (insert y T) \<subseteq> span S" + by (metis (no_types) \<open>T \<subseteq> span S\<close> subsetD insert_subset span_inc span_mono span_span) + with \<open>dim T = n\<close> \<open>subspace T\<close> y show ?thesis + apply (rule_tac x="span(insert y T)" in exI) + apply (auto simp: dim_insert) + using span_eq by blast + qed + qed + with that show ?thesis by blast +qed + +lemma choose_affine_subset: + assumes "affine S" "-1 \<le> d" and dle: "d \<le> aff_dim S" + obtains T where "affine T" "T \<subseteq> S" "aff_dim T = d" +proof (cases "d = -1 \<or> 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 \<in> S" "0 \<le> d" by auto + with assms have ss: "subspace (op + (- a) ` S)" + by (simp add: affine_diffs_subspace) + have "nat d \<le> dim (op + (- 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 \<subseteq> span (op + (- a) ` S)" + and Tdim: "dim T = nat d" + using choose_subspace_of_subspace [of "nat d" "op + (- a) ` S"] by blast + then have "affine T" + using subspace_affine by blast + then have "affine (op + a ` T)" + by (metis affine_hull_eq affine_hull_translation) + moreover have "op + a ` T \<subseteq> S" + proof - + have "T \<subseteq> op + (- a) ` S" + by (metis (no_types) span_eq Tsb ss) + then show "op + a ` T \<subseteq> S" + using add_ac by auto + qed + moreover have "aff_dim (op + a ` T) = d" + by (simp add: aff_dim_subspace Tdim \<open>0 \<le> d\<close> \<open>subspace T\<close> aff_dim_translation_eq) + ultimately show ?thesis + by (rule that) +qed + +subsection\<open>Several Variants of Paracompactness\<close> + +proposition paracompact: + fixes S :: "'a :: euclidean_space set" + assumes "S \<subseteq> \<Union>\<C>" and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T" + obtains \<C>' where "S \<subseteq> \<Union> \<C>'" + and "\<And>U. U \<in> \<C>' \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<C> \<and> U \<subseteq> T)" + and "\<And>x. x \<in> S + \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> + finite {U. U \<in> \<C>' \<and> (U \<inter> V \<noteq> {})}" +proof (cases "S = {}") + case True with that show ?thesis by blast +next + case False + have "\<exists>T U. x \<in> U \<and> open U \<and> closure U \<subseteq> T \<and> T \<in> \<C>" if "x \<in> S" for x + proof - + obtain T where "x \<in> T" "T \<in> \<C>" "open T" + using assms \<open>x \<in> S\<close> by blast + then obtain e where "e > 0" "cball x e \<subseteq> T" + by (force simp: open_contains_cball) + then show ?thesis + apply (rule_tac x = T in exI) + apply (rule_tac x = "ball x e" in exI) + using \<open>T \<in> \<C>\<close> + apply (simp add: closure_minimal) + done + qed + then obtain F G where Gin: "x \<in> G x" and oG: "open (G x)" + and clos: "closure (G x) \<subseteq> F x" and Fin: "F x \<in> \<C>" + if "x \<in> S" for x + by metis + then obtain \<F> where "\<F> \<subseteq> G ` S" "countable \<F>" "\<Union>\<F> = UNION S G" + using Lindelof [of "G ` S"] by (metis image_iff) + then obtain K where K: "K \<subseteq> S" "countable K" and eq: "UNION K G = UNION S G" + by (metis countable_subset_image) + with False Gin have "K \<noteq> {}" by force + then obtain a :: "nat \<Rightarrow> 'a" where "range a = K" + by (metis range_from_nat_into \<open>countable K\<close>) + then have odif: "\<And>n. open (F (a n) - \<Union>{closure (G (a m)) |m. m < n})" + using \<open>K \<subseteq> S\<close> Fin opC by (fastforce simp add:) + let ?C = "range (\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n})" + have enum_S: "\<exists>n. x \<in> F(a n) \<and> x \<in> G(a n)" if "x \<in> S" for x + proof - + have "\<exists>y \<in> K. x \<in> G y" using eq that Gin by fastforce + then show ?thesis + using clos K \<open>range a = K\<close> closure_subset by blast + qed + have 1: "S \<subseteq> Union ?C" + proof + fix x assume "x \<in> S" + define n where "n \<equiv> LEAST n. x \<in> F(a n)" + have n: "x \<in> F(a n)" + using enum_S [OF \<open>x \<in> S\<close>] by (force simp: n_def intro: LeastI) + have notn: "x \<notin> F(a m)" if "m < n" for m + using that not_less_Least by (force simp: n_def) + then have "x \<notin> \<Union>{closure (G (a m)) |m. m < n}" + using n \<open>K \<subseteq> S\<close> \<open>range a = K\<close> clos notn by fastforce + with n show "x \<in> Union ?C" + by blast + qed + have 3: "\<exists>V. open V \<and> x \<in> V \<and> finite {U. U \<in> ?C \<and> (U \<inter> V \<noteq> {})}" if "x \<in> S" for x + proof - + obtain n where n: "x \<in> F(a n)" "x \<in> G(a n)" + using \<open>x \<in> S\<close> enum_S by auto + have "{U \<in> ?C. U \<inter> G (a n) \<noteq> {}} \<subseteq> (\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n}) ` atMost n" + proof clarsimp + fix k assume "(F (a k) - \<Union>{closure (G (a m)) |m. m < k}) \<inter> G (a n) \<noteq> {}" + then have "k \<le> n" + by auto (metis closure_subset not_le subsetCE) + then show "F (a k) - \<Union>{closure (G (a m)) |m. m < k} + \<in> (\<lambda>n. F (a n) - \<Union>{closure (G (a m)) |m. m < n}) ` {..n}" + by force + qed + moreover have "finite ((\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n}) ` atMost n)" + by force + ultimately have *: "finite {U \<in> ?C. U \<inter> G (a n) \<noteq> {}}" + using finite_subset by blast + show ?thesis + apply (rule_tac x="G (a n)" in exI) + apply (intro conjI oG n *) + using \<open>K \<subseteq> S\<close> \<open>range a = K\<close> apply blast + done + qed + show ?thesis + apply (rule that [OF 1 _ 3]) + using Fin \<open>K \<subseteq> S\<close> \<open>range a = K\<close> apply (auto simp: odif) + done +qed + + +corollary paracompact_closedin: + fixes S :: "'a :: euclidean_space set" + assumes cin: "closedin (subtopology euclidean U) S" + and oin: "\<And>T. T \<in> \<C> \<Longrightarrow> openin (subtopology euclidean U) T" + and "S \<subseteq> \<Union>\<C>" + obtains \<C>' where "S \<subseteq> \<Union> \<C>'" + and "\<And>V. V \<in> \<C>' \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)" + and "\<And>x. x \<in> U + \<Longrightarrow> \<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and> + finite {X. X \<in> \<C>' \<and> (X \<inter> V \<noteq> {})}" +proof - + have "\<exists>Z. open Z \<and> (T = U \<inter> Z)" if "T \<in> \<C>" for T + using oin [OF that] by (auto simp: openin_open) + then obtain F where opF: "open (F T)" and intF: "U \<inter> F T = T" if "T \<in> \<C>" for T + by metis + obtain K where K: "closed K" "U \<inter> K = S" + using cin by (auto simp: closedin_closed) + have 1: "U \<subseteq> \<Union>insert (- K) (F ` \<C>)" + by clarsimp (metis Int_iff Union_iff \<open>U \<inter> K = S\<close> \<open>S \<subseteq> \<Union>\<C>\<close> subsetD intF) + have 2: "\<And>T. T \<in> insert (- K) (F ` \<C>) \<Longrightarrow> open T" + using \<open>closed K\<close> by (auto simp: opF) + obtain \<D> where "U \<subseteq> \<Union>\<D>" + and D1: "\<And>U. U \<in> \<D> \<Longrightarrow> open U \<and> (\<exists>T. T \<in> insert (- K) (F ` \<C>) \<and> U \<subseteq> T)" + and D2: "\<And>x. x \<in> U \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<D>. U \<inter> V \<noteq> {}}" + using paracompact [OF 1 2] by auto + let ?C = "{U \<inter> V |V. V \<in> \<D> \<and> (V \<inter> K \<noteq> {})}" + show ?thesis + proof (rule_tac \<C>' = "{U \<inter> V |V. V \<in> \<D> \<and> (V \<inter> K \<noteq> {})}" in that) + show "S \<subseteq> \<Union>?C" + using \<open>U \<inter> K = S\<close> \<open>U \<subseteq> \<Union>\<D>\<close> K by (blast dest!: subsetD) + show "\<And>V. V \<in> ?C \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)" + using D1 intF by fastforce + have *: "{X. (\<exists>V. X = U \<inter> V \<and> V \<in> \<D> \<and> V \<inter> K \<noteq> {}) \<and> X \<inter> (U \<inter> V) \<noteq> {}} \<subseteq> + (\<lambda>x. U \<inter> x) ` {U \<in> \<D>. U \<inter> V \<noteq> {}}" for V + by blast + show "\<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and> finite {X \<in> ?C. X \<inter> V \<noteq> {}}" + if "x \<in> U" for x + using D2 [OF that] + apply clarify + apply (rule_tac x="U \<inter> V" in exI) + apply (auto intro: that finite_subset [OF *]) + done + qed +qed + +corollary paracompact_closed: + fixes S :: "'a :: euclidean_space set" + assumes "closed S" + and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T" + and "S \<subseteq> \<Union>\<C>" + obtains \<C>' where "S \<subseteq> \<Union>\<C>'" + and "\<And>U. U \<in> \<C>' \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<C> \<and> U \<subseteq> T)" + and "\<And>x. \<exists>V. open V \<and> x \<in> V \<and> + finite {U. U \<in> \<C>' \<and> (U \<inter> V \<noteq> {})}" +using paracompact_closedin [of UNIV S \<C>] assms +by (auto simp: open_openin [symmetric] closed_closedin [symmetric]) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Derivative.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,2910 @@ +(* Title: HOL/Analysis/Derivative.thy + Author: John Harrison + Author: Robert Himmelmann, TU Muenchen (translation from HOL Light) +*) + +section \<open>Multivariate calculus in Euclidean space\<close> + +theory Derivative +imports Brouwer_Fixpoint Operator_Norm Uniform_Limit Bounded_Linear_Function +begin + +lemma onorm_inner_left: + assumes "bounded_linear r" + shows "onorm (\<lambda>x. r x \<bullet> f) \<le> onorm r * norm f" +proof (rule onorm_bound) + fix x + have "norm (r x \<bullet> f) \<le> norm (r x) * norm f" + by (simp add: Cauchy_Schwarz_ineq2) + also have "\<dots> \<le> onorm r * norm x * norm f" + by (intro mult_right_mono onorm assms norm_ge_zero) + finally show "norm (r x \<bullet> f) \<le> onorm r * norm f * norm x" + by (simp add: ac_simps) +qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le assms) + +lemma onorm_inner_right: + assumes "bounded_linear r" + shows "onorm (\<lambda>x. f \<bullet> r x) \<le> norm f * onorm r" + apply (subst inner_commute) + apply (rule onorm_inner_left[OF assms, THEN order_trans]) + apply simp + done + +declare has_derivative_bounded_linear[dest] + +subsection \<open>Derivatives\<close> + +subsubsection \<open>Combining theorems.\<close> + +lemmas has_derivative_id = has_derivative_ident +lemmas has_derivative_neg = has_derivative_minus +lemmas has_derivative_sub = has_derivative_diff +lemmas scaleR_right_has_derivative = has_derivative_scaleR_right +lemmas scaleR_left_has_derivative = has_derivative_scaleR_left +lemmas inner_right_has_derivative = has_derivative_inner_right +lemmas inner_left_has_derivative = has_derivative_inner_left +lemmas mult_right_has_derivative = has_derivative_mult_right +lemmas mult_left_has_derivative = has_derivative_mult_left + +lemma has_derivative_add_const: + "(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. f x + c) has_derivative f') net" + by (intro derivative_eq_intros) auto + + +subsection \<open>Derivative with composed bilinear function.\<close> + +lemma has_derivative_bilinear_within: + assumes "(f has_derivative f') (at x within s)" + and "(g has_derivative g') (at x within s)" + and "bounded_bilinear h" + shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x within s)" + using bounded_bilinear.FDERIV[OF assms(3,1,2)] . + +lemma has_derivative_bilinear_at: + assumes "(f has_derivative f') (at x)" + and "(g has_derivative g') (at x)" + and "bounded_bilinear h" + shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x)" + using has_derivative_bilinear_within[of f f' x UNIV g g' h] assms by simp + +text \<open>These are the only cases we'll care about, probably.\<close> + +lemma has_derivative_within: "(f has_derivative f') (at x within s) \<longleftrightarrow> + bounded_linear f' \<and> ((\<lambda>y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \<longlongrightarrow> 0) (at x within s)" + unfolding has_derivative_def Lim + by (auto simp add: netlimit_within field_simps) + +lemma has_derivative_at: "(f has_derivative f') (at x) \<longleftrightarrow> + bounded_linear f' \<and> ((\<lambda>y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \<longlongrightarrow> 0) (at x)" + using has_derivative_within [of f f' x UNIV] + by simp + +text \<open>More explicit epsilon-delta forms.\<close> + +lemma has_derivative_within': + "(f has_derivative f')(at x within s) \<longleftrightarrow> + bounded_linear f' \<and> + (\<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. 0 < norm (x' - x) \<and> norm (x' - x) < d \<longrightarrow> + norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" + unfolding has_derivative_within Lim_within dist_norm + unfolding diff_0_right + by (simp add: diff_diff_eq) + +lemma has_derivative_at': + "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and> + (\<forall>e>0. \<exists>d>0. \<forall>x'. 0 < norm (x' - x) \<and> norm (x' - x) < d \<longrightarrow> + 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_at_within: + "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f') (at x within s)" + unfolding has_derivative_within' has_derivative_at' + by blast + +lemma has_derivative_within_open: + "a \<in> s \<Longrightarrow> open s \<Longrightarrow> + (f has_derivative f') (at a within s) \<longleftrightarrow> (f has_derivative f') (at a)" + by (simp only: at_within_interior interior_open) + +lemma has_derivative_right: + fixes f :: "real \<Rightarrow> real" + and y :: "real" + shows "(f has_derivative (op * y)) (at x within ({x <..} \<inter> I)) \<longleftrightarrow> + ((\<lambda>t. (f x - f t) / (x - t)) \<longlongrightarrow> y) (at x within ({x <..} \<inter> I))" +proof - + have "((\<lambda>t. (f t - (f x + y * (t - x))) / \<bar>t - x\<bar>) \<longlongrightarrow> 0) (at x within ({x<..} \<inter> I)) \<longleftrightarrow> + ((\<lambda>t. (f t - f x) / (t - x) - y) \<longlongrightarrow> 0) (at x within ({x<..} \<inter> I))" + by (intro Lim_cong_within) (auto simp add: diff_divide_distrib add_divide_distrib) + also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f t - f x) / (t - x)) \<longlongrightarrow> y) (at x within ({x<..} \<inter> I))" + by (simp add: Lim_null[symmetric]) + also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f x - f t) / (x - t)) \<longlongrightarrow> y) (at x within ({x<..} \<inter> 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 \<open>Caratheodory characterization\<close> + +lemmas DERIV_within_iff = has_field_derivative_iff + +lemma DERIV_caratheodory_within: + "(f has_field_derivative l) (at x within s) \<longleftrightarrow> + (\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> continuous (at x within s) g \<and> 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 "\<forall>z. f z - f x = ?g z * (z-x)" by simp + show "continuous (at x within s) ?g" using \<open>?lhs\<close> + by (auto simp add: continuous_within DERIV_within_iff cong: Lim_cong_within) + show "?g x = l" by simp + qed +next + assume ?rhs + then obtain g where + "(\<forall>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 DERIV_within_iff cong: Lim_cong_within) +qed + +subsubsection \<open>Limit transformation for derivatives\<close> + +lemma has_derivative_transform_within: + assumes "(f has_derivative f') (at x within s)" + and "0 < d" + and "x \<in> s" + and "\<And>x'. \<lbrakk>x' \<in> s; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'" + shows "(g has_derivative f') (at x within s)" + using assms + unfolding has_derivative_within + by (force simp add: intro: Lim_transform_within) + +lemma has_derivative_transform_within_open: + assumes "(f has_derivative f') (at x)" + and "open s" + and "x \<in> s" + and "\<And>x. x\<in>s \<Longrightarrow> f x = g x" + shows "(g has_derivative f') (at x)" + using assms unfolding has_derivative_at + by (force simp add: intro: Lim_transform_within_open) + +subsection \<open>Differentiability\<close> + +definition + differentiable_on :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool" + (infix "differentiable'_on" 50) + where "f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable (at x within s))" + +lemma differentiableI: "(f has_derivative f') net \<Longrightarrow> f differentiable net" + unfolding differentiable_def + by auto + +lemma differentiable_onD: "\<lbrakk>f differentiable_on S; x \<in> S\<rbrakk> \<Longrightarrow> f differentiable (at x within S)" + using differentiable_on_def by blast + +lemma differentiable_at_withinI: "f differentiable (at x) \<Longrightarrow> f differentiable (at x within s)" + unfolding differentiable_def + using has_derivative_at_within + by blast + +lemma differentiable_at_imp_differentiable_on: + "(\<And>x. x \<in> s \<Longrightarrow> f differentiable at x) \<Longrightarrow> f differentiable_on s" + by (metis differentiable_at_withinI differentiable_on_def) + +corollary differentiable_iff_scaleR: + fixes f :: "real \<Rightarrow> 'a::real_normed_vector" + shows "f differentiable F \<longleftrightarrow> (\<exists>d. (f has_derivative (\<lambda>x. x *\<^sub>R d)) F)" + by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR) + +lemma differentiable_within_open: (* TODO: delete *) + assumes "a \<in> s" + and "open s" + shows "f differentiable (at a within s) \<longleftrightarrow> f differentiable (at a)" + using assms + by (simp only: at_within_interior interior_open) + +lemma differentiable_on_eq_differentiable_at: + "open s \<Longrightarrow> f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>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 \<in> s" + and "\<And>x'. \<lbrakk>x'\<in>s; dist x' x < d\<rbrakk> \<Longrightarrow> 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]: "(\<lambda>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_compose: + "\<lbrakk>g differentiable_on S; f differentiable_on (g ` S)\<rbrakk> \<Longrightarrow> (\<lambda>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 \<Longrightarrow> f differentiable_on S" + by (simp add: differentiable_on_def bounded_linear_imp_differentiable) + +lemma linear_imp_differentiable_on: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + shows "linear f \<Longrightarrow> f differentiable_on S" +by (simp add: differentiable_on_def linear_imp_differentiable) + +lemma differentiable_on_minus [simp, derivative_intros]: + "f differentiable_on S \<Longrightarrow> (\<lambda>z. -(f z)) differentiable_on S" +by (simp add: differentiable_on_def) + +lemma differentiable_on_add [simp, derivative_intros]: + "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) differentiable_on S" +by (simp add: differentiable_on_def) + +lemma differentiable_on_diff [simp, derivative_intros]: + "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>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 \<Rightarrow> 'b :: real_normed_field" + shows "f differentiable_on S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> f x \<noteq> 0) \<Longrightarrow> (\<lambda>x. inverse (f x)) differentiable_on S" +by (simp add: differentiable_on_def) + +lemma differentiable_on_scaleR [derivative_intros, simp]: + "\<lbrakk>f differentiable_on S; g differentiable_on S\<rbrakk> \<Longrightarrow> (\<lambda>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]: + "((\<lambda>x. (norm x)\<^sup>2) has_derivative (\<lambda>x. 2 *\<^sub>R (a \<bullet> x))) (at a)" +using has_derivative_bilinear_at [of id id a id id "op \<bullet>"] +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 "(\<lambda>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 "(\<lambda>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 \<noteq> 0 \<Longrightarrow> 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 \<notin> S \<Longrightarrow> norm differentiable_on S" +by (metis differentiable_at_imp_differentiable_on differentiable_norm_at) + + +subsection \<open>Frechet derivative and Jacobian matrix\<close> + +definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)" + +lemma frechet_derivative_works: + "f differentiable net \<longleftrightarrow> (f has_derivative (frechet_derivative f net)) net" + unfolding frechet_derivative_def differentiable_def + unfolding some_eq_ex[of "\<lambda> f' . (f has_derivative f') net"] .. + +lemma linear_frechet_derivative: "f differentiable net \<Longrightarrow> linear (frechet_derivative f net)" + unfolding frechet_derivative_works has_derivative_def + by (auto intro: bounded_linear.linear) + + +subsection \<open>Differentiability implies continuity\<close> + +lemma differentiable_imp_continuous_within: + "f differentiable (at x within s) \<Longrightarrow> continuous (at x within s) f" + by (auto simp: differentiable_def intro: has_derivative_continuous) + +lemma differentiable_imp_continuous_on: + "f differentiable_on s \<Longrightarrow> 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 \<Longrightarrow> s \<subseteq> t \<Longrightarrow> 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 + +text \<open>Results about neighborhoods filter.\<close> + +lemma eventually_nhds_metric_le: + "eventually P (nhds a) = (\<exists>d>0. \<forall>x. dist x a \<le> d \<longrightarrow> P x)" + unfolding eventually_nhds_metric by (safe, rule_tac x="d / 2" in exI, auto) + +lemma le_nhds: "F \<le> nhds a \<longleftrightarrow> (\<forall>S. open S \<and> a \<in> S \<longrightarrow> eventually (\<lambda>x. x \<in> S) F)" + unfolding le_filter_def eventually_nhds by (fast elim: eventually_mono) + +lemma le_nhds_metric: "F \<le> nhds a \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist x a < e) F)" + unfolding le_filter_def eventually_nhds_metric by (fast elim: eventually_mono) + +lemma le_nhds_metric_le: "F \<le> nhds a \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist x a \<le> e) F)" + unfolding le_filter_def eventually_nhds_metric_le by (fast elim: eventually_mono) + +text \<open>Several results are easier using a "multiplied-out" variant. +(I got this idea from Dieudonne's proof of the chain rule).\<close> + +lemma has_derivative_within_alt: + "(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and> + (\<forall>e>0. \<exists>d>0. \<forall>y\<in>s. norm(y - x) < d \<longrightarrow> norm (f y - f x - f' (y - x)) \<le> 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) \<longleftrightarrow> bounded_linear f' \<and> + (\<forall>e>0. eventually (\<lambda>y. norm (f y - f x - f' (y - x)) \<le> 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) \<longleftrightarrow> + bounded_linear f' \<and> + (\<forall>e>0. \<exists>d>0. \<forall>y. norm(y - x) < d \<longrightarrow> norm (f y - f x - f'(y - x)) \<le> e * norm (y - x))" + using has_derivative_within_alt[where s=UNIV] + by simp + + +subsection \<open>The chain rule\<close> + +lemma 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 \<circ> f) has_derivative (g' \<circ> 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) \<Longrightarrow> + (g has_derivative g') (at (f x)) \<Longrightarrow> ((g \<circ> f) has_derivative (g' \<circ> f')) (at x)" + using has_derivative_compose[of f f' x UNIV g g'] + by (simp add: comp_def) + + +subsection \<open>Composition rules stated just for differentiability\<close> + +lemma differentiable_chain_at: + "f differentiable (at x) \<Longrightarrow> + g differentiable (at (f x)) \<Longrightarrow> (g \<circ> f) differentiable (at x)" + unfolding differentiable_def + by (meson diff_chain_at) + +lemma differentiable_chain_within: + "f differentiable (at x within s) \<Longrightarrow> + g differentiable (at(f x) within (f ` s)) \<Longrightarrow> (g \<circ> f) differentiable (at x within s)" + unfolding differentiable_def + by (meson diff_chain_within) + + +subsection \<open>Uniqueness of derivative\<close> + + +text \<open> + The general result is a bit messy because we need approachability of the + limit point from any direction. But OK for nontrivial intervals etc. +\<close> + +lemma frechet_derivative_unique_within: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "(f has_derivative f') (at x within s)" + and "(f has_derivative f'') (at x within s)" + and "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> (x + d *\<^sub>R i) \<in> 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 (rule, rule) + fix e :: real + assume "e > 0" + obtain d where "0 < \<bar>d\<bar>" and "\<bar>d\<bar> < e" and "x + d *\<^sub>R (SOME i. i \<in> Basis) \<in> s" + using assms(3) SOME_Basis \<open>e>0\<close> by blast + then show "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e" + apply (rule_tac x="x + d *\<^sub>R (SOME i. i \<in> Basis)" in bexI) + unfolding dist_norm + apply (auto simp: SOME_Basis nonzero_Basis) + done + qed + then have *: "netlimit (at x within s) = x" + apply (auto intro!: netlimit_within) + by (metis trivial_limit_within) + show ?thesis + apply (rule linear_eq_stdbasis) + unfolding linear_conv_bounded_linear + apply (rule as(1,2)[THEN conjunct1])+ + proof (rule, rule ccontr) + fix i :: 'a + assume i: "i \<in> Basis" + define e where "e = norm (f' i - f'' i)" + assume "f' i \<noteq> f'' i" + then have "e > 0" + unfolding e_def by auto + obtain d where d: + "0 < d" + "(\<And>xa. xa\<in>s \<longrightarrow> 0 < dist xa x \<and> dist xa x < d \<longrightarrow> + dist ((f xa - f x - f' (xa - x)) /\<^sub>R norm (xa - x) - + (f xa - f x - f'' (xa - x)) /\<^sub>R norm (xa - x)) (0 - 0) < e)" + using tendsto_diff [OF as(1,2)[THEN conjunct2]] + unfolding * Lim_within + using \<open>e>0\<close> by blast + obtain c where c: "0 < \<bar>c\<bar>" "\<bar>c\<bar> < d \<and> x + c *\<^sub>R i \<in> s" + using assms(3) i d(1) by blast + have *: "norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R i)) = + norm ((1 / \<bar>c\<bar>) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))" + unfolding scaleR_right_distrib by auto + also have "\<dots> = norm ((1 / \<bar>c\<bar>) *\<^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 "\<dots> = 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 + +lemma frechet_derivative_unique_at: + "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f'') (at x) \<Longrightarrow> f' = f''" + by (rule has_derivative_unique) + +lemma frechet_derivative_unique_within_closed_interval: + fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" + and "x \<in> 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''" + apply(rule frechet_derivative_unique_within) + apply(rule assms(3,4))+ +proof (rule, rule, rule) + fix e :: real + fix i :: 'a + assume "e > 0" and i: "i \<in> Basis" + then show "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> cbox a b" + proof (cases "x\<bullet>i = a\<bullet>i") + case True + then show ?thesis + apply (rule_tac x="(min (b\<bullet>i - a\<bullet>i) e) / 2" in exI) + using assms(1)[THEN bspec[where x=i]] and \<open>e>0\<close> and assms(2) + unfolding mem_box + using i + apply (auto simp add: field_simps inner_simps inner_Basis) + done + next + note * = assms(2)[unfolded mem_box, THEN bspec, OF i] + case False + moreover have "a \<bullet> i < x \<bullet> i" + using False * by auto + moreover { + have "a \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e \<le> a\<bullet>i *2 + x\<bullet>i - a\<bullet>i" + by auto + also have "\<dots> = a\<bullet>i + x\<bullet>i" + by auto + also have "\<dots> \<le> 2 * (x\<bullet>i)" + using * by auto + finally have "a \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e \<le> x \<bullet> i * 2" + by auto + } + moreover have "min (x \<bullet> i - a \<bullet> i) e \<ge> 0" + using * and \<open>e>0\<close> by auto + then have "x \<bullet> i * 2 \<le> b \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e" + using * by auto + ultimately show ?thesis + apply (rule_tac x="- (min (x\<bullet>i - a\<bullet>i) e) / 2" in exI) + using assms(1)[THEN bspec, OF i] and \<open>e>0\<close> and assms(2) + unfolding mem_box + using i + apply (auto simp add: field_simps inner_simps inner_Basis) + done + qed +qed + +lemma frechet_derivative_unique_within_open_interval: + fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "x \<in> box a b" + and "(f has_derivative f' ) (at x within box a b)" + and "(f has_derivative f'') (at x within box a b)" + shows "f' = f''" +proof - + from assms(1) have *: "at x within box a b = at x" + by (metis at_within_interior interior_open open_box) + from assms(2,3) [unfolded *] show "f' = f''" + by (rule frechet_derivative_unique_at) +qed + +lemma frechet_derivative_at: + "(f has_derivative f') (at x) \<Longrightarrow> f' = frechet_derivative f (at x)" + apply (rule frechet_derivative_unique_at[of f]) + apply assumption + unfolding frechet_derivative_works[symmetric] + using differentiable_def + apply auto + done + +lemma frechet_derivative_within_cbox: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" + and "x \<in> 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) + + +subsection \<open>The traditional Rolle theorem in one dimension\<close> + +text \<open>Derivatives of local minima and maxima are zero.\<close> + +lemma has_derivative_local_min: + fixes f :: "'a::real_normed_vector \<Rightarrow> real" + assumes deriv: "(f has_derivative f') (at x)" + assumes min: "eventually (\<lambda>y. f x \<le> f y) (at x)" + shows "f' = (\<lambda>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") + assume "h \<noteq> 0" + from min obtain d where d1: "0 < d" and d2: "\<forall>y\<in>ball x d. f x \<le> f y" + unfolding eventually_at by (force simp: dist_commute) + have "FDERIV (\<lambda>r. x + r *\<^sub>R h) 0 :> (\<lambda>r. r *\<^sub>R h)" + by (intro derivative_eq_intros) auto + then have "FDERIV (\<lambda>r. f (x + r *\<^sub>R h)) 0 :> (\<lambda>k. f' (k *\<^sub>R h))" + by (rule has_derivative_compose, simp add: deriv) + then have "DERIV (\<lambda>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 \<open>h \<noteq> 0\<close> by simp + moreover have "\<forall>y. \<bar>0 - y\<bar> < d / norm h \<longrightarrow> f (x + 0 *\<^sub>R h) \<le> f (x + y *\<^sub>R h)" + using \<open>h \<noteq> 0\<close> by (auto simp add: d2 dist_norm pos_less_divide_eq) + ultimately show "f' h = 0" + by (rule DERIV_local_min) + qed (simp add: f'.zero) +qed + +lemma has_derivative_local_max: + fixes f :: "'a::real_normed_vector \<Rightarrow> real" + assumes "(f has_derivative f') (at x)" + assumes "eventually (\<lambda>y. f y \<le> f x) (at x)" + shows "f' = (\<lambda>h. 0)" + using has_derivative_local_min [of "\<lambda>x. - f x" "\<lambda>h. - f' h" "x"] + using assms unfolding fun_eq_iff by simp + +lemma differential_zero_maxmin: + fixes f::"'a::real_normed_vector \<Rightarrow> real" + assumes "x \<in> s" + and "open s" + and deriv: "(f has_derivative f') (at x)" + and mono: "(\<forall>y\<in>s. f y \<le> f x) \<or> (\<forall>y\<in>s. f x \<le> f y)" + shows "f' = (\<lambda>v. 0)" + using mono +proof + assume "\<forall>y\<in>s. f y \<le> f x" + with \<open>x \<in> s\<close> and \<open>open s\<close> have "eventually (\<lambda>y. f y \<le> f x) (at x)" + unfolding eventually_at_topological by auto + with deriv show ?thesis + by (rule has_derivative_local_max) +next + assume "\<forall>y\<in>s. f x \<le> f y" + with \<open>x \<in> s\<close> and \<open>open s\<close> have "eventually (\<lambda>y. f x \<le> 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: (* TODO: delete? *) + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes k: "k \<in> Basis" + and ball: "0 < e" "(\<forall>y \<in> ball x e. (f y)\<bullet>k \<le> (f x)\<bullet>k) \<or> (\<forall>y\<in>ball x e. (f x)\<bullet>k \<le> (f y)\<bullet>k)" + and diff: "f differentiable (at x)" + shows "(\<Sum>j\<in>Basis. (frechet_derivative f (at x) j \<bullet> k) *\<^sub>R j) = (0::'a)" (is "?D k = 0") +proof - + let ?f' = "frechet_derivative f (at x)" + have "x \<in> ball x e" using \<open>0 < e\<close> by simp + moreover have "open (ball x e)" by simp + moreover have "((\<lambda>x. f x \<bullet> k) has_derivative (\<lambda>h. ?f' h \<bullet> k)) (at x)" + using bounded_linear_inner_left diff[unfolded frechet_derivative_works] + by (rule bounded_linear.has_derivative) + ultimately have "(\<lambda>h. frechet_derivative f (at x) h \<bullet> k) = (\<lambda>v. 0)" + using ball(2) by (rule differential_zero_maxmin) + then show ?thesis + unfolding fun_eq_iff by simp +qed + +lemma rolle: + fixes f :: "real \<Rightarrow> real" + assumes "a < b" + and "f a = f b" + and "continuous_on {a .. b} f" + and "\<forall>x\<in>{a <..< b}. (f has_derivative f' x) (at x)" + shows "\<exists>x\<in>{a <..< b}. f' x = (\<lambda>v. 0)" +proof - + have "\<exists>x\<in>box a b. (\<forall>y\<in>box a b. f x \<le> f y) \<or> (\<forall>y\<in>box a b. f y \<le> f x)" + proof - + have "(a + b) / 2 \<in> {a .. b}" + using assms(1) by auto + then have *: "{a .. b} \<noteq> {}" + by auto + obtain d where d: + "d \<in>cbox a b" + "\<forall>y\<in>cbox a b. f y \<le> f d" + using continuous_attains_sup[OF compact_Icc * assms(3)] by auto + obtain c where c: + "c \<in> cbox a b" + "\<forall>y\<in>cbox a b. f c \<le> f y" + using continuous_attains_inf[OF compact_Icc * assms(3)] by auto + show ?thesis + proof (cases "d \<in> box a b \<or> c \<in> box a b") + case True + then show ?thesis + by (metis c(2) d(2) box_subset_cbox subset_iff) + next + define e where "e = (a + b) /2" + case False + then have "f d = f c" + using d c assms(2) by auto + then have "\<And>x. x \<in> {a..b} \<Longrightarrow> f x = f d" + using c d + by force + then show ?thesis + apply (rule_tac x=e in bexI) + unfolding e_def + using assms(1) + apply auto + done + qed + qed + then obtain x where x: "x \<in> {a <..< b}" "(\<forall>y\<in>{a <..< b}. f x \<le> f y) \<or> (\<forall>y\<in>{a <..< b}. f y \<le> f x)" + by auto + then have "f' x = (\<lambda>v. 0)" + apply (rule_tac differential_zero_maxmin[of x "box a b" f "f' x"]) + using assms + apply auto + done + then show ?thesis + by (metis x(1)) +qed + + +subsection \<open>One-dimensional mean value theorem\<close> + +lemma mvt: + fixes f :: "real \<Rightarrow> real" + assumes "a < b" + and "continuous_on {a..b} f" + assumes "\<forall>x\<in>{a<..<b}. (f has_derivative (f' x)) (at x)" + shows "\<exists>x\<in>{a<..<b}. f b - f a = (f' x) (b - a)" +proof - + have "\<exists>x\<in>{a <..< b}. (\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa) = (\<lambda>v. 0)" + proof (intro rolle[OF assms(1), of "\<lambda>x. f x - (f b - f a) / (b - a) * x"] ballI) + fix x + assume x: "x \<in> {a <..< b}" + show "((\<lambda>x. f x - (f b - f a) / (b - a) * x) has_derivative + (\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa)) (at x)" + by (intro derivative_intros assms(3)[rule_format,OF x]) + qed (insert assms(1,2), auto intro!: continuous_intros simp: field_simps) + then obtain x where + "x \<in> {a <..< b}" + "(\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa) = (\<lambda>v. 0)" .. + then show ?thesis + by (metis (hide_lams) assms(1) diff_gt_0_iff_gt eq_iff_diff_eq_0 + zero_less_mult_iff nonzero_mult_divide_cancel_right not_real_square_gt_zero + times_divide_eq_left) +qed + +lemma mvt_simple: + fixes f :: "real \<Rightarrow> real" + assumes "a < b" + and "\<forall>x\<in>{a..b}. (f has_derivative f' x) (at x within {a..b})" + shows "\<exists>x\<in>{a<..<b}. f b - f a = f' x (b - a)" +proof (rule mvt) + have "f differentiable_on {a..b}" + using assms(2) unfolding differentiable_on_def differentiable_def by fast + then show "continuous_on {a..b} f" + by (rule differentiable_imp_continuous_on) + show "\<forall>x\<in>{a<..<b}. (f has_derivative f' x) (at x)" + proof + fix x + assume x: "x \<in> {a <..< b}" + show "(f has_derivative f' x) (at x)" + unfolding at_within_open[OF x open_greaterThanLessThan,symmetric] + apply (rule has_derivative_within_subset) + apply (rule assms(2)[rule_format]) + using x + apply auto + done + qed +qed (rule assms(1)) + +lemma mvt_very_simple: + fixes f :: "real \<Rightarrow> real" + assumes "a \<le> b" + and "\<forall>x\<in>{a .. b}. (f has_derivative f' x) (at x within {a .. b})" + shows "\<exists>x\<in>{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 + apply (rule_tac x=a in bexI) + using assms(2)[THEN bspec[where x=a]] + unfolding has_derivative_def + unfolding True + using zero + apply auto + done +next + case False + then show ?thesis + using mvt_simple[OF _ assms(2)] + using assms(1) + by auto +qed + +text \<open>A nice generalization (see Havin's proof of 5.19 from Rudin's book).\<close> + +lemma mvt_general: + fixes f :: "real \<Rightarrow> 'a::real_inner" + assumes "a < b" + and "continuous_on {a .. b} f" + and "\<forall>x\<in>{a<..<b}. (f has_derivative f'(x)) (at x)" + shows "\<exists>x\<in>{a<..<b}. norm (f b - f a) \<le> norm (f' x (b - a))" +proof - + have "\<exists>x\<in>{a<..<b}. (f b - f a) \<bullet> f b - (f b - f a) \<bullet> f a = (f b - f a) \<bullet> f' x (b - a)" + apply (rule mvt) + apply (rule assms(1)) + apply (intro continuous_intros assms(2)) + using assms(3) + apply (fast intro: has_derivative_inner_right) + done + then obtain x where x: + "x \<in> {a<..<b}" + "(f b - f a) \<bullet> f b - (f b - f a) \<bullet> f a = (f b - f a) \<bullet> 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 "\<dots> = (f b - f a) \<bullet> (f b - f a)" + unfolding power2_norm_eq_inner .. + also have "\<dots> = (f b - f a) \<bullet> f' x (b - a)" + using x(2) by (simp only: inner_diff_right) + also have "\<dots> \<le> 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 assms(1) + apply (rule_tac x="(a + b) /2" in bexI) + apply auto + done + qed +qed + + +subsection \<open>More general bound theorems\<close> + +lemma differentiable_bound_general: + fixes f :: "real \<Rightarrow> 'a::real_normed_vector" + assumes "a < b" + and f_cont: "continuous_on {a .. b} f" + and phi_cont: "continuous_on {a .. b} \<phi>" + and f': "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> (f has_vector_derivative f' x) (at x)" + and phi': "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> (\<phi> has_vector_derivative \<phi>' x) (at x)" + and bnd: "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> norm (f' x) \<le> \<phi>' x" + shows "norm (f b - f a) \<le> \<phi> b - \<phi> a" +proof - + { + fix x assume x: "a < x" "x < b" + have "0 \<le> norm (f' x)" by simp + also have "\<dots> \<le> \<phi>' x" using x by (auto intro!: bnd) + finally have "0 \<le> \<phi>' 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 \<open>e > 0\<close> have "e2 > 0" by simp + let ?le = "\<lambda>x1. norm (f x1 - f a) \<le> \<phi> x1 - \<phi> a + e * (x1 - a) + e" + define A where "A = {x2. a \<le> x2 \<and> x2 \<le> b \<and> (\<forall>x1\<in>{a ..< x2}. ?le x1)}" + have A_subset: "A \<subseteq> {a .. b}" by (auto simp: A_def) + { + fix x2 + assume a: "a \<le> x2" "x2 \<le> b" and le: "\<forall>x1\<in>{a..<x2}. ?le x1" + have "?le x2" using \<open>e > 0\<close> + proof cases + assume "x2 \<noteq> a" with a have "a < x2" by simp + have "at x2 within {a <..<x2}\<noteq> bot" + using \<open>a < x2\<close> + by (auto simp: trivial_limit_within islimpt_in_closure) + moreover + have "((\<lambda>x1. (\<phi> x1 - \<phi> a) + e * (x1 - a) + e) \<longlongrightarrow> (\<phi> x2 - \<phi> a) + e * (x2 - a) + e) (at x2 within {a <..<x2})" + "((\<lambda>x1. norm (f x1 - f a)) \<longlongrightarrow> norm (f x2 - f a)) (at x2 within {a <..<x2})" + using a + by (auto intro!: tendsto_eq_intros f_tendsto phi_tendsto + intro: tendsto_within_subset[where S="{a .. b}"]) + moreover + have "eventually (\<lambda>x. x > a) (at x2 within {a <..<x2})" + by (auto simp: eventually_at_filter) + hence "eventually ?le (at x2 within {a <..<x2})" + unfolding eventually_at_filter + by eventually_elim (insert le, auto) + ultimately + show ?thesis + by (rule tendsto_le) + qed simp + } note le_cont = this + have "a \<in> A" + using assms by (auto simp: A_def) + hence [simp]: "A \<noteq> {}" by auto + have A_ivl: "\<And>x1 x2. x2 \<in> A \<Longrightarrow> x1 \<in> {a ..x2} \<Longrightarrow> x1 \<in> A" + by (simp add: A_def) + have [simp]: "bdd_above A" by (auto simp: A_def) + define y where "y = Sup A" + have "y \<le> b" + unfolding y_def + by (simp add: cSup_le_iff) (simp add: A_def) + have leI: "\<And>x x1. a \<le> x1 \<Longrightarrow> x \<in> A \<Longrightarrow> x1 < x \<Longrightarrow> ?le x1" + by (auto simp: A_def intro!: le_cont) + have y_all_le: "\<forall>x1\<in>{a..<y}. ?le x1" + by (auto simp: y_def less_cSup_iff leI) + have "a \<le> y" + by (metis \<open>a \<in> A\<close> \<open>bdd_above A\<close> cSup_upper y_def) + have "y \<in> A" + using y_all_le \<open>a \<le> y\<close> \<open>y \<le> b\<close> + 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 \<open>a \<le> y\<close> \<open>y \<le> b\<close> y_all_le] have le_y: "?le y" . + { + assume "a \<noteq> y" with \<open>a \<le> y\<close> have "a < y" by simp + have "y = b" + proof (rule ccontr) + assume "y \<noteq> b" + hence "y < b" using \<open>y \<le> b\<close> by simp + let ?F = "at y within {y..<b}" + from f' phi' + have "(f has_vector_derivative f' y) ?F" + and "(\<phi> has_vector_derivative \<phi>' y) ?F" + using \<open>a < y\<close> \<open>y < b\<close> + by (auto simp add: at_within_open[of _ "{a<..<b}"] has_vector_derivative_def + intro!: has_derivative_subset[where s="{a<..<b}" and t="{y..<b}"]) + hence "\<forall>\<^sub>F x1 in ?F. norm (f x1 - f y - (x1 - y) *\<^sub>R f' y) \<le> e2 * \<bar>x1 - y\<bar>" + "\<forall>\<^sub>F x1 in ?F. norm (\<phi> x1 - \<phi> y - (x1 - y) *\<^sub>R \<phi>' y) \<le> e2 * \<bar>x1 - y\<bar>" + using \<open>e2 > 0\<close> + by (auto simp: has_derivative_within_alt2 has_vector_derivative_def) + moreover + have "\<forall>\<^sub>F x1 in ?F. y \<le> x1" "\<forall>\<^sub>F x1 in ?F. x1 < b" + by (auto simp: eventually_at_filter) + ultimately + have "\<forall>\<^sub>F x1 in ?F. norm (f x1 - f y) \<le> (\<phi> x1 - \<phi> y) + e * \<bar>x1 - y\<bar>" + (is "\<forall>\<^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) \<le> norm (f' y) * \<bar>x1 - y\<bar> + e2 * \<bar>x1 - y\<bar>" + by (simp add: ac_simps) + also have "norm (f' y) \<le> \<phi>' y" using bnd \<open>a < y\<close> \<open>y < b\<close> by simp + also + from elim have "\<phi>' y * \<bar>x1 - y\<bar> \<le> \<phi> x1 - \<phi> y + e2 * \<bar>x1 - y\<bar>" + by (simp add: ac_simps) + finally + have "norm (f x1 - f y) \<le> \<phi> x1 - \<phi> y + e2 * \<bar>x1 - y\<bar> + e2 * \<bar>x1 - y\<bar>" + 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 \<in> S" "\<And>x. x\<in>S \<Longrightarrow> x \<in> {y..<b} \<Longrightarrow> ?le' x" + unfolding eventually_at_topological + by metis + from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0" + by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \<open>y \<in> S\<close>]) + define d' where "d' = min ((y + b)/2) (y + (d/2))" + have "d' \<in> A" + unfolding A_def + proof safe + show "a \<le> d'" using \<open>a < y\<close> \<open>0 < d\<close> \<open>y < b\<close> by (simp add: d'_def) + show "d' \<le> b" using \<open>y < b\<close> by (simp add: d'_def min_def) + fix x1 + assume x1: "x1 \<in> {a..<d'}" + { + assume "x1 < y" + hence "?le x1" + using \<open>x1 \<in> {a..<d'}\<close> y_all_le by auto + } moreover { + assume "x1 \<ge> y" + hence x1': "x1 \<in> S" "x1 \<in> {y..<b}" using x1 + by (auto simp: d'_def dist_real_def intro!: d) + have "norm (f x1 - f a) \<le> 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 have "?le x1" + using \<open>x1 \<ge> y\<close> by (auto simp: algebra_simps) + } ultimately show "?le x1" by arith + qed + hence "d' \<le> y" + unfolding y_def + by (rule cSup_upper) simp + thus False using \<open>d > 0\<close> \<open>y < b\<close> + by (simp add: d'_def min_def split: if_split_asm) + qed + } moreover { + assume "a = y" + with \<open>a < b\<close> have "y < b" by simp + with \<open>a = y\<close> f_cont phi_cont \<open>e2 > 0\<close> + have 1: "\<forall>\<^sub>F x in at y within {y..b}. dist (f x) (f y) < e2" + and 2: "\<forall>\<^sub>F x in at y within {y..b}. dist (\<phi> x) (\<phi> y) < e2" + by (auto simp: continuous_on_def tendsto_iff) + have 3: "eventually (\<lambda>x. y < x) (at y within {y..b})" + by (auto simp: eventually_at_filter) + have 4: "eventually (\<lambda>x::real. x < b) (at y within {y..b})" + using _ \<open>y < b\<close> + by (rule order_tendstoD) (auto intro!: tendsto_eq_intros) + from 1 2 3 4 + have eventually_le: "eventually (\<lambda>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: \<open>a = y\<close>) + also have "norm (f x1 - f y) \<le> e2" + using elim \<open>a = y\<close> by (auto simp : dist_norm intro!: less_imp_le) + also have "\<dots> \<le> e2 + (\<phi> x1 - \<phi> a + e2 + e * (x1 - a))" + using \<open>0 < e\<close> elim + by (intro add_increasing2[OF add_nonneg_nonneg order.refl]) + (auto simp: \<open>a = y\<close> dist_norm intro!: mult_nonneg_nonneg) + also have "\<dots> = \<phi> x1 - \<phi> a + e * (x1 - a) + e" + by (simp add: e2_def) + finally show "?le x1" . + qed + from this[unfolded eventually_at_topological] \<open>?le y\<close> + obtain S + where S: "open S" "y \<in> S" "\<And>x. x\<in>S \<Longrightarrow> x \<in> {y..b} \<Longrightarrow> ?le x" + by metis + from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0" + by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \<open>y \<in> S\<close>]) + define d' where "d' = min b (y + (d/2))" + have "d' \<in> A" + unfolding A_def + proof safe + show "a \<le> d'" using \<open>a = y\<close> \<open>0 < d\<close> \<open>y < b\<close> by (simp add: d'_def) + show "d' \<le> b" by (simp add: d'_def) + fix x1 + assume "x1 \<in> {a..<d'}" + hence "x1 \<in> S" "x1 \<in> {y..b}" + by (auto simp: \<open>a = y\<close> d'_def dist_real_def intro!: d ) + thus "?le x1" + by (rule S) + qed + hence "d' \<le> y" + unfolding y_def + by (rule cSup_upper) simp + hence "y = b" using \<open>d > 0\<close> \<open>y < b\<close> + by (simp add: d'_def) + } ultimately have "y = b" + by auto + with le_y have "norm (f b - f a) \<le> \<phi> b - \<phi> a + e * (b - a + 1)" + by (simp add: algebra_simps) + } note * = this + { + fix e::real assume "e > 0" + hence "norm (f b - f a) \<le> \<phi> b - \<phi> a + e" + using *[of "e / (b - a + 1)"] \<open>a < b\<close> by simp + } thus ?thesis + by (rule field_le_epsilon) +qed + +lemma differentiable_bound: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "convex s" + and "\<forall>x\<in>s. (f has_derivative f' x) (at x within s)" + and "\<forall>x\<in>s. onorm (f' x) \<le> B" + and x: "x \<in> s" + and y: "y \<in> s" + shows "norm (f x - f y) \<le> B * norm (x - y)" +proof - + let ?p = "\<lambda>u. x + u *\<^sub>R (y - x)" + let ?\<phi> = "\<lambda>h. h * B * norm (x - y)" + have *: "\<And>u. u\<in>{0..1} \<Longrightarrow> x + u *\<^sub>R (y - x) \<in> s" + using assms(1)[unfolded convex_alt,rule_format,OF x y] + unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib + by (auto simp add: algebra_simps) + have 0: "continuous_on (?p ` {0..1}) f" + using * + unfolding continuous_on_eq_continuous_within + apply - + apply rule + apply (rule differentiable_imp_continuous_within) + unfolding differentiable_def + apply (rule_tac x="f' xa" in exI) + apply (rule has_derivative_within_subset) + apply (rule assms(2)[rule_format]) + apply auto + done + from * have 1: "continuous_on {0 .. 1} (f \<circ> ?p)" + by (intro continuous_intros 0)+ + { + fix u::real assume u: "u \<in>{0 <..< 1}" + let ?u = "?p u" + interpret linear "(f' ?u)" + using u by (auto intro!: has_derivative_linear assms(2)[rule_format] *) + have "(f \<circ> ?p has_derivative (f' ?u) \<circ> (\<lambda>u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)" + apply (rule diff_chain_within) + apply (rule derivative_intros)+ + apply (rule has_derivative_within_subset) + apply (rule assms(2)[rule_format]) + using u * + apply auto + done + hence "((f \<circ> ?p) has_vector_derivative f' ?u (y - x)) (at u)" + by (simp add: has_derivative_within_open[OF u open_greaterThanLessThan] + scaleR has_vector_derivative_def o_def) + } note 2 = this + { + have "continuous_on {0..1} ?\<phi>" + by (rule continuous_intros)+ + } note 3 = this + { + fix u::real assume u: "u \<in>{0 <..< 1}" + have "(?\<phi> has_vector_derivative B * norm (x - y)) (at u)" + by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros) + } note 4 = this + { + fix u::real assume u: "u \<in>{0 <..< 1}" + let ?u = "?p u" + interpret bounded_linear "(f' ?u)" + using u by (auto intro!: has_derivative_bounded_linear assms(2)[rule_format] *) + have "norm (f' ?u (y - x)) \<le> onorm (f' ?u) * norm (y - x)" + by (rule onorm) fact + also have "onorm (f' ?u) \<le> B" + using u by (auto intro!: assms(3)[rule_format] *) + finally have "norm ((f' ?u) (y - x)) \<le> B * norm (x - y)" + by (simp add: mult_right_mono norm_minus_commute) + } note 5 = this + have "norm (f x - f y) = norm ((f \<circ> (\<lambda>u. x + u *\<^sub>R (y - x))) 1 - (f \<circ> (\<lambda>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 \<circ> ?p) 1 - (f \<circ> ?p) 0) \<le> B * norm (x - y)" + by simp + finally show ?thesis . +qed + +lemma + differentiable_bound_segment: + fixes f::"'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "\<And>t. t \<in> {0..1} \<Longrightarrow> x0 + t *\<^sub>R a \<in> G" + assumes f': "\<And>x. x \<in> G \<Longrightarrow> (f has_derivative f' x) (at x within G)" + assumes B: "\<forall>x\<in>{0..1}. onorm (f' (x0 + x *\<^sub>R a)) \<le> B" + shows "norm (f (x0 + a) - f x0) \<le> norm a * B" +proof - + let ?G = "(\<lambda>x. x0 + x *\<^sub>R a) ` {0..1}" + have "?G = op + x0 ` (\<lambda>x. x *\<^sub>R a) ` {0..1}" by auto + also have "convex \<dots>" + by (intro convex_translation convex_scaled convex_real_interval) + finally have "convex ?G" . + moreover have "?G \<subseteq> G" "x0 \<in> ?G" "x0 + a \<in> ?G" using assms by (auto intro: image_eqI[where x=1]) + ultimately show ?thesis + using has_derivative_subset[OF f' \<open>?G \<subseteq> G\<close>] B + differentiable_bound[of "(\<lambda>x. x0 + x *\<^sub>R a) ` {0..1}" f f' B "x0 + a" x0] + by (auto simp: ac_simps) +qed + +lemma differentiable_bound_linearization: + fixes f::"'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "\<And>t. t \<in> {0..1} \<Longrightarrow> a + t *\<^sub>R (b - a) \<in> S" + assumes f'[derivative_intros]: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" + assumes B: "\<forall>x\<in>S. onorm (f' x - f' x0) \<le> B" + assumes "x0 \<in> S" + shows "norm (f b - f a - f' x0 (b - a)) \<le> norm (b - a) * B" +proof - + define g where [abs_def]: "g x = f x - f' x0 x" for x + have g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative (\<lambda>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 B: "\<forall>x\<in>{0..1}. onorm (\<lambda>i. f' (a + x *\<^sub>R (b - a)) i - f' x0 i) \<le> B" + using assms by (auto simp: fun_diff_def) + from differentiable_bound_segment[OF assms(1) g B] \<open>x0 \<in> S\<close> + show ?thesis + by (simp add: g_def field_simps linear_diff[OF has_derivative_linear[OF f']]) +qed + +text \<open>In particular.\<close> + +lemma has_derivative_zero_constant: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "convex s" + and "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)" + shows "\<exists>c. \<forall>x\<in>s. f x = c" +proof - + { fix x y assume "x \<in> s" "y \<in> s" + then have "norm (f x - f y) \<le> 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" "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative 0) (at x within s)" + shows "\<exists>c. \<forall>x\<in>s. f (x) = (c :: 'a :: real_normed_field)" +proof (rule has_derivative_zero_constant) + have A: "op * 0 = (\<lambda>_. 0 :: 'a)" by (intro ext) simp + fix x assume "x \<in> s" thus "(f has_derivative (\<lambda>h. 0)) (at x within s)" + using assms(2)[of x] by (simp add: has_field_derivative_def A) +qed fact + +lemma has_derivative_zero_unique: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "convex s" + and "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)" + and "x \<in> s" "y \<in> 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 \<Rightarrow> 'b::real_normed_vector" + assumes "open s" "connected s" + assumes f: "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative (\<lambda>x. 0)) (at x)" + assumes "x \<in> s" "y \<in> s" + shows "f x = f y" +proof (rule connected_local_const[where f=f, OF \<open>connected s\<close> \<open>x\<in>s\<close> \<open>y\<in>s\<close>]) + show "\<forall>a\<in>s. eventually (\<lambda>b. f a = f b) (at a within s)" + proof + fix a assume "a \<in> s" + with \<open>open s\<close> obtain e where "0 < e" "ball a e \<subseteq> s" + by (rule openE) + then have "\<exists>c. \<forall>x\<in>ball a e. f x = c" + by (intro has_derivative_zero_constant) + (auto simp: at_within_open[OF _ open_ball] f convex_ball) + with \<open>0<e\<close> have "\<forall>x\<in>ball a e. f a = f x" + by auto + then show "eventually (\<lambda>b. f a = f b) (at a within s)" + using \<open>0<e\<close> unfolding eventually_at_topological + by (intro exI[of _ "ball a e"]) auto + qed +qed + +subsection \<open>Differentiability of inverse function (most basic form)\<close> + +lemma has_derivative_inverse_basic: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "(f has_derivative f') (at (g y))" + and "bounded_linear g'" + and "g' \<circ> f' = id" + and "continuous (at y) g" + and "open t" + and "y \<in> t" + and "\<forall>z\<in>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" "\<And>x. norm (g' x) \<le> norm x * C" + using bounded_linear.pos_bounded[OF assms(2)] by blast + have lem1: "\<forall>e>0. \<exists>d>0. \<forall>z. + norm (z - y) < d \<longrightarrow> norm (g z - g y - g'(z - y)) \<le> e * norm (g z - g y)" + proof (rule, rule) + fix e :: real + assume "e > 0" + with C(1) have *: "e / C > 0" by auto + obtain d0 where d0: + "0 < d0" + "\<forall>ya. norm (ya - g y) < d0 \<longrightarrow> norm (f ya - f (g y) - f' (ya - g y)) \<le> e / C * norm (ya - g y)" + using assms(1) + unfolding has_derivative_at_alt + using * by blast + obtain d1 where d1: + "0 < d1" + "\<forall>x. 0 < dist x y \<and> dist x y < d1 \<longrightarrow> dist (g x) (g y) < d0" + using assms(4) + unfolding continuous_at Lim_at + using d0(1) by blast + obtain d2 where d2: + "0 < d2" + "\<forall>ya. dist ya y < d2 \<longrightarrow> ya \<in> t" + using assms(5) + unfolding open_dist + using assms(6) by blast + obtain d where d: "0 < d" "d < d1" "d < d2" + using real_lbound_gt_zero[OF d1(1) d2(1)] by blast + then show "\<exists>d>0. \<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)" + apply (rule_tac x=d in exI) + apply rule + defer + apply rule + apply rule + proof - + fix z + assume as: "norm (z - y) < d" + then have "z \<in> t" + using d2 d unfolding dist_norm by auto + have "norm (g z - g y - g' (z - y)) \<le> 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] + unfolding assms(7)[rule_format,OF \<open>z\<in>t\<close>] + apply (subst norm_minus_cancel[symmetric]) + apply auto + done + also have "\<dots> \<le> norm (f (g z) - y - f' (g z - g y)) * C" + by (rule C(2)) + also have "\<dots> \<le> (e / C) * norm (g z - g y) * C" + apply (rule mult_right_mono) + apply (rule d0(2)[rule_format,unfolded assms(7)[rule_format,OF \<open>y\<in>t\<close>]]) + apply (cases "z = y") + defer + apply (rule d1(2)[unfolded dist_norm,rule_format]) + using as d C d0 + apply auto + done + also have "\<dots> \<le> e * norm (g z - g y)" + using C by (auto simp add: field_simps) + finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)" + by simp + qed auto + qed + have *: "(0::real) < 1 / 2" + by auto + obtain d where d: + "0 < d" + "\<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> 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) \<le> B * norm (z - y)" if z: "norm(z - y) < d" for z + proof - + have "norm (g z - g y) \<le> norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" + by (rule norm_triangle_sub) + also have "\<dots> \<le> norm (g' (z - y)) + 1 / 2 * norm (g z - g y)" + apply (rule add_left_mono) + using d and z + apply auto + done + also have "\<dots> \<le> norm (z - y) * C + 1 / 2 * norm (g z - g y)" + apply (rule add_right_mono) + using C + apply auto + done + finally show "norm (g z - g y) \<le> B * norm (z - y)" + unfolding B_def + by (auto simp add: field_simps) + qed + show ?thesis + unfolding has_derivative_at_alt + apply rule + apply (rule assms) + apply rule + apply rule + proof - + fix e :: real + assume "e > 0" + then have *: "e / B > 0" by (metis \<open>B > 0\<close> divide_pos_pos) + obtain d' where d': + "0 < d'" + "\<forall>z. norm (z - y) < d' \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e / B * norm (g z - g y)" + using lem1 * by blast + obtain k where k: "0 < k" "k < d" "k < d'" + using real_lbound_gt_zero[OF d(1) d'(1)] by blast + show "\<exists>d>0. \<forall>ya. norm (ya - y) < d \<longrightarrow> norm (g ya - g y - g' (ya - y)) \<le> e * norm (ya - y)" + apply (rule_tac x=k in exI) + apply auto + proof - + fix z + assume as: "norm (z - y) < k" + then have "norm (g z - g y - g' (z - y)) \<le> e / B * norm(g z - g y)" + using d' k by auto + also have "\<dots> \<le> e * norm (z - y)" + unfolding times_divide_eq_left pos_divide_le_eq[OF \<open>B>0\<close>] + using lem2[of z] + using k as using \<open>e > 0\<close> + by (auto simp add: field_simps) + finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (z - y)" + by simp + qed(insert k, auto) + qed +qed + +text \<open>Simply rewrite that based on the domain point x.\<close> + +lemma has_derivative_inverse_basic_x: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "(f has_derivative f') (at x)" + and "bounded_linear g'" + and "g' \<circ> f' = id" + and "continuous (at (f x)) g" + and "g (f x) = x" + and "open t" + and "f x \<in> t" + and "\<forall>y\<in>t. f (g y) = y" + shows "(g has_derivative g') (at (f x))" + apply (rule has_derivative_inverse_basic) + using assms + apply auto + done + +text \<open>This is the version in Dieudonne', assuming continuity of f and g.\<close> + +lemma has_derivative_inverse_dieudonne: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "open s" + and "open (f ` s)" + and "continuous_on s f" + and "continuous_on (f ` s) g" + and "\<forall>x\<in>s. g (f x) = x" + and "x \<in> s" + and "(f has_derivative f') (at x)" + and "bounded_linear g'" + and "g' \<circ> 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 \<open>Here's the simplest way of not assuming much about g.\<close> + +lemma has_derivative_inverse: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "compact s" + and "x \<in> s" + and "f x \<in> interior (f ` s)" + and "continuous_on s f" + and "\<forall>y\<in>s. g (f y) = y" + and "(f has_derivative f') (at x)" + and "bounded_linear g'" + and "g' \<circ> f' = id" + shows "(g has_derivative g') (at (f x))" +proof - + { + fix y + assume "y \<in> interior (f ` s)" + then obtain x where "x \<in> s" and *: "y = f x" + unfolding image_iff + using interior_subset + by auto + have "f (g y) = y" + unfolding * and assms(5)[rule_format,OF \<open>x\<in>s\<close>] .. + } note * = this + show ?thesis + apply (rule has_derivative_inverse_basic_x[OF assms(6-8)]) + apply (rule continuous_on_interior[OF _ assms(3)]) + apply (rule continuous_on_inv[OF assms(4,1)]) + apply (rule assms(2,5) assms(5)[rule_format] open_interior assms(3))+ + apply (metis *) + done +qed + + +subsection \<open>Proving surjectivity via Brouwer fixpoint theorem\<close> + +lemma brouwer_surjective: + fixes f :: "'n::euclidean_space \<Rightarrow> 'n" + assumes "compact t" + and "convex t" + and "t \<noteq> {}" + and "continuous_on t f" + and "\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t" + and "x \<in> s" + shows "\<exists>y\<in>t. f y = x" +proof - + have *: "\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" + by (auto simp add: algebra_simps) + show ?thesis + unfolding * + apply (rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"]) + apply (rule continuous_intros assms)+ + using assms(4-6) + apply auto + done +qed + +lemma brouwer_surjective_cball: + fixes f :: "'n::euclidean_space \<Rightarrow> 'n" + assumes "e > 0" + and "continuous_on (cball a e) f" + and "\<forall>x\<in>s. \<forall>y\<in>cball a e. x + (y - f y) \<in> cball a e" + and "x \<in> s" + shows "\<exists>y\<in>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 + +text \<open>See Sussmann: "Multidifferential calculus", Theorem 2.1.1\<close> + +lemma sussmann_open_mapping: + fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" + assumes "open s" + and "continuous_on s f" + and "x \<in> s" + and "(f has_derivative f') (at x)" + and "bounded_linear g'" "f' \<circ> g' = id" + and "t \<subseteq> s" + and "x \<in> interior t" + shows "f x \<in> 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" "\<forall>x. norm (g' x) \<le> 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" + "\<forall>y. norm (y - x) < e0 \<longrightarrow> norm (f y - f x - f' (y - x)) \<le> 1 / (2 * B) * norm (y - x)" + using assms(4) + unfolding has_derivative_at_alt + using * by blast + obtain e1 where e1: "0 < e1" "cball x e1 \<subseteq> t" + using assms(8) + unfolding mem_interior_cball + 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 real_lbound_gt_zero[OF *] by blast + have "\<forall>z\<in>cball (f x) (e / 2). \<exists>y\<in>cball (f x) e. f (x + g' (y - f x)) = z" + apply rule + apply (rule brouwer_surjective_cball[where s="cball (f x) (e/2)"]) + prefer 3 + apply rule + apply rule + proof- + show "continuous_on (cball (f x) e) (\<lambda>y. f (x + g' (y - f x)))" + unfolding g'.diff + apply (rule continuous_on_compose[of _ _ f, unfolded o_def]) + apply (rule continuous_intros linear_continuous_on[OF assms(5)])+ + apply (rule continuous_on_subset[OF assms(2)]) + apply rule + apply (unfold image_iff) + apply (erule bexE) + proof- + fix y z + assume as: "y \<in>cball (f x) e" "z = x + (g' y - g' (f x))" + have "dist x z = norm (g' (f x) - g' y)" + unfolding as(2) and dist_norm by auto + also have "\<dots> \<le> norm (f x - y) * B" + unfolding g'.diff[symmetric] + using B + by auto + also have "\<dots> \<le> e * B" + using as(1)[unfolded mem_cball dist_norm] + using B + by auto + also have "\<dots> \<le> e1" + using e + unfolding less_divide_eq + using B + by auto + finally have "z \<in> cball x e1" + unfolding mem_cball + by force + then show "z \<in> s" + using e1 assms(7) by auto + qed + next + fix y z + assume as: "y \<in> cball (f x) (e / 2)" "z \<in> cball (f x) e" + have "norm (g' (z - f x)) \<le> norm (z - f x) * B" + using B by auto + also have "\<dots> \<le> e * B" + apply (rule mult_right_mono) + using as(2)[unfolded mem_cball dist_norm] and B + unfolding norm_minus_commute + apply auto + done + also have "\<dots> < e0" + using e and B + unfolding less_divide_eq + by auto + 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))))) \<le> + 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 "\<dots> \<le> 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 "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" + using as(1)[unfolded mem_cball dist_norm] + by auto + also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" + using * and B + by (auto simp add: field_simps) + also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" + by auto + also have "\<dots> \<le> e/2 + e/2" + apply (rule add_right_mono) + using as(2)[unfolded mem_cball dist_norm] + unfolding norm_minus_commute + apply auto + done + finally show "y + (z - f (x + g' (z - f x))) \<in> cball (f x) e" + unfolding mem_cball dist_norm + by auto + qed (insert e, auto) note lem = this + show ?thesis + unfolding mem_interior + apply (rule_tac x="e/2" in exI) + apply rule + apply (rule divide_pos_pos) + prefer 3 + proof + fix y + assume "y \<in> ball (f x) (e / 2)" + then have *: "y \<in> cball (f x) (e / 2)" + by auto + obtain z where z: "z \<in> cball (f x) e" "f (x + g' (z - f x)) = y" + using lem * by blast + then have "norm (g' (z - f x)) \<le> norm (z - f x) * B" + using B + by (auto simp add: field_simps) + also have "\<dots> \<le> e * B" + apply (rule mult_right_mono) + using z(1) + unfolding mem_cball dist_norm norm_minus_commute + using B + apply auto + done + also have "\<dots> \<le> e1" + using e B unfolding less_divide_eq by auto + finally have "x + g'(z - f x) \<in> t" + apply - + apply (rule e1(2)[unfolded subset_eq,rule_format]) + unfolding mem_cball dist_norm + apply auto + done + then show "y \<in> f ` t" + using z by auto + qed (insert e, auto) +qed + +text \<open>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 \<open>f' \<circ> g = I\<close> but this happens to fit with the minimal linear + algebra theory I've set up so far.\<close> + +(* move before left_inverse_linear in Euclidean_Space*) + +lemma right_inverse_linear: + fixes f :: "'a::euclidean_space \<Rightarrow> 'a" + assumes lf: "linear f" + and gf: "f \<circ> g = id" + shows "linear g" +proof - + from gf have fi: "surj f" + by (auto simp add: surj_def o_def id_def) metis + from linear_surjective_isomorphism[OF lf fi] + obtain h:: "'a \<Rightarrow> 'a" where h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" + by blast + have "h = g" + apply (rule ext) + using gf h(2,3) + apply (simp add: o_def id_def fun_eq_iff) + apply metis + done + with h(1) show ?thesis by blast +qed + +lemma has_derivative_inverse_strong: + fixes f :: "'n::euclidean_space \<Rightarrow> 'n" + assumes "open s" + and "x \<in> s" + and "continuous_on s f" + and "\<forall>x\<in>s. g (f x) = x" + and "(f has_derivative f') (at x)" + and "f' \<circ> g' = id" + shows "(g has_derivative g') (at (f x))" +proof - + have linf: "bounded_linear f'" + using assms(5) unfolding has_derivative_def by auto + then have ling: "bounded_linear g'" + unfolding linear_conv_bounded_linear[symmetric] + apply - + apply (rule right_inverse_linear) + using assms(6) + apply auto + done + moreover have "g' \<circ> f' = id" + using assms(6) linf ling + unfolding linear_conv_bounded_linear[symmetric] + using linear_inverse_left + by auto + moreover have *:"\<forall>t\<subseteq>s. x \<in> interior t \<longrightarrow> f x \<in> interior (f ` t)" + apply clarify + 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 \<in> interior (f ` (ball x e \<inter> s))" + using *[rule_format,of "ball x e \<inter> s"] \<open>x \<in> s\<close> + 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 \<subseteq> f ` (ball x e \<inter> s)" + unfolding mem_interior by blast + show "\<exists>d>0. \<forall>y. 0 < dist y (f x) \<and> dist y (f x) < d \<longrightarrow> dist (g y) (g (f x)) < e" + apply (rule_tac x=d in exI) + apply rule + apply (rule d(1)) + apply rule + apply rule + proof - + fix y + assume "0 < dist y (f x) \<and> dist y (f x) < d" + then have "g y \<in> g ` f ` (ball x e \<inter> s)" + using d(2)[unfolded subset_eq,THEN bspec[where x=y]] + by (auto simp add: dist_commute) + then have "g y \<in> ball x e \<inter> s" + using assms(4) by auto + then show "dist (g y) (g (f x)) < e" + using assms(4)[rule_format,OF \<open>x \<in> s\<close>] + by (auto simp add: dist_commute) + qed + qed + moreover have "f x \<in> interior (f ` s)" + apply (rule sussmann_open_mapping) + apply (rule assms ling)+ + using interior_open[OF assms(1)] and \<open>x \<in> s\<close> + apply auto + done + moreover have "f (g y) = y" if "y \<in> interior (f ` s)" for y + proof - + from that have "y \<in> f ` s" + using interior_subset by auto + then obtain z where "z \<in> s" "y = f z" unfolding image_iff .. + then show ?thesis + using assms(4) by auto + qed + ultimately show ?thesis using assms + by (metis has_derivative_inverse_basic_x open_interior) +qed + +text \<open>A rewrite based on the other domain.\<close> + +lemma has_derivative_inverse_strong_x: + fixes f :: "'a::euclidean_space \<Rightarrow> 'a" + assumes "open s" + and "g y \<in> s" + and "continuous_on s f" + and "\<forall>x\<in>s. g (f x) = x" + and "(f has_derivative f') (at (g y))" + and "f' \<circ> 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 \<open>On a region.\<close> + +lemma has_derivative_inverse_on: + fixes f :: "'n::euclidean_space \<Rightarrow> 'n" + assumes "open s" + and "\<forall>x\<in>s. (f has_derivative f'(x)) (at x)" + and "\<forall>x\<in>s. g (f x) = x" + and "f' x \<circ> g' x = id" + and "x \<in> s" + shows "(g has_derivative g'(x)) (at (f x))" + apply (rule has_derivative_inverse_strong[where g'="g' x" and f=f]) + apply (rule assms)+ + unfolding continuous_on_eq_continuous_at[OF assms(1)] + apply rule + apply (rule differentiable_imp_continuous_within) + unfolding differentiable_def + using assms + apply auto + done + +text \<open>Invertible derivative continous 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.\<close> + +proposition has_derivative_locally_injective: + fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" + assumes "a \<in> s" + and "open s" + and "bounded_linear g'" + and "g' \<circ> f' a = id" + and "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative f' x) (at x)" + and "\<And>e. e > 0 \<Longrightarrow> \<exists>d>0. \<forall>x. dist a x < d \<longrightarrow> onorm (\<lambda>v. f' x v - f' a v) < e" + obtains r where "r > 0" "ball a r \<subseteq> 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 (\<Sum>Basis)) = (\<Sum>Basis)" "(\<Sum>Basis) \<noteq> (0::'n)" + defer + apply (subst euclidean_eq_iff) + using f'g' + apply auto + done + 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" + "\<And>x. dist a x < d1 \<Longrightarrow> onorm (\<lambda>v. f' x v - f' a v) < k" + using assms(6) * by blast + from \<open>open s\<close> obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s" + using \<open>a\<in>s\<close> .. + obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s" + using assms(2,1) .. + obtain d2 where d2: "0 < d2" "ball a d2 \<subseteq> s" + using assms(2) + unfolding open_contains_ball + using \<open>a\<in>s\<close> by blast + obtain d where d: "0 < d" "d < d1" "d < d2" + using real_lbound_gt_zero[OF d1(1) d2(1)] by blast + show ?thesis + proof + show "0 < d" by (fact d) + show "ball a d \<subseteq> s" + using \<open>d < d2\<close> \<open>ball a d2 \<subseteq> s\<close> by auto + show "inj_on f (ball a d)" + unfolding inj_on_def + proof (intro strip) + fix x y + assume as: "x \<in> ball a d" "y \<in> 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' \<circ> (\<lambda>w. f' a w - (f w - f x))" + unfolding ph_def o_def + unfolding diff + using f'g' + by (auto simp: algebra_simps) + have "norm (ph x - ph y) \<le> (1 / 2) * norm (x - y)" + apply (rule differentiable_bound[OF convex_ball _ _ as(1-2), where f'="\<lambda>x v. v - g'(f' x v)"]) + apply (rule_tac[!] ballI) + proof - + fix u + assume u: "u \<in> ball a d" + then have "u \<in> s" + using d d2 by auto + have *: "(\<lambda>v. v - g' (f' u v)) = g' \<circ> (\<lambda>w. f' a w - f' u w)" + unfolding o_def and diff + using f'g' by auto + show "(ph has_derivative (\<lambda>v. v - g' (f' u v))) (at u within ball a d)" + unfolding ph' * + apply (simp add: comp_def) + apply (rule bounded_linear.has_derivative[OF assms(3)]) + apply (rule derivative_intros) + defer + apply (rule has_derivative_sub[where g'="\<lambda>x.0",unfolded diff_0_right]) + apply (rule has_derivative_at_within) + using assms(5) and \<open>u \<in> s\<close> \<open>a \<in> s\<close> + apply (auto intro!: derivative_intros bounded_linear.has_derivative[of _ "\<lambda>x. x"] has_derivative_bounded_linear) + done + have **: "bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)" + apply (rule_tac[!] bounded_linear_sub) + apply (rule_tac[!] has_derivative_bounded_linear) + using assms(5) \<open>u \<in> s\<close> \<open>a \<in> s\<close> + apply auto + done + have "onorm (\<lambda>v. v - g' (f' u v)) \<le> onorm g' * onorm (\<lambda>w. f' a w - f' u w)" + unfolding * + apply (rule onorm_compose) + apply (rule assms(3) **)+ + done + also have "\<dots> \<le> onorm g' * k" + apply (rule mult_left_mono) + using d1(2)[of u] + using onorm_neg[where f="\<lambda>x. f' u x - f' a x"] + using d and u and onorm_pos_le[OF assms(3)] + apply (auto simp: algebra_simps) + done + also have "\<dots> \<le> 1 / 2" + unfolding k_def by auto + finally show "onorm (\<lambda>v. v - g' (f' u v)) \<le> 1 / 2" . + qed + moreover have "norm (ph y - ph x) = norm (y - x)" + apply (rule arg_cong[where f=norm]) + unfolding ph_def + using diff + unfolding as + apply auto + done + ultimately show "x = y" + unfolding norm_minus_commute by auto + qed + qed +qed + + +subsection \<open>Uniformly convergent sequence of derivatives\<close> + +lemma has_derivative_sequence_lipschitz_lemma: + fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "convex s" + and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" + and "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" + and "0 \<le> e" + shows "\<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm ((f m x - f n x) - (f m y - f n y)) \<le> 2 * e * norm (x - y)" +proof rule+ + fix m n x y + assume as: "N \<le> m" "N \<le> n" "x \<in> s" "y \<in> s" + show "norm ((f m x - f n x) - (f m y - f n y)) \<le> 2 * e * norm (x - y)" + apply (rule differentiable_bound[where f'="\<lambda>x h. f' m x h - f' n x h", OF assms(1) _ _ as(3-4)]) + apply (rule_tac[!] ballI) + proof - + fix x + assume "x \<in> s" + show "((\<lambda>a. f m a - f n a) has_derivative (\<lambda>h. f' m x h - f' n x h)) (at x within s)" + by (rule derivative_intros assms(2)[rule_format] \<open>x\<in>s\<close>)+ + show "onorm (\<lambda>h. f' m x h - f' n x h) \<le> 2 * e" + proof (rule onorm_bound) + fix h + have "norm (f' m x h - f' n x h) \<le> 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"] + unfolding norm_minus_commute + by (auto simp add: algebra_simps) + also have "\<dots> \<le> e * norm h + e * norm h" + using assms(3)[rule_format,OF \<open>N \<le> m\<close> \<open>x \<in> s\<close>, of h] + using assms(3)[rule_format,OF \<open>N \<le> n\<close> \<open>x \<in> s\<close>, of h] + by (auto simp add: field_simps) + finally show "norm (f' m x h - f' n x h) \<le> 2 * e * norm h" + by auto + qed (simp add: \<open>0 \<le> e\<close>) + qed +qed + +lemma has_derivative_sequence_lipschitz: + fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + assumes "convex s" + and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" + and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" + shows "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. + norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)" +proof (rule, rule) + fix e :: real + assume "e > 0" + then have *: "2 * (1/2* e) = e" "1/2 * e >0" + by auto + obtain N where "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> 1 / 2 * e * norm h" + using assms(3) *(2) by blast + then show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)" + apply (rule_tac x=N in exI) + apply (rule has_derivative_sequence_lipschitz_lemma[where e="1/2 *e", unfolded *]) + using assms \<open>e > 0\<close> + apply auto + done +qed + +lemma has_derivative_sequence: + fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach" + assumes "convex s" + and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" + and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" + and "x0 \<in> s" + and "((\<lambda>n. f n x0) \<longlongrightarrow> l) sequentially" + shows "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g'(x)) (at x within s)" +proof - + have lem1: "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. + norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)" + using assms(1,2,3) by (rule has_derivative_sequence_lipschitz) + have "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially" + apply (rule bchoice) + unfolding convergent_eq_cauchy + proof + fix x + assume "x \<in> s" + show "Cauchy (\<lambda>n. f n x)" + proof (cases "x = x0") + case True + then show ?thesis + using LIMSEQ_imp_Cauchy[OF assms(5)] by auto + next + case False + show ?thesis + unfolding Cauchy_def + proof (rule, rule) + fix e :: real + assume "e > 0" + hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto + obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x0) (f n x0) < e / 2" + using LIMSEQ_imp_Cauchy[OF assms(5)] + unfolding Cauchy_def + using *(1) by blast + obtain N where N: + "\<forall>m\<ge>N. \<forall>n\<ge>N. + \<forall>xa\<in>s. \<forall>y\<in>s. norm (f m xa - f n xa - (f m y - f n y)) \<le> + e / 2 / norm (x - x0) * norm (xa - y)" + using lem1 *(2) by blast + show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x) (f n x) < e" + apply (rule_tac x="max M N" in exI) + proof rule+ + fix m n + assume as: "max M N \<le>m" "max M N\<le>n" + have "dist (f m x) (f n x) \<le> + 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 "\<dots> \<le> norm (f m x0 - f n x0) + e / 2" + using N[rule_format,OF _ _ \<open>x\<in>s\<close> \<open>x0\<in>s\<close>, of m n] and as and False + by auto + also have "\<dots> < e / 2 + e / 2" + apply (rule add_strict_right_mono) + using as and M[rule_format] + unfolding dist_norm + apply auto + done + finally show "dist (f m x) (f n x) < e" + by auto + qed + qed + qed + qed + then obtain g where g: "\<forall>x\<in>s. (\<lambda>n. f n x) \<longlonglongrightarrow> g x" .. + have lem2: "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm ((f n x - f n y) - (g x - g y)) \<le> e * norm (x - y)" + proof (rule, rule) + fix e :: real + assume *: "e > 0" + obtain N where + N: "\<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)" + using lem1 * by blast + show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" + apply (rule_tac x=N in exI) + proof rule+ + fix n x y + assume as: "N \<le> n" "x \<in> s" "y \<in> s" + have "((\<lambda>m. norm (f n x - f n y - (f m x - f m y))) \<longlongrightarrow> norm (f n x - f n y - (g x - g y))) sequentially" + by (intro tendsto_intros g[rule_format] as) + moreover have "eventually (\<lambda>m. norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)) sequentially" + unfolding eventually_sequentially + apply (rule_tac x=N in exI) + apply rule + apply rule + proof - + fix m + assume "N \<le> m" + then show "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)" + using N[rule_format, of n m x y] and as + by (auto simp add: algebra_simps) + qed + ultimately show "norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" + by (rule tendsto_ge_const[OF trivial_limit_sequentially]) + qed + qed + have "\<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g' x) (at x within s)" + unfolding has_derivative_within_alt2 + proof (intro ballI conjI) + fix x + assume "x \<in> s" + then show "((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially" + by (simp add: g) + have lem3: "\<forall>u. ((\<lambda>n. f' n x u) \<longlongrightarrow> g' x u) sequentially" + unfolding filterlim_def le_nhds_metric_le eventually_filtermap dist_norm + proof (intro allI impI) + fix u + fix e :: real + assume "e > 0" + show "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e) sequentially" + proof (cases "u = 0") + case True + have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e * norm u) sequentially" + using assms(3)[folded eventually_sequentially] and \<open>0 < e\<close> and \<open>x \<in> s\<close> + by (fast elim: eventually_mono) + then show ?thesis + using \<open>u = 0\<close> and \<open>0 < e\<close> by (auto elim: eventually_mono) + next + case False + with \<open>0 < e\<close> have "0 < e / norm u" by simp + then have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e / norm u * norm u) sequentially" + using assms(3)[folded eventually_sequentially] and \<open>x \<in> s\<close> + by (fast elim: eventually_mono) + then show ?thesis + using \<open>u \<noteq> 0\<close> 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 \<open>x\<in>s\<close>,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]) + apply (rule lem3[rule_format]) + unfolding lin[THEN bounded_linear.linear, THEN linear_cmul] + apply (intro tendsto_intros) + apply (rule lem3[rule_format]) + done + show "g' x (y + z) = g' x y + g' x z" + apply (rule tendsto_unique[OF trivial_limit_sequentially]) + apply (rule lem3[rule_format]) + unfolding lin[THEN bounded_linear.linear, THEN linear_add] + apply (rule tendsto_add) + apply (rule lem3[rule_format])+ + done + obtain N where N: "\<forall>h. norm (f' N x h - g' x h) \<le> 1 * norm h" + using assms(3) \<open>x \<in> s\<close> by (fast intro: zero_less_one) + have "bounded_linear (f' N x)" + using assms(2) \<open>x \<in> s\<close> by fast + from bounded_linear.bounded [OF this] + obtain K where K: "\<forall>h. norm (f' N x h) \<le> 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 "\<dots> \<le> norm (f' N x h) + norm (f' N x h - g' x h)" + by (rule norm_triangle_ineq4) + also have "\<dots> \<le> norm h * K + 1 * norm h" + using N K by (fast intro: add_mono) + finally have "norm (g' x h) \<le> norm h * (K + 1)" + by (simp add: ring_distribs) + } + then show "\<exists>K. \<forall>h. norm (g' x h) \<le> norm h * K" by fast + qed + show "\<forall>e>0. eventually (\<lambda>y. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)) (at x within s)" + proof (rule, rule) + fix e :: real + assume "e > 0" + then have *: "e / 3 > 0" + by auto + obtain N1 where N1: "\<forall>n\<ge>N1. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e / 3 * norm h" + using assms(3) * by blast + obtain N2 where + N2: "\<forall>n\<ge>N2. \<forall>x\<in>s. \<forall>y\<in>s. norm (f n x - f n y - (g x - g y)) \<le> e / 3 * norm (x - y)" + using lem2 * by blast + let ?N = "max N1 N2" + have "eventually (\<lambda>y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)) (at x within s)" + using assms(2)[unfolded has_derivative_within_alt2] and \<open>x \<in> s\<close> and * by fast + moreover have "eventually (\<lambda>y. y \<in> s) (at x within s)" + unfolding eventually_at by (fast intro: zero_less_one) + ultimately show "\<forall>\<^sub>F y in at x within s. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)" + proof (rule eventually_elim2) + fix y + assume "y \<in> s" + assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)" + moreover have "norm (g y - g x - (f ?N y - f ?N x)) \<le> e / 3 * norm (y - x)" + using N2[rule_format, OF _ \<open>y \<in> s\<close> \<open>x \<in> s\<close>] + by (simp add: norm_minus_commute) + ultimately have "norm (g y - g x - f' ?N x (y - x)) \<le> 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)) \<le> e / 3 * norm (y - x)" + using N1 \<open>x \<in> s\<close> by auto + ultimately show "norm (g y - g x - g' x (y - x)) \<le> 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 \<open>Can choose to line up antiderivatives if we want.\<close> + +lemma has_antiderivative_sequence: + fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach" + assumes "convex s" + and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" + and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" + shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)" +proof (cases "s = {}") + case False + then obtain a where "a \<in> s" + by auto + have *: "\<And>P Q. \<exists>g. \<forall>x\<in>s. P g x \<and> Q g x \<Longrightarrow> \<exists>g. \<forall>x\<in>s. Q g x" + by auto + show ?thesis + apply (rule *) + apply (rule has_derivative_sequence[OF assms(1) _ assms(3), of "\<lambda>n x. f n x + (f 0 a - f n a)"]) + apply (metis assms(2) has_derivative_add_const) + apply (rule \<open>a \<in> s\<close>) + apply auto + done +qed auto + +lemma has_antiderivative_limit: + fixes g' :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'b::banach" + assumes "convex s" + and "\<forall>e>0. \<exists>f f'. \<forall>x\<in>s. + (f has_derivative (f' x)) (at x within s) \<and> (\<forall>h. norm (f' x h - g' x h) \<le> e * norm h)" + shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)" +proof - + have *: "\<forall>n. \<exists>f f'. \<forall>x\<in>s. + (f has_derivative (f' x)) (at x within s) \<and> + (\<forall>h. norm(f' x h - g' x h) \<le> inverse (real (Suc n)) * norm h)" + by (simp add: assms(2)) + obtain f where + *: "\<forall>x. \<exists>f'. \<forall>xa\<in>s. (f x has_derivative f' xa) (at xa within s) \<and> + (\<forall>h. norm (f' xa h - g' xa h) \<le> inverse (real (Suc x)) * norm h)" + using *[THEN choice] .. + obtain f' where + f: "\<forall>x. \<forall>xa\<in>s. (f x has_derivative f' x xa) (at xa within s) \<and> + (\<forall>h. norm (f' x xa h - g' xa h) \<le> inverse (real (Suc x)) * norm h)" + using *[THEN choice] .. + show ?thesis + apply (rule has_antiderivative_sequence[OF assms(1), of f f']) + defer + apply rule + apply rule + proof - + fix e :: real + assume "e > 0" + obtain N where N: "inverse (real (Suc N)) < e" + using reals_Archimedean[OF \<open>e>0\<close>] .. + show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" + apply (rule_tac x=N in exI) + apply rule + apply rule + apply rule + apply rule + proof - + fix n x h + assume n: "N \<le> n" and x: "x \<in> s" + have *: "inverse (real (Suc n)) \<le> 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) \<le> e * norm h" + using f[rule_format,THEN conjunct2, OF x, of n, THEN spec[where x=h]] + apply (rule order_trans) + using N * + apply (cases "h = 0") + apply auto + done + qed + qed (insert f, auto) +qed + + +subsection \<open>Differentiation of a series\<close> + +lemma has_derivative_series: + fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach" + assumes "convex s" + and "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within s)" + and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (setsum (\<lambda>i. f' i x h) {..<n} - g' x h) \<le> e * norm h" + and "x \<in> s" + and "(\<lambda>n. f n x) sums l" + shows "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums (g x) \<and> (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_setsum) + using assms(4-5) + unfolding sums_def + apply auto + done + +lemma has_field_derivative_series: + fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" + assumes "convex s" + assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" + assumes "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" + assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" + shows "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (g has_field_derivative g' x) (at x within s)" +unfolding has_field_derivative_def +proof (rule has_derivative_series) + show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" + proof (intro allI impI) + fix e :: real assume "e > 0" + with assms(3) obtain N where N: "\<And>n x. n \<ge> N \<Longrightarrow> x \<in> s \<Longrightarrow> norm ((\<Sum>i<n. f' i x) - g' x) < e" + unfolding uniform_limit_iff eventually_at_top_linorder dist_norm by blast + { + fix n :: nat and x h :: 'a assume nx: "n \<ge> N" "x \<in> s" + have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) = norm ((\<Sum>i<n. f' i x) - g' x) * norm h" + by (simp add: norm_mult [symmetric] ring_distribs setsum_left_distrib) + also from N[OF nx] have "norm ((\<Sum>i<n. f' i x) - g' x) \<le> e" by simp + hence "norm ((\<Sum>i<n. f' i x) - g' x) * norm h \<le> e * norm h" + by (intro mult_right_mono) simp_all + finally have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" . + } + thus "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" by blast + qed +qed (insert assms, auto simp: has_field_derivative_def) + +lemma has_field_derivative_series': + fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" + assumes "convex s" + assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" + assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" + assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" "x \<in> interior s" + shows "summable (\<lambda>n. f n x)" "((\<lambda>x. \<Sum>n. f n x) has_field_derivative (\<Sum>n. f' n x)) (at x)" +proof - + from \<open>x \<in> interior s\<close> have "x \<in> s" using interior_subset by blast + define g' where [abs_def]: "g' x = (\<Sum>i. f' i x)" for x + from assms(3) have "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" + by (simp add: uniformly_convergent_uniform_limit_iff suminf_eq_lim g'_def) + from has_field_derivative_series[OF assms(1,2) this assms(4,5)] obtain g where g: + "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x" + "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast + from g(1)[OF \<open>x \<in> s\<close>] show "summable (\<lambda>n. f n x)" by (simp add: sums_iff) + from g(2)[OF \<open>x \<in> s\<close>] \<open>x \<in> interior s\<close> 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) \<longleftrightarrow> + ((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' x) (at x)" + using eventually_nhds_in_nhd[OF \<open>x \<in> interior s\<close>] interior_subset[of s] g(1) + by (intro DERIV_cong_ev) (auto elim!: eventually_mono simp: sums_iff) + finally show "((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' x) (at x)" . +qed + +lemma differentiable_series: + fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" + assumes "convex s" "open s" + assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" + assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" and x: "x \<in> s" + shows "summable (\<lambda>n. f n x)" and "(\<lambda>x. \<Sum>n. f n x) differentiable (at x)" +proof - + from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" + unfolding uniformly_convergent_on_def by blast + from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open) + have "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (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: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x" + "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast + from g[OF x] show "summable (\<lambda>n. f n x)" by (auto simp: summable_def) + from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)" + by (simp add: has_field_derivative_def s) + have "((\<lambda>x. \<Sum>n. f n x) has_derivative op * (g' x)) (at x)" + by (rule has_derivative_transform_within_open[OF g' \<open>open s\<close> x]) + (insert g, auto simp: sums_iff) + thus "(\<lambda>x. \<Sum>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 \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" + assumes "convex s" "open s" + assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" + assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" + assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" + shows "(\<lambda>x. \<Sum>n. f n x) differentiable (at x0)" + using differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+ + +text \<open>Considering derivative @{typ "real \<Rightarrow> 'b::real_normed_vector"} as a vector.\<close> + +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 \<noteq> 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 "(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')" + proof (rule frechet_derivative_unique_within) + show "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> s" + proof clarsimp + fix e :: real assume "0 < e" + with islimpt_approachable_real[of x s] not_bot + obtain x' where "x' \<in> s" "x' \<noteq> x" "\<bar>x' - x\<bar> < e" + by (auto simp add: trivial_limit_within) + then show "\<exists>d. d \<noteq> 0 \<and> \<bar>d\<bar> < e \<and> x + d \<in> s" + by (intro exI[of _ "x' - x"]) auto + qed + qed (insert f' f'', 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) \<Longrightarrow> (f has_vector_derivative f'') (at x) \<Longrightarrow> f' = f''" + by (rule vector_derivative_unique_within) auto + +lemma differentiableI_vector: "(f has_vector_derivative y) F \<Longrightarrow> f differentiable F" + by (auto simp: differentiable_def has_vector_derivative_def) + +lemma vector_derivative_works: + "f differentiable net \<longleftrightarrow> (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 \<open>?l\<close> 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 \<noteq> 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)) = (\<lambda>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 \<Rightarrow> real" + assumes "(f has_derivative f') F" + obtains c where "(f has_real_derivative c) F" +proof - + obtain c where "f' = (\<lambda>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 \<Rightarrow> real" + shows "(\<exists>c. (f has_real_derivative c) F) = (\<exists>D. (f has_derivative D) F)" + by (metis has_field_derivative_def has_real_derivative) + +definition deriv :: "('a \<Rightarrow> 'a::real_normed_field) \<Rightarrow> 'a \<Rightarrow> 'a" where + "deriv f x \<equiv> SOME D. DERIV f x :> D" + +lemma DERIV_imp_deriv: "DERIV f x :> f' \<Longrightarrow> 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 \<longleftrightarrow> (\<exists>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 \<longleftrightarrow> f differentiable at x" + unfolding differentiable_def by (metis DERIV_imp_deriv has_real_derivative_iff) + +lemma real_derivative_chain: + fixes x :: real + shows "f differentiable at x \<Longrightarrow> g differentiable at (f x) + \<Longrightarrow> 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) + +lemma islimpt_closure_open: + fixes s :: "'a::perfect_space set" + assumes "open s" and t: "t = closure s" "x \<in> t" + shows "x islimpt t" +proof cases + assume "x \<in> s" + { fix T assume "x \<in> T" "open T" + then have "open (s \<inter> T)" + using \<open>open s\<close> by auto + then have "s \<inter> T \<noteq> {x}" + using not_open_singleton[of x] by auto + with \<open>x \<in> T\<close> \<open>x \<in> s\<close> have "\<exists>y\<in>t. y \<in> T \<and> y \<noteq> x" + using closure_subset[of s] by (auto simp: t) } + then show ?thesis + by (auto intro!: islimptI) +next + assume "x \<notin> 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 \<in> 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) \<Longrightarrow> vector_derivative f (at x) = f'" + by (intro vector_derivative_within at_neq_bot) + +lemma has_vector_derivative_id_at [simp]: "vector_derivative (\<lambda>x. x) (at a) = 1" + by (simp add: vector_derivative_at) + +lemma vector_derivative_minus_at [simp]: + "f differentiable at a + \<Longrightarrow> vector_derivative (\<lambda>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]: + "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> + \<Longrightarrow> vector_derivative (\<lambda>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]: + "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> + \<Longrightarrow> vector_derivative (\<lambda>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 \<Rightarrow> 'a :: real_normed_algebra" + shows "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> + \<Longrightarrow> vector_derivative (\<lambda>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]: + "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> + \<Longrightarrow> vector_derivative (\<lambda>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_closed_interval: + assumes ab: "a < b" "x \<in> 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 has_vector_derivative_within_subset: + "(f has_vector_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_vector_derivative f') (at x within t)" + by (auto simp: has_vector_derivative_def intro: has_derivative_within_subset) + +lemma has_vector_derivative_at_within: + "(f has_vector_derivative f') (at x) \<Longrightarrow> (f has_vector_derivative f') (at x within s)" + unfolding has_vector_derivative_def + by (rule has_derivative_at_within) + +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 \<in> s" "s \<subseteq> t" + and "\<And>x. x \<in> s \<Longrightarrow> 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) \<longleftrightarrow> (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 \<open>s \<subseteq> t\<close>] by simp +qed + +lemma has_vector_derivative_transform_within: + assumes "(f has_vector_derivative f') (at x within s)" + and "0 < d" + and "x \<in> s" + and "\<And>x'. \<lbrakk>x'\<in>s; dist x' x < d\<rbrakk> \<Longrightarrow> 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 \<in> s" + and "\<And>y. y\<in>s \<Longrightarrow> 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 vector_diff_chain_at: + assumes "(f has_vector_derivative f') (at x)" + and "(g has_vector_derivative g') (at (f x))" + shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x)" + using assms(2) + unfolding has_vector_derivative_def + apply - + apply (drule diff_chain_at[OF assms(1)[unfolded has_vector_derivative_def]]) + apply (simp only: o_def real_scaleR_def scaleR_scaleR) + done + +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 \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)" + using assms(2) + unfolding has_vector_derivative_def + apply - + apply (drule diff_chain_within[OF assms(1)[unfolded has_vector_derivative_def]]) + apply (simp only: o_def real_scaleR_def scaleR_scaleR) + done + +lemma vector_derivative_const_at [simp]: "vector_derivative (\<lambda>x. c) (at a) = 0" + by (simp add: vector_derivative_at) + +lemma vector_derivative_at_within_ivl: + "(f has_vector_derivative f') (at x) \<Longrightarrow> + a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> a<b \<Longrightarrow> vector_derivative f (at x within {a..b}) = f'" +using has_vector_derivative_at_within vector_derivative_within_closed_interval by fastforce + +lemma vector_derivative_chain_at: + assumes "f differentiable at x" "(g differentiable at (f x))" + shows "vector_derivative (g \<circ> 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 \<circ> 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_at_general: (*thanks to Wenda Li*) + assumes "f differentiable at x" "(\<exists>g'. (g has_field_derivative g') (at (f x)))" + shows "vector_derivative (g \<circ> 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 +by (auto simp: vector_derivative_works DERIV_deriv_iff_has_field_derivative) + +lemma exp_scaleR_has_vector_derivative_right: + "((\<lambda>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 \<inter> {t - 1 <..< t + 1})" + have *: "at t within T = ?F" + by (rule at_within_nhd[where S="{t - 1 <..< t + 1}"]) auto + let ?e = "\<lambda>i x. (inverse (1 + real i) * inverse (fact i) * (x - t) ^ i) *\<^sub>R (A * A ^ i)" + have "\<forall>\<^sub>F n in sequentially. + \<forall>x\<in>T \<inter> {t - 1<..<t + 1}. norm (?e n x) \<le> norm (A ^ (n + 1) /\<^sub>R fact (n + 1))" + by (auto simp: divide_simps power_abs intro!: mult_left_le_one_le power_le_one eventuallyI) + then have "uniform_limit (T \<inter> {t - 1<..<t + 1}) (\<lambda>n x. \<Sum>i<n. ?e i x) (\<lambda>x. \<Sum>i. ?e i x) sequentially" + by (rule weierstrass_m_test_ev) (intro summable_ignore_initial_segment summable_norm_exp) + moreover + have "\<forall>\<^sub>F x in sequentially. x > 0" + by (metis eventually_gt_at_top) + then have + "\<forall>\<^sub>F n in sequentially. ((\<lambda>x. \<Sum>i<n. ?e i x) \<longlongrightarrow> A) ?F" + by eventually_elim + (auto intro!: tendsto_eq_intros + simp: power_0_left if_distrib cond_application_beta setsum.delta + cong: if_cong) + ultimately + have [tendsto_intros]: "((\<lambda>x. \<Sum>i. ?e i x) \<longlongrightarrow> A) ?F" + by (auto intro!: swap_uniform_limit[where f="\<lambda>n x. \<Sum>i < n. ?e i x" and F = sequentially]) + have [tendsto_intros]: "((\<lambda>x. if x = t then 0 else 1) \<longlongrightarrow> 1) ?F" + by (rule Lim_eventually) (simp add: eventually_at_filter) + have "((\<lambda>y. ((y - t) / abs (y - t)) *\<^sub>R ((\<Sum>n. ?e n y) - A)) \<longlongrightarrow> 0) (at t within T)" + unfolding * + by (rule tendsto_norm_zero_cancel) (auto intro!: tendsto_eq_intros) + + moreover + + have "\<forall>\<^sub>F x in at t within T. x \<noteq> t" + by (simp add: eventually_at_filter) + then have "\<forall>\<^sub>F x in at t within T. ((x - t) / \<bar>x - t\<bar>) *\<^sub>R ((\<Sum>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) = + ((\<Sum>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 (\<lambda>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 "(\<Sum>n. (x - t) *\<^sub>R ?e n x) = (x - t) *\<^sub>R (\<Sum>n. ?e n x)" + by (rule suminf_scaleR_right[symmetric]) + also have "(\<dots> - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = (x - t) *\<^sub>R ((\<Sum>n. ?e n x) - A) /\<^sub>R norm (x - t)" + by (simp add: algebra_simps) + finally show ?case + by (simp add: divide_simps) + qed + + ultimately + + have "((\<lambda>y. (exp ((y - t) *\<^sub>R A) - 1 - (y - t) *\<^sub>R A) /\<^sub>R norm (y - t)) \<longlongrightarrow> 0) (at t within T)" + by (rule Lim_transform_eventually[rotated]) + from tendsto_mult_right_zero[OF this, where c="exp (t *\<^sub>R A)"] + show "((\<lambda>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)) \<longlongrightarrow> 0) + (at t within T)" + by (rule Lim_transform_eventually[rotated]) + (auto simp: algebra_simps divide_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: "((\<lambda>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) + + +subsection \<open>Relation between convexity and derivative\<close> + +(* TODO: Generalise to real vector spaces? *) +lemma convex_on_imp_above_tangent: + assumes convex: "convex_on A f" and connected: "connected A" + assumes c: "c \<in> interior A" and x : "x \<in> A" + assumes deriv: "(f has_field_derivative f') (at c within A)" + shows "f x - f c \<ge> f' * (x - c)" +proof (cases x c rule: linorder_cases) + assume xc: "x > c" + let ?A' = "interior A \<inter> {c<..}" + from c have "c \<in> interior A \<inter> closure {c<..}" by auto + also have "\<dots> \<subseteq> closure (interior A \<inter> {c<..})" by (intro open_Int_closure_subset) auto + finally have "at c within ?A' \<noteq> bot" by (subst at_within_eq_bot_iff) auto + moreover from deriv have "((\<lambda>y. (f y - f c) / (y - c)) \<longlongrightarrow> f') (at c within ?A')" + unfolding DERIV_within_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) + moreover from eventually_at_right_real[OF xc] + have "eventually (\<lambda>y. (f y - f c) / (y - c) \<le> (f x - f c) / (x - c)) (at_right c)" + proof eventually_elim + fix y assume y: "y \<in> {c<..<x}" + with convex connected x c have "f y \<le> (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 \<le> (f x - f c) / (x - c) * (y - c)" by simp + thus "(f y - f c) / (y - c) \<le> (f x - f c) / (x - c)" using y xc by (simp add: divide_simps) + qed + hence "eventually (\<lambda>y. (f y - f c) / (y - c) \<le> (f x - f c) / (x - c)) (at c within ?A')" + by (blast intro: filter_leD at_le) + ultimately have "f' \<le> (f x - f c) / (x - c)" by (rule tendsto_ge_const) + thus ?thesis using xc by (simp add: field_simps) +next + assume xc: "x < c" + let ?A' = "interior A \<inter> {..<c}" + from c have "c \<in> interior A \<inter> closure {..<c}" by auto + also have "\<dots> \<subseteq> closure (interior A \<inter> {..<c})" by (intro open_Int_closure_subset) auto + finally have "at c within ?A' \<noteq> bot" by (subst at_within_eq_bot_iff) auto + moreover from deriv have "((\<lambda>y. (f y - f c) / (y - c)) \<longlongrightarrow> f') (at c within ?A')" + unfolding DERIV_within_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) + moreover from eventually_at_left_real[OF xc] + have "eventually (\<lambda>y. (f y - f c) / (y - c) \<ge> (f x - f c) / (x - c)) (at_left c)" + proof eventually_elim + fix y assume y: "y \<in> {x<..<c}" + with convex connected x c have "f y \<le> (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 \<le> (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) \<ge> (f x - f c) / (x - c)" using y xc + by (simp add: divide_simps) + qed + hence "eventually (\<lambda>y. (f y - f c) / (y - c) \<ge> (f x - f c) / (x - c)) (at c within ?A')" + by (blast intro: filter_leD at_le) + ultimately have "f' \<ge> (f x - f c) / (x - c)" by (rule tendsto_le_const) + thus ?thesis using xc by (simp add: field_simps) +qed simp_all + + +subsection \<open>Partial derivatives\<close> + +lemma eventually_at_Pair_within_TimesI1: + fixes x::"'a::metric_space" + assumes "\<forall>\<^sub>F x' in at x within X. P x'" + assumes "P x" + shows "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P x'" +proof - + from assms[unfolded eventually_at_topological] + obtain S where S: "open S" "x \<in> S" "\<And>x'. x' \<in> X \<Longrightarrow> x' \<in> S \<Longrightarrow> P x'" + by metis + show "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P x'" + unfolding eventually_at_topological + by (auto intro!: exI[where x="S \<times> UNIV"] S open_Times) +qed + +lemma eventually_at_Pair_within_TimesI2: + fixes x::"'a::metric_space" + assumes "\<forall>\<^sub>F y' in at y within Y. P y'" + assumes "P y" + shows "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P y'" +proof - + from assms[unfolded eventually_at_topological] + obtain S where S: "open S" "y \<in> S" "\<And>y'. y' \<in> Y \<Longrightarrow> y' \<in> S \<Longrightarrow> P y'" + by metis + show "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P y'" + unfolding eventually_at_topological + by (auto intro!: exI[where x="UNIV \<times> S"] S open_Times) +qed + +lemma has_derivative_partialsI: + assumes fx: "\<And>x y. x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> ((\<lambda>x. f x y) has_derivative blinfun_apply (fx x y)) (at x within X)" + assumes fy: "\<And>x y. x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> ((\<lambda>y. f x y) has_derivative blinfun_apply (fy x y)) (at y within Y)" + assumes fx_cont: "continuous_on (X \<times> Y) (\<lambda>(x, y). fx x y)" + assumes fy_cont: "continuous_on (X \<times> Y) (\<lambda>(x, y). fy x y)" + assumes "x \<in> X" "y \<in> Y" + assumes "convex X" "convex Y" + shows "((\<lambda>(x, y). f x y) has_derivative (\<lambda>(tx, ty). fx x y tx + fy x y ty)) (at (x, y) within X \<times> Y)" +proof (safe intro!: has_derivativeI tendstoI, goal_cases) + case (2 e') + define e where "e = e' / 9" + have "e > 0" using \<open>e' > 0\<close> by (simp add: e_def) + + have "(x, y) \<in> X \<times> Y" using assms by auto + from fy_cont[unfolded continuous_on_eq_continuous_within, rule_format, OF this, + unfolded continuous_within, THEN tendstoD, OF \<open>e > 0\<close>] + have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. dist (fy x' y') (fy x y) < e" + by (auto simp: split_beta') + from this[unfolded eventually_at] obtain d' where + "d' > 0" + "\<And>x' y'. x' \<in> X \<Longrightarrow> y' \<in> Y \<Longrightarrow> (x', y') \<noteq> (x, y) \<Longrightarrow> dist (x', y') (x, y) < d' \<Longrightarrow> + dist (fy x' y') (fy x y) < e" + by auto + then + have d': "x' \<in> X \<Longrightarrow> y' \<in> Y \<Longrightarrow> dist (x', y') (x, y) < d' \<Longrightarrow> dist (fy x' y') (fy x y) < e" + for x' y' + using \<open>0 < e\<close> + by (cases "(x', y') = (x, y)") auto + define d where "d = d' / sqrt 2" + have "d > 0" using \<open>0 < d'\<close> by (simp add: d_def) + have d: "x' \<in> X \<Longrightarrow> y' \<in> Y \<Longrightarrow> dist x' x < d \<Longrightarrow> dist y' y < d \<Longrightarrow> 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 \<inter> Y" + have "convex ?S" + by (auto intro!: convex_Int \<open>convex Y\<close>) + { + fix x'::'a and y'::'b + assume x': "x' \<in> X" and y': "y' \<in> Y" + assume dx': "dist x' x < d" and dy': "dist y' y < d" + have "norm (fy x' y' - fy x' y) \<le> 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' \<open>d > 0\<close> \<open>y \<in> Y\<close> 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: "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. (x', y') \<in> X \<times> Y" + using \<open>x \<in> X\<close> \<open>y \<in> Y\<close> + by (auto simp: eventually_at intro!: zero_less_one) + moreover + have ev_dist: "\<forall>\<^sub>F xy in at (x, y) within X \<times> 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 \<open>0 < d\<close>] + ultimately + have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. + norm (f x' y' - f x' y - (fy x' y) (y' - y)) \<le> norm (y' - y) * (e + e)" + proof (eventually_elim, safe) + fix x' y' + assume "x' \<in> X" and y': "y' \<in> 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 \<in> {0 .. 1}" + then have "y + t *\<^sub>R (y' - y) \<in> closed_segment y y'" + by (auto simp: closed_segment_def algebra_simps intro!: exI[where x=t]) + also + have "\<dots> \<subseteq> ball y d \<inter> Y" + using \<open>y \<in> Y\<close> \<open>0 < d\<close> dy y' + by (intro \<open>convex ?S\<close>[unfolded convex_contains_segment, rule_format, of y y']) + (auto simp: dist_commute) + finally have "y + t *\<^sub>R (y' - y) \<in> ?S" . + } note seg = this + + have "\<forall>x\<in>ball y d \<inter> Y. onorm (blinfun_apply (fy x' x) - blinfun_apply (fy x' y)) \<le> e + e" + by (safe intro!: onorm less_imp_le \<open>x' \<in> X\<close> dx) (auto simp: dist_commute \<open>0 < d\<close> \<open>y \<in> Y\<close>) + with seg has_derivative_within_subset[OF assms(2)[OF \<open>x' \<in> X\<close>]] + show "norm (f x' y' - f x' y - (fy x' y) (y' - y)) \<le> norm (y' - y) * (e + e)" + by (rule differentiable_bound_linearization[where S="?S"]) + (auto intro!: \<open>0 < d\<close> \<open>y \<in> Y\<close>) + qed + moreover + let ?le = "\<lambda>x'. norm (f x' y - f x y - (fx x y) (x' - x)) \<le> norm (x' - x) * e" + from fx[OF \<open>x \<in> X\<close> \<open>y \<in> Y\<close>, unfolded has_derivative_within, THEN conjunct2, THEN tendstoD, OF \<open>0 < e\<close>] + have "\<forall>\<^sub>F x' in at x within X. ?le x'" + by eventually_elim + (auto simp: dist_norm divide_simps blinfun.bilinear_simps field_simps split: if_split_asm) + then have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. ?le x'" + by (rule eventually_at_Pair_within_TimesI1) + (simp add: blinfun.bilinear_simps) + moreover have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. norm ((x', y') - (x, y)) \<noteq> 0" + unfolding norm_eq_zero right_minus_eq + by (auto simp: eventually_at intro!: zero_less_one) + moreover + from fy_cont[unfolded continuous_on_eq_continuous_within, rule_format, OF SigmaI[OF \<open>x \<in> X\<close> \<open>y \<in> Y\<close>], + unfolded continuous_within, THEN tendstoD, OF \<open>0 < e\<close>] + have "\<forall>\<^sub>F x' in at x within X. norm (fy x' y - fy x y) < e" + unfolding eventually_at + using \<open>y \<in> Y\<close> + by (auto simp: dist_prod_def dist_norm) + then have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. norm (fy x' y - fy x y) < e" + by (rule eventually_at_Pair_within_TimesI1) + (simp add: blinfun.bilinear_simps \<open>0 < e\<close>) + ultimately + have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. + norm ((f x' y' - f x y - (fx x y (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 y (x' - x) + fy x y (y' - y))) \<le> + 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 y (x' - x))" + by norm + also + assume nz: "norm ((x', y') - (x, y)) \<noteq> 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)) \<le> norm (y' - y) * (e + e)" + also assume "norm (f x' y - f x y - blinfun_apply (fx x y) (x' - x)) \<le> norm (x' - x) * e" + also + have "norm ((fy x y) (y' - y) - (fy x' y) (y' - y)) \<le> norm ((fy x y) - (fy x' y)) * norm (y' - y)" + by (auto simp: blinfun.bilinear_simps[symmetric] intro!: norm_blinfun) + also have "\<dots> \<le> (e + e) * norm (y' - y)" + using \<open>e > 0\<close> nfy + by (auto simp: norm_minus_commute intro!: mult_right_mono) + also have "norm (x' - x) * e \<le> norm (x' - x) * (e + e)" + using \<open>0 < e\<close> by simp + also have "norm (y' - y) * (e + e) + (e + e) * norm (y' - y) + norm (x' - x) * (e + e) \<le> + (norm (y' - y) + norm (x' - x)) * (4 * e)" + using \<open>e > 0\<close> + by (simp add: algebra_simps) + also have "\<dots> \<le> 2 * norm ((x', y') - (x, y)) * (4 * e)" + using \<open>0 < e\<close> 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 "\<dots> \<le> norm ((x', y') - (x, y)) * (8 * e)" + by simp + also have "\<dots> < norm ((x', y') - (x, y)) * e'" + using \<open>0 < e'\<close> nz + by (auto simp: e_def) + finally show "norm ((f x' y' - f x y - (fx x y (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" + by (auto simp: divide_simps dist_norm mult.commute) + qed + then show ?case + by eventually_elim (auto simp: dist_norm field_simps) +qed (auto intro!: bounded_linear_intros simp: split_beta') + + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Determinants.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,1249 @@ +(* Title: HOL/Analysis/Determinants.thy + Author: Amine Chaieb, University of Cambridge +*) + +section \<open>Traces, Determinant of square matrices and some properties\<close> + +theory Determinants +imports + Cartesian_Euclidean_Space + "~~/src/HOL/Library/Permutations" +begin + +subsection \<open>Trace\<close> + +definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" + where "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)" + +lemma trace_0: "trace (mat 0) = 0" + by (simp add: trace_def mat_def) + +lemma trace_I: "trace (mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))" + by (simp add: trace_def mat_def) + +lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" + by (simp add: trace_def setsum.distrib) + +lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" + by (simp add: trace_def setsum_subtractf) + +lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)" + apply (simp add: trace_def matrix_matrix_mult_def) + apply (subst setsum.commute) + apply (simp add: mult.commute) + done + +text \<open>Definition of determinant.\<close> + +definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where + "det A = + setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) + {p. p permutes (UNIV :: 'n set)}" + +text \<open>A few general lemmas we need below.\<close> + +lemma setprod_permute: + assumes p: "p permutes S" + shows "setprod f S = setprod (f \<circ> p) S" + using assms by (fact setprod.permute) + +lemma setproduct_permute_nat_interval: + fixes m n :: nat + shows "p permutes {m..n} \<Longrightarrow> setprod f {m..n} = setprod (f \<circ> p) {m..n}" + by (blast intro!: setprod_permute) + +text \<open>Basic determinant properties.\<close> + +lemma det_transpose: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)" +proof - + let ?di = "\<lambda>A i j. A$i$j" + let ?U = "(UNIV :: 'n set)" + have fU: "finite ?U" by simp + { + fix p + assume p: "p \<in> {p. p permutes ?U}" + from p have pU: "p permutes ?U" + by blast + have sth: "sign (inv p) = sign p" + by (metis sign_inverse fU p mem_Collect_eq permutation_permutes) + from permutes_inj[OF pU] + have pi: "inj_on p ?U" + by (blast intro: subset_inj_on) + from permutes_image[OF pU] + have "setprod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U = + setprod (\<lambda>i. ?di (transpose A) i (inv p i)) (p ` ?U)" + by simp + also have "\<dots> = setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U" + unfolding setprod.reindex[OF pi] .. + also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U" + proof - + { + fix i + assume i: "i \<in> ?U" + from i permutes_inv_o[OF pU] permutes_in_image[OF pU] + have "((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) i = ?di A i (p i)" + unfolding transpose_def by (simp add: fun_eq_iff) + } + then show "setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U = + setprod (\<lambda>i. ?di A i (p i)) ?U" + by (auto intro: setprod.cong) + qed + finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U) = + of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" + using sth by simp + } + then show ?thesis + unfolding det_def + apply (subst setsum_permutations_inverse) + apply (rule setsum.cong) + apply (rule refl) + apply blast + done +qed + +lemma det_lowerdiagonal: + fixes A :: "'a::comm_ring_1^('n::{finite,wellorder})^('n::{finite,wellorder})" + assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0" + shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)" +proof - + let ?U = "UNIV:: 'n set" + let ?PU = "{p. p permutes ?U}" + let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)" + have fU: "finite ?U" + by simp + from finite_permutations[OF fU] have fPU: "finite ?PU" . + have id0: "{id} \<subseteq> ?PU" + by (auto simp add: permutes_id) + { + fix p + assume p: "p \<in> ?PU - {id}" + from p have pU: "p permutes ?U" and pid: "p \<noteq> id" + by blast+ + from permutes_natset_le[OF pU] pid obtain i where i: "p i > i" + by (metis not_le) + from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" + by blast + from setprod_zero[OF fU ex] have "?pp p = 0" + by simp + } + then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0" + by blast + from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis + unfolding det_def by (simp add: sign_id) +qed + +lemma det_upperdiagonal: + fixes A :: "'a::comm_ring_1^'n::{finite,wellorder}^'n::{finite,wellorder}" + assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0" + shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)" +proof - + let ?U = "UNIV:: 'n set" + let ?PU = "{p. p permutes ?U}" + let ?pp = "(\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))" + have fU: "finite ?U" + by simp + from finite_permutations[OF fU] have fPU: "finite ?PU" . + have id0: "{id} \<subseteq> ?PU" + by (auto simp add: permutes_id) + { + fix p + assume p: "p \<in> ?PU - {id}" + from p have pU: "p permutes ?U" and pid: "p \<noteq> id" + by blast+ + from permutes_natset_ge[OF pU] pid obtain i where i: "p i < i" + by (metis not_le) + from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" + by blast + from setprod_zero[OF fU ex] have "?pp p = 0" + by simp + } + then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0" + by blast + from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis + unfolding det_def by (simp add: sign_id) +qed + +lemma det_diagonal: + fixes A :: "'a::comm_ring_1^'n^'n" + assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0" + shows "det A = setprod (\<lambda>i. A$i$i) (UNIV::'n set)" +proof - + let ?U = "UNIV:: 'n set" + let ?PU = "{p. p permutes ?U}" + let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)" + have fU: "finite ?U" by simp + from finite_permutations[OF fU] have fPU: "finite ?PU" . + have id0: "{id} \<subseteq> ?PU" + by (auto simp add: permutes_id) + { + fix p + assume p: "p \<in> ?PU - {id}" + then have "p \<noteq> id" + by simp + then obtain i where i: "p i \<noteq> i" + unfolding fun_eq_iff by auto + from ld [OF i [symmetric]] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" + by blast + from setprod_zero [OF fU ex] have "?pp p = 0" + by simp + } + then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0" + by blast + from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis + unfolding det_def by (simp add: sign_id) +qed + +lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1" +proof - + let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n" + let ?U = "UNIV :: 'n set" + let ?f = "\<lambda>i j. ?A$i$j" + { + fix i + assume i: "i \<in> ?U" + have "?f i i = 1" + using i by (vector mat_def) + } + then have th: "setprod (\<lambda>i. ?f i i) ?U = setprod (\<lambda>x. 1) ?U" + by (auto intro: setprod.cong) + { + fix i j + assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i \<noteq> j" + have "?f i j = 0" using i j ij + by (vector mat_def) + } + then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" + using det_diagonal by blast + also have "\<dots> = 1" + unfolding th setprod.neutral_const .. + finally show ?thesis . +qed + +lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0" + by (simp add: det_def setprod_zero) + +lemma det_permute_rows: + fixes A :: "'a::comm_ring_1^'n^'n" + assumes p: "p permutes (UNIV :: 'n::finite set)" + shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" + apply (simp add: det_def setsum_right_distrib mult.assoc[symmetric]) + apply (subst sum_permutations_compose_right[OF p]) +proof (rule setsum.cong) + let ?U = "UNIV :: 'n set" + let ?PU = "{p. p permutes ?U}" + fix q + assume qPU: "q \<in> ?PU" + have fU: "finite ?U" + by simp + from qPU have q: "q permutes ?U" + by blast + from p q have pp: "permutation p" and qp: "permutation q" + by (metis fU permutation_permutes)+ + from permutes_inv[OF p] have ip: "inv p permutes ?U" . + have "setprod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = setprod ((\<lambda>i. A$p i$(q \<circ> p) i) \<circ> inv p) ?U" + by (simp only: setprod_permute[OF ip, symmetric]) + also have "\<dots> = setprod (\<lambda>i. A $ (p \<circ> inv p) i $ (q \<circ> (p \<circ> inv p)) i) ?U" + by (simp only: o_def) + also have "\<dots> = setprod (\<lambda>i. A$i$q i) ?U" + by (simp only: o_def permutes_inverses[OF p]) + finally have thp: "setprod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U" + by blast + show "of_int (sign (q \<circ> p)) * setprod (\<lambda>i. A$ p i$ (q \<circ> p) i) ?U = + of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U" + by (simp only: thp sign_compose[OF qp pp] mult.commute of_int_mult) +qed rule + +lemma det_permute_columns: + fixes A :: "'a::comm_ring_1^'n^'n" + assumes p: "p permutes (UNIV :: 'n set)" + shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A" +proof - + let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n" + let ?At = "transpose A" + have "of_int (sign p) * det A = det (transpose (\<chi> i. transpose A $ p i))" + unfolding det_permute_rows[OF p, of ?At] det_transpose .. + moreover + have "?Ap = transpose (\<chi> i. transpose A $ p i)" + by (simp add: transpose_def vec_eq_iff) + ultimately show ?thesis + by simp +qed + +lemma det_identical_rows: + fixes A :: "'a::linordered_idom^'n^'n" + assumes ij: "i \<noteq> j" + and r: "row i A = row j A" + shows "det A = 0" +proof- + have tha: "\<And>(a::'a) b. a = b \<Longrightarrow> b = - a \<Longrightarrow> a = 0" + by simp + have th1: "of_int (-1) = - 1" by simp + let ?p = "Fun.swap i j id" + let ?A = "\<chi> i. A $ ?p i" + from r have "A = ?A" by (simp add: vec_eq_iff row_def Fun.swap_def) + then have "det A = det ?A" by simp + moreover have "det A = - det ?A" + by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1) + ultimately show "det A = 0" by (metis tha) +qed + +lemma det_identical_columns: + fixes A :: "'a::linordered_idom^'n^'n" + assumes ij: "i \<noteq> j" + and r: "column i A = column j A" + shows "det A = 0" + apply (subst det_transpose[symmetric]) + apply (rule det_identical_rows[OF ij]) + apply (metis row_transpose r) + done + +lemma det_zero_row: + fixes A :: "'a::{idom, ring_char_0}^'n^'n" + assumes r: "row i A = 0" + shows "det A = 0" + using r + apply (simp add: row_def det_def vec_eq_iff) + apply (rule setsum.neutral) + apply (auto simp: sign_nz) + done + +lemma det_zero_column: + fixes A :: "'a::{idom,ring_char_0}^'n^'n" + assumes r: "column i A = 0" + shows "det A = 0" + apply (subst det_transpose[symmetric]) + apply (rule det_zero_row [of i]) + apply (metis row_transpose r) + done + +lemma det_row_add: + fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n" + shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) = + det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + + det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" + unfolding det_def vec_lambda_beta setsum.distrib[symmetric] +proof (rule setsum.cong) + let ?U = "UNIV :: 'n set" + let ?pU = "{p. p permutes ?U}" + let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" + let ?g = "(\<lambda> i. if i = k then a i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" + let ?h = "(\<lambda> i. if i = k then b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" + fix p + assume p: "p \<in> ?pU" + let ?Uk = "?U - {k}" + from p have pU: "p permutes ?U" + by blast + have kU: "?U = insert k ?Uk" + by blast + { + fix j + assume j: "j \<in> ?Uk" + from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j" + by simp_all + } + then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk" + and th2: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?h i $ p i) ?Uk" + apply - + apply (rule setprod.cong, simp_all)+ + done + have th3: "finite ?Uk" "k \<notin> ?Uk" + by auto + have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" + unfolding kU[symmetric] .. + also have "\<dots> = ?f k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" + apply (rule setprod.insert) + apply simp + apply blast + done + also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" + by (simp add: field_simps) + also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" + by (metis th1 th2) + also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)" + unfolding setprod.insert[OF th3] by simp + finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U" + unfolding kU[symmetric] . + then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = + of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U" + by (simp add: field_simps) +qed rule + +lemma det_row_mul: + fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n" + shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) = + c * det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" + unfolding det_def vec_lambda_beta setsum_right_distrib +proof (rule setsum.cong) + let ?U = "UNIV :: 'n set" + let ?pU = "{p. p permutes ?U}" + let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n" + let ?g = "(\<lambda> i. if i = k then a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n" + fix p + assume p: "p \<in> ?pU" + let ?Uk = "?U - {k}" + from p have pU: "p permutes ?U" + by blast + have kU: "?U = insert k ?Uk" + by blast + { + fix j + assume j: "j \<in> ?Uk" + from j have "?f j $ p j = ?g j $ p j" + by simp + } + then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk" + apply - + apply (rule setprod.cong) + apply simp_all + done + have th3: "finite ?Uk" "k \<notin> ?Uk" + by auto + have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" + unfolding kU[symmetric] .. + also have "\<dots> = ?f k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" + apply (rule setprod.insert) + apply simp + apply blast + done + also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" + by (simp add: field_simps) + also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)" + unfolding th1 by (simp add: ac_simps) + also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))" + unfolding setprod.insert[OF th3] by simp + finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" + unfolding kU[symmetric] . + then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = + c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)" + by (simp add: field_simps) +qed rule + +lemma det_row_0: + fixes b :: "'n::finite \<Rightarrow> _ ^ 'n" + shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0" + using det_row_mul[of k 0 "\<lambda>i. 1" b] + apply simp + apply (simp only: vector_smult_lzero) + done + +lemma det_row_operation: + fixes A :: "'a::linordered_idom^'n^'n" + assumes ij: "i \<noteq> j" + shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A" +proof - + let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n" + have th: "row i ?Z = row j ?Z" by (vector row_def) + have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A" + by (vector row_def) + show ?thesis + unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2 + by simp +qed + +lemma det_row_span: + fixes A :: "real^'n^'n" + assumes x: "x \<in> span {row j A |j. j \<noteq> i}" + shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A" +proof - + let ?U = "UNIV :: 'n set" + let ?S = "{row j A |j. j \<noteq> i}" + let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)" + let ?P = "\<lambda>x. ?d (row i A + x) = det A" + { + fix k + have "(if k = i then row i A + 0 else row k A) = row k A" + by simp + } + then have P0: "?P 0" + apply - + apply (rule cong[of det, OF refl]) + apply (vector row_def) + done + moreover + { + fix c z y + assume zS: "z \<in> ?S" and Py: "?P y" + from zS obtain j where j: "z = row j A" "i \<noteq> j" + by blast + let ?w = "row i A + y" + have th0: "row i A + (c*s z + y) = ?w + c*s z" + by vector + have thz: "?d z = 0" + apply (rule det_identical_rows[OF j(2)]) + using j + apply (vector row_def) + done + have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" + unfolding th0 .. + then have "?P (c*s z + y)" + unfolding thz Py det_row_mul[of i] det_row_add[of i] + by simp + } + ultimately show ?thesis + apply - + apply (rule span_induct_alt[of ?P ?S, OF P0, folded scalar_mult_eq_scaleR]) + apply blast + apply (rule x) + done +qed + +text \<open> + May as well do this, though it's a bit unsatisfactory since it ignores + exact duplicates by considering the rows/columns as a set. +\<close> + +lemma det_dependent_rows: + fixes A:: "real^'n^'n" + assumes d: "dependent (rows A)" + shows "det A = 0" +proof - + let ?U = "UNIV :: 'n set" + from d obtain i where i: "row i A \<in> span (rows A - {row i A})" + unfolding dependent_def rows_def by blast + { + fix j k + assume jk: "j \<noteq> k" and c: "row j A = row k A" + from det_identical_rows[OF jk c] have ?thesis . + } + moreover + { + assume H: "\<And> i j. i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A" + have th0: "- row i A \<in> span {row j A|j. j \<noteq> i}" + apply (rule span_neg) + apply (rule set_rev_mp) + apply (rule i) + apply (rule span_mono) + using H i + apply (auto simp add: rows_def) + done + from det_row_span[OF th0] + have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)" + unfolding right_minus vector_smult_lzero .. + with det_row_mul[of i "0::real" "\<lambda>i. 1"] + have "det A = 0" by simp + } + ultimately show ?thesis by blast +qed + +lemma det_dependent_columns: + assumes d: "dependent (columns (A::real^'n^'n))" + shows "det A = 0" + by (metis d det_dependent_rows rows_transpose det_transpose) + +text \<open>Multilinearity and the multiplication formula.\<close> + +lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)" + by (rule iffD1[OF vec_lambda_unique]) vector + +lemma det_linear_row_setsum: + assumes fS: "finite S" + shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = + setsum (\<lambda>j. det ((\<chi> i. if i = k then a i j else c i)::'a^'n^'n)) S" +proof (induct rule: finite_induct[OF fS]) + case 1 + then show ?case + apply simp + unfolding setsum.empty det_row_0[of k] + apply rule + done +next + case (2 x F) + then show ?case + by (simp add: det_row_add cong del: if_weak_cong) +qed + +lemma finite_bounded_functions: + assumes fS: "finite S" + shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}" +proof (induct k) + case 0 + have th: "{f. \<forall>i. f i = i} = {id}" + by auto + show ?case + by (auto simp add: th) +next + case (Suc k) + let ?f = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i" + let ?S = "?f ` (S \<times> {f. (\<forall>i\<in>{1..k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)})" + have "?S = {f. (\<forall>i\<in>{1.. Suc k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. Suc k} \<longrightarrow> f i = i)}" + apply (auto simp add: image_iff) + apply (rule_tac x="x (Suc k)" in bexI) + apply (rule_tac x = "\<lambda>i. if i = Suc k then i else x i" in exI) + apply auto + done + with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f] + show ?case + by metis +qed + + +lemma det_linear_rows_setsum_lemma: + assumes fS: "finite S" + and fT: "finite T" + shows "det ((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n) = + setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)) + {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}" + using fT +proof (induct T arbitrary: a c set: finite) + case empty + have th0: "\<And>x y. (\<chi> i. if i \<in> {} then x i else y i) = (\<chi> i. y i)" + by vector + from empty.prems show ?case + unfolding th0 by (simp add: eq_id_iff) +next + case (insert z T a c) + let ?F = "\<lambda>T. {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}" + let ?h = "\<lambda>(y,g) i. if i = z then y else g i" + let ?k = "\<lambda>h. (h(z),(\<lambda>i. if i = z then i else h i))" + let ?s = "\<lambda> k a c f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)" + let ?c = "\<lambda>j i. if i = z then a i j else c i" + have thif: "\<And>a b c d. (if a \<or> b then c else d) = (if a then c else if b then c else d)" + by simp + have thif2: "\<And>a b c d e. (if a then b else if c then d else e) = + (if c then (if a then b else d) else (if a then b else e))" + by simp + from \<open>z \<notin> T\<close> have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False" + by auto + have "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) = + det (\<chi> i. if i = z then setsum (a i) S else if i \<in> T then setsum (a i) S else c i)" + unfolding insert_iff thif .. + also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then setsum (a i) S else if i = z then a i j else c i))" + unfolding det_linear_row_setsum[OF fS] + apply (subst thif2) + using nz + apply (simp cong del: if_weak_cong cong add: if_cong) + done + finally have tha: + "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) = + (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i) + else if i = z then a i j + else c i))" + unfolding insert.hyps unfolding setsum.cartesian_product by blast + show ?case unfolding tha + using \<open>z \<notin> T\<close> + by (intro setsum.reindex_bij_witness[where i="?k" and j="?h"]) + (auto intro!: cong[OF refl[of det]] simp: vec_eq_iff) +qed + +lemma det_linear_rows_setsum: + fixes S :: "'n::finite set" + assumes fS: "finite S" + shows "det (\<chi> i. setsum (a i) S) = + setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \<forall>i. f i \<in> S}" +proof - + have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" + by vector + from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] + show ?thesis by simp +qed + +lemma matrix_mul_setsum_alt: + fixes A B :: "'a::comm_ring_1^'n^'n" + shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))" + by (vector matrix_matrix_mult_def setsum_component) + +lemma det_rows_mul: + "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) = + setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)" +proof (simp add: det_def setsum_right_distrib cong add: setprod.cong, rule setsum.cong) + let ?U = "UNIV :: 'n set" + let ?PU = "{p. p permutes ?U}" + fix p + assume pU: "p \<in> ?PU" + let ?s = "of_int (sign p)" + from pU have p: "p permutes ?U" + by blast + have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U" + unfolding setprod.distrib .. + then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) = + setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" + by (simp add: field_simps) +qed rule + +lemma det_mul: + fixes A B :: "'a::linordered_idom^'n^'n" + shows "det (A ** B) = det A * det B" +proof - + let ?U = "UNIV :: 'n set" + let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}" + let ?PU = "{p. p permutes ?U}" + have fU: "finite ?U" + by simp + have fF: "finite ?F" + by (rule finite) + { + fix p + assume p: "p permutes ?U" + have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p] + using p[unfolded permutes_def] by simp + } + then have PUF: "?PU \<subseteq> ?F" by blast + { + fix f + assume fPU: "f \<in> ?F - ?PU" + have fUU: "f ` ?U \<subseteq> ?U" + using fPU by auto + from fPU have f: "\<forall>i \<in> ?U. f i \<in> ?U" "\<forall>i. i \<notin> ?U \<longrightarrow> f i = i" "\<not>(\<forall>y. \<exists>!x. f x = y)" + unfolding permutes_def by auto + + let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n" + let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n" + { + assume fni: "\<not> inj_on f ?U" + then obtain i j where ij: "f i = f j" "i \<noteq> j" + unfolding inj_on_def by blast + from ij + have rth: "row i ?B = row j ?B" + by (vector row_def) + from det_identical_rows[OF ij(2) rth] + have "det (\<chi> i. A$i$f i *s B$f i) = 0" + unfolding det_rows_mul by simp + } + moreover + { + assume fi: "inj_on f ?U" + from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j" + unfolding inj_on_def by metis + note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]] + { + fix y + from fs f have "\<exists>x. f x = y" + by blast + then obtain x where x: "f x = y" + by blast + { + fix z + assume z: "f z = y" + from fith x z have "z = x" + by metis + } + with x have "\<exists>!x. f x = y" + by blast + } + with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0" + by blast + } + ultimately have "det (\<chi> i. A$i$f i *s B$f i) = 0" + by blast + } + then have zth: "\<forall> f\<in> ?F - ?PU. det (\<chi> i. A$i$f i *s B$f i) = 0" + by simp + { + fix p + assume pU: "p \<in> ?PU" + from pU have p: "p permutes ?U" + by blast + let ?s = "\<lambda>p. of_int (sign p)" + let ?f = "\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))" + have "(setsum (\<lambda>q. ?s q * + (\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = + (setsum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)" + unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] + proof (rule setsum.cong) + fix q + assume qU: "q \<in> ?PU" + then have q: "q permutes ?U" + by blast + from p q have pp: "permutation p" and pq: "permutation q" + unfolding permutation_permutes by auto + have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" + "\<And>a. of_int (sign p) * (of_int (sign p) * a) = a" + unfolding mult.assoc[symmetric] + unfolding of_int_mult[symmetric] + by (simp_all add: sign_idempotent) + have ths: "?s q = ?s p * ?s (q \<circ> inv p)" + using pp pq permutation_inverse[OF pp] sign_inverse[OF pp] + by (simp add: th00 ac_simps sign_idempotent sign_compose) + have th001: "setprod (\<lambda>i. B$i$ q (inv p i)) ?U = setprod ((\<lambda>i. B$i$ q (inv p i)) \<circ> p) ?U" + by (rule setprod_permute[OF p]) + have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = + setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U" + unfolding th001 setprod.distrib[symmetric] o_def permutes_inverses[OF p] + apply (rule setprod.cong[OF refl]) + using permutes_in_image[OF q] + apply vector + done + show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = + ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q \<circ> inv p) * setprod (\<lambda>i. B$i$(q \<circ> inv p) i) ?U)" + using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp] + by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose) + qed rule + } + then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B" + unfolding det_def setsum_product + by (rule setsum.cong [OF refl]) + have "det (A**B) = setsum (\<lambda>f. det (\<chi> i. A $ i $ f i *s B $ f i)) ?F" + unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] + by simp + also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU" + using setsum.mono_neutral_cong_left[OF fF PUF zth, symmetric] + unfolding det_rows_mul by auto + finally show ?thesis unfolding th2 . +qed + +text \<open>Relation to invertibility.\<close> + +lemma invertible_left_inverse: + fixes A :: "real^'n^'n" + shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)" + by (metis invertible_def matrix_left_right_inverse) + +lemma invertible_righ_inverse: + fixes A :: "real^'n^'n" + shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)" + by (metis invertible_def matrix_left_right_inverse) + +lemma invertible_det_nz: + fixes A::"real ^'n^'n" + shows "invertible A \<longleftrightarrow> det A \<noteq> 0" +proof - + { + assume "invertible A" + then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1" + unfolding invertible_righ_inverse by blast + then have "det (A ** B) = det (mat 1 :: real ^'n^'n)" + by simp + then have "det A \<noteq> 0" + by (simp add: det_mul det_I) algebra + } + moreover + { + assume H: "\<not> invertible A" + let ?U = "UNIV :: 'n set" + have fU: "finite ?U" + by simp + from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0" + and iU: "i \<in> ?U" + and ci: "c i \<noteq> 0" + unfolding invertible_righ_inverse + unfolding matrix_right_invertible_independent_rows + by blast + have *: "\<And>(a::real^'n) b. a + b = 0 \<Longrightarrow> -a = b" + apply (drule_tac f="op + (- a)" in cong[OF refl]) + apply (simp only: ab_left_minus add.assoc[symmetric]) + apply simp + done + from c ci + have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})" + unfolding setsum.remove[OF fU iU] setsum_cmul + apply - + apply (rule vector_mul_lcancel_imp[OF ci]) + apply (auto simp add: field_simps) + unfolding * + apply rule + done + have thr: "- row i A \<in> span {row j A| j. j \<noteq> i}" + unfolding thr0 + apply (rule span_setsum) + apply simp + apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+ + apply (rule span_superset) + apply auto + done + let ?B = "(\<chi> k. if k = i then 0 else row k A) :: real ^'n^'n" + have thrb: "row i ?B = 0" using iU by (vector row_def) + have "det A = 0" + unfolding det_row_span[OF thr, symmetric] right_minus + unfolding det_zero_row[OF thrb] .. + } + ultimately show ?thesis + by blast +qed + +text \<open>Cramer's rule.\<close> + +lemma cramer_lemma_transpose: + fixes A:: "real^'n^'n" + and x :: "real^'n" + shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set) + else row i A)::real^'n^'n) = x$k * det A" + (is "?lhs = ?rhs") +proof - + let ?U = "UNIV :: 'n set" + let ?Uk = "?U - {k}" + have U: "?U = insert k ?Uk" + by blast + have fUk: "finite ?Uk" + by simp + have kUk: "k \<notin> ?Uk" + by simp + have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s" + by (vector field_simps) + have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" + by auto + have "(\<chi> i. row i A) = A" by (vector row_def) + then have thd1: "det (\<chi> i. row i A) = det A" + by simp + have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A" + apply (rule det_row_span) + apply (rule span_setsum) + apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+ + apply (rule span_superset) + apply auto + done + show "?lhs = x$k * det A" + apply (subst U) + unfolding setsum.insert[OF fUk kUk] + apply (subst th00) + unfolding add.assoc + apply (subst det_row_add) + unfolding thd0 + unfolding det_row_mul + unfolding th001[of k "\<lambda>i. row i A"] + unfolding thd1 + apply (simp add: field_simps) + done +qed + +lemma cramer_lemma: + fixes A :: "real^'n^'n" + shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: real^'n^'n) = x$k * det A" +proof - + let ?U = "UNIV :: 'n set" + have *: "\<And>c. setsum (\<lambda>i. c i *s row i (transpose A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U" + by (auto simp add: row_transpose intro: setsum.cong) + show ?thesis + unfolding matrix_mult_vsum + unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric] + unfolding *[of "\<lambda>i. x$i"] + apply (subst det_transpose[symmetric]) + apply (rule cong[OF refl[of det]]) + apply (vector transpose_def column_def row_def) + done +qed + +lemma cramer: + fixes A ::"real^'n^'n" + assumes d0: "det A \<noteq> 0" + shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)" +proof - + from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" + unfolding invertible_det_nz[symmetric] invertible_def + by blast + have "(A ** B) *v b = b" + by (simp add: B matrix_vector_mul_lid) + then have "A *v (B *v b) = b" + by (simp add: matrix_vector_mul_assoc) + then have xe: "\<exists>x. A *v x = b" + by blast + { + fix x + assume x: "A *v x = b" + have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)" + unfolding x[symmetric] + using d0 by (simp add: vec_eq_iff cramer_lemma field_simps) + } + with xe show ?thesis + by auto +qed + +text \<open>Orthogonality of a transformation and matrix.\<close> + +definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)" + +lemma orthogonal_transformation: + "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^_). norm (f v) = norm v)" + unfolding orthogonal_transformation_def + apply auto + apply (erule_tac x=v in allE)+ + apply (simp add: norm_eq_sqrt_inner) + apply (simp add: dot_norm linear_add[symmetric]) + done + +definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> + transpose Q ** Q = mat 1 \<and> Q ** transpose Q = mat 1" + +lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1" + by (metis matrix_left_right_inverse orthogonal_matrix_def) + +lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)" + by (simp add: orthogonal_matrix_def transpose_mat matrix_mul_lid) + +lemma orthogonal_matrix_mul: + fixes A :: "real ^'n^'n" + assumes oA : "orthogonal_matrix A" + and oB: "orthogonal_matrix B" + shows "orthogonal_matrix(A ** B)" + using oA oB + unfolding orthogonal_matrix matrix_transpose_mul + apply (subst matrix_mul_assoc) + apply (subst matrix_mul_assoc[symmetric]) + apply (simp add: matrix_mul_rid) + done + +lemma orthogonal_transformation_matrix: + fixes f:: "real^'n \<Rightarrow> real^'n" + shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)" + (is "?lhs \<longleftrightarrow> ?rhs") +proof - + let ?mf = "matrix f" + let ?ot = "orthogonal_transformation f" + let ?U = "UNIV :: 'n set" + have fU: "finite ?U" by simp + let ?m1 = "mat 1 :: real ^'n^'n" + { + assume ot: ?ot + from ot have lf: "linear f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w" + unfolding orthogonal_transformation_def orthogonal_matrix by blast+ + { + fix i j + let ?A = "transpose ?mf ** ?mf" + have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" + "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" + by simp_all + from fd[rule_format, of "axis i 1" "axis j 1", + simplified matrix_works[OF lf, symmetric] dot_matrix_vector_mul] + have "?A$i$j = ?m1 $ i $ j" + by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def + th0 setsum.delta[OF fU] mat_def axis_def) + } + then have "orthogonal_matrix ?mf" + unfolding orthogonal_matrix + by vector + with lf have ?rhs + by blast + } + moreover + { + assume lf: "linear f" and om: "orthogonal_matrix ?mf" + from lf om have ?lhs + apply (simp only: orthogonal_matrix_def norm_eq orthogonal_transformation) + apply (simp only: matrix_works[OF lf, symmetric]) + apply (subst dot_matrix_vector_mul) + apply (simp add: dot_matrix_product matrix_mul_lid) + done + } + ultimately show ?thesis + by blast +qed + +lemma det_orthogonal_matrix: + fixes Q:: "'a::linordered_idom^'n^'n" + assumes oQ: "orthogonal_matrix Q" + shows "det Q = 1 \<or> det Q = - 1" +proof - + have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x") + proof - + fix x:: 'a + have th0: "x * x - 1 = (x - 1) * (x + 1)" + by (simp add: field_simps) + have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0" + apply (subst eq_iff_diff_eq_0) + apply simp + done + have "x * x = 1 \<longleftrightarrow> x * x - 1 = 0" + by simp + also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1" + unfolding th0 th1 by simp + finally show "?ths x" .. + qed + from oQ have "Q ** transpose Q = mat 1" + by (metis orthogonal_matrix_def) + then have "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" + by simp + then have "det Q * det Q = 1" + by (simp add: det_mul det_I det_transpose) + then show ?thesis unfolding th . +qed + +text \<open>Linearity of scaling, and hence isometry, that preserves origin.\<close> + +lemma scaling_linear: + fixes f :: "real ^'n \<Rightarrow> real ^'n" + assumes f0: "f 0 = 0" + and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y" + shows "linear f" +proof - + { + fix v w + { + fix x + note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] + } + note th0 = this + have "f v \<bullet> f w = c\<^sup>2 * (v \<bullet> w)" + unfolding dot_norm_neg dist_norm[symmetric] + unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)} + note fc = this + show ?thesis + unfolding linear_iff vector_eq[where 'a="real^'n"] scalar_mult_eq_scaleR + by (simp add: inner_add fc field_simps) +qed + +lemma isometry_linear: + "f (0:: real^'n) = (0:: real^'n) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y \<Longrightarrow> linear f" + by (rule scaling_linear[where c=1]) simp_all + +text \<open>Hence another formulation of orthogonal transformation.\<close> + +lemma orthogonal_transformation_isometry: + "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)" + unfolding orthogonal_transformation + apply (rule iffI) + apply clarify + apply (clarsimp simp add: linear_0 linear_diff[symmetric] dist_norm) + apply (rule conjI) + apply (rule isometry_linear) + apply simp + apply simp + apply clarify + apply (erule_tac x=v in allE) + apply (erule_tac x=0 in allE) + apply (simp add: dist_norm) + done + +text \<open>Can extend an isometry from unit sphere.\<close> + +lemma isometry_sphere_extend: + fixes f:: "real ^'n \<Rightarrow> real ^'n" + assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1" + and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y" + shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)" +proof - + { + fix x y x' y' x0 y0 x0' y0' :: "real ^'n" + assume H: + "x = norm x *\<^sub>R x0" + "y = norm y *\<^sub>R y0" + "x' = norm x *\<^sub>R x0'" "y' = norm y *\<^sub>R y0'" + "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1" + "norm(x0' - y0') = norm(x0 - y0)" + then have *: "x0 \<bullet> y0 = x0' \<bullet> y0' + y0' \<bullet> x0' - y0 \<bullet> x0 " + by (simp add: norm_eq norm_eq_1 inner_add inner_diff) + have "norm(x' - y') = norm(x - y)" + apply (subst H(1)) + apply (subst H(2)) + apply (subst H(3)) + apply (subst H(4)) + using H(5-9) + apply (simp add: norm_eq norm_eq_1) + apply (simp add: inner_diff scalar_mult_eq_scaleR) + unfolding * + apply (simp add: field_simps) + done + } + note th0 = this + let ?g = "\<lambda>x. if x = 0 then 0 else norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" + { + fix x:: "real ^'n" + assume nx: "norm x = 1" + have "?g x = f x" + using nx by auto + } + then have thfg: "\<forall>x. norm x = 1 \<longrightarrow> ?g x = f x" + by blast + have g0: "?g 0 = 0" + by simp + { + fix x y :: "real ^'n" + { + assume "x = 0" "y = 0" + then have "dist (?g x) (?g y) = dist x y" + by simp + } + moreover + { + assume "x = 0" "y \<noteq> 0" + then have "dist (?g x) (?g y) = dist x y" + apply (simp add: dist_norm) + apply (rule f1[rule_format]) + apply (simp add: field_simps) + done + } + moreover + { + assume "x \<noteq> 0" "y = 0" + then have "dist (?g x) (?g y) = dist x y" + apply (simp add: dist_norm) + apply (rule f1[rule_format]) + apply (simp add: field_simps) + done + } + moreover + { + assume z: "x \<noteq> 0" "y \<noteq> 0" + have th00: + "x = norm x *\<^sub>R (inverse (norm x) *\<^sub>R x)" + "y = norm y *\<^sub>R (inverse (norm y) *\<^sub>R y)" + "norm x *\<^sub>R f ((inverse (norm x) *\<^sub>R x)) = norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" + "norm y *\<^sub>R f (inverse (norm y) *\<^sub>R y) = norm y *\<^sub>R f (inverse (norm y) *\<^sub>R y)" + "norm (inverse (norm x) *\<^sub>R x) = 1" + "norm (f (inverse (norm x) *\<^sub>R x)) = 1" + "norm (inverse (norm y) *\<^sub>R y) = 1" + "norm (f (inverse (norm y) *\<^sub>R y)) = 1" + "norm (f (inverse (norm x) *\<^sub>R x) - f (inverse (norm y) *\<^sub>R y)) = + norm (inverse (norm x) *\<^sub>R x - inverse (norm y) *\<^sub>R y)" + using z + by (auto simp add: field_simps intro: f1[rule_format] fd1[rule_format, unfolded dist_norm]) + from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" + by (simp add: dist_norm) + } + ultimately have "dist (?g x) (?g y) = dist x y" + by blast + } + note thd = this + show ?thesis + apply (rule exI[where x= ?g]) + unfolding orthogonal_transformation_isometry + using g0 thfg thd + apply metis + done +qed + +text \<open>Rotation, reflection, rotoinversion.\<close> + +definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1" +definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1" + +lemma orthogonal_rotation_or_rotoinversion: + fixes Q :: "'a::linordered_idom^'n^'n" + shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q" + by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) + +text \<open>Explicit formulas for low dimensions.\<close> + +lemma setprod_neutral_const: "setprod f {(1::nat)..1} = f 1" + by simp + +lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" + by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute) + +lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" + by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute) + +lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" + by (simp add: det_def of_nat_Suc sign_id) + +lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" +proof - + have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto + show ?thesis + unfolding det_def UNIV_2 + unfolding setsum_over_permutations_insert[OF f12] + unfolding permutes_sing + by (simp add: sign_swap_id sign_id swap_id_eq) +qed + +lemma det_3: + "det (A::'a::comm_ring_1^3^3) = + A$1$1 * A$2$2 * A$3$3 + + A$1$2 * A$2$3 * A$3$1 + + A$1$3 * A$2$1 * A$3$2 - + A$1$1 * A$2$3 * A$3$2 - + A$1$2 * A$2$1 * A$3$3 - + A$1$3 * A$2$2 * A$3$1" +proof - + have f123: "finite {2::3, 3}" "1 \<notin> {2::3, 3}" + by auto + have f23: "finite {3::3}" "2 \<notin> {3::3}" + by auto + + show ?thesis + unfolding det_def UNIV_3 + unfolding setsum_over_permutations_insert[OF f123] + unfolding setsum_over_permutations_insert[OF f23] + unfolding permutes_sing + by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Embed_Measure.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,398 @@ +(* Title: HOL/Analysis/Embed_Measure.thy + Author: Manuel Eberl, TU München + + Defines measure embeddings with injective functions, i.e. lifting a measure on some values + to a measure on "tagged" values (e.g. embed_measure M Inl lifts a measure on values 'a to a + measure on the left part of the sum type 'a + 'b) +*) + +section \<open>Embed Measure Spaces with a Function\<close> + +theory Embed_Measure +imports Binary_Product_Measure +begin + +definition embed_measure :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where + "embed_measure M f = measure_of (f ` space M) {f ` A |A. A \<in> sets M} + (\<lambda>A. emeasure M (f -` A \<inter> space M))" + +lemma space_embed_measure: "space (embed_measure M f) = f ` space M" + unfolding embed_measure_def + by (subst space_measure_of) (auto dest: sets.sets_into_space) + +lemma sets_embed_measure': + assumes inj: "inj_on f (space M)" + shows "sets (embed_measure M f) = {f ` A |A. A \<in> sets M}" + unfolding embed_measure_def +proof (intro sigma_algebra.sets_measure_of_eq sigma_algebra_iff2[THEN iffD2] conjI allI ballI impI) + fix s assume "s \<in> {f ` A |A. A \<in> sets M}" + then obtain s' where s'_props: "s = f ` s'" "s' \<in> sets M" by auto + hence "f ` space M - s = f ` (space M - s')" using inj + by (auto dest: inj_onD sets.sets_into_space) + also have "... \<in> {f ` A |A. A \<in> sets M}" using s'_props by auto + finally show "f ` space M - s \<in> {f ` A |A. A \<in> sets M}" . +next + fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> {f ` A |A. A \<in> sets M}" + then obtain A' where A': "\<And>i. A i = f ` A' i" "\<And>i. A' i \<in> sets M" + by (auto simp: subset_eq choice_iff) + then have "(\<Union>x. f ` A' x) = f ` (\<Union>x. A' x)" by blast + with A' show "(\<Union>i. A i) \<in> {f ` A |A. A \<in> sets M}" + by simp blast +qed (auto dest: sets.sets_into_space) + +lemma the_inv_into_vimage: + "inj_on f X \<Longrightarrow> A \<subseteq> X \<Longrightarrow> the_inv_into X f -` A \<inter> (f`X) = f ` A" + by (auto simp: the_inv_into_f_f) + +lemma sets_embed_eq_vimage_algebra: + assumes "inj_on f (space M)" + shows "sets (embed_measure M f) = sets (vimage_algebra (f`space M) (the_inv_into (space M) f) M)" + by (auto simp: sets_embed_measure'[OF assms] Pi_iff the_inv_into_f_f assms sets_vimage_algebra2 simple_image + dest: sets.sets_into_space + intro!: image_cong the_inv_into_vimage[symmetric]) + +lemma sets_embed_measure: + assumes inj: "inj f" + shows "sets (embed_measure M f) = {f ` A |A. A \<in> sets M}" + using assms by (subst sets_embed_measure') (auto intro!: inj_onI dest: injD) + +lemma in_sets_embed_measure: "A \<in> sets M \<Longrightarrow> f ` A \<in> sets (embed_measure M f)" + unfolding embed_measure_def + by (intro in_measure_of) (auto dest: sets.sets_into_space) + +lemma measurable_embed_measure1: + assumes g: "(\<lambda>x. g (f x)) \<in> measurable M N" + shows "g \<in> measurable (embed_measure M f) N" + unfolding measurable_def +proof safe + fix A assume "A \<in> sets N" + with g have "(\<lambda>x. g (f x)) -` A \<inter> space M \<in> sets M" + by (rule measurable_sets) + then have "f ` ((\<lambda>x. g (f x)) -` A \<inter> space M) \<in> sets (embed_measure M f)" + by (rule in_sets_embed_measure) + also have "f ` ((\<lambda>x. g (f x)) -` A \<inter> space M) = g -` A \<inter> space (embed_measure M f)" + by (auto simp: space_embed_measure) + finally show "g -` A \<inter> space (embed_measure M f) \<in> sets (embed_measure M f)" . +qed (insert measurable_space[OF assms], auto simp: space_embed_measure) + +lemma measurable_embed_measure2': + assumes "inj_on f (space M)" + shows "f \<in> measurable M (embed_measure M f)" +proof- + { + fix A assume A: "A \<in> sets M" + also from A have "A = A \<inter> space M" by auto + also have "... = f -` f ` A \<inter> space M" using A assms + by (auto dest: inj_onD sets.sets_into_space) + finally have "f -` f ` A \<inter> space M \<in> sets M" . + } + thus ?thesis using assms unfolding embed_measure_def + by (intro measurable_measure_of) (auto dest: sets.sets_into_space) +qed + +lemma measurable_embed_measure2: + assumes [simp]: "inj f" shows "f \<in> measurable M (embed_measure M f)" + by (auto simp: inj_vimage_image_eq embed_measure_def + intro!: measurable_measure_of dest: sets.sets_into_space) + +lemma embed_measure_eq_distr': + assumes "inj_on f (space M)" + shows "embed_measure M f = distr M (embed_measure M f) f" +proof- + have "distr M (embed_measure M f) f = + measure_of (f ` space M) {f ` A |A. A \<in> sets M} + (\<lambda>A. emeasure M (f -` A \<inter> space M))" unfolding distr_def + by (simp add: space_embed_measure sets_embed_measure'[OF assms]) + also have "... = embed_measure M f" unfolding embed_measure_def .. + finally show ?thesis .. +qed + +lemma embed_measure_eq_distr: + "inj f \<Longrightarrow> embed_measure M f = distr M (embed_measure M f) f" + by (rule embed_measure_eq_distr') (auto intro!: inj_onI dest: injD) + +lemma nn_integral_embed_measure': + "inj_on f (space M) \<Longrightarrow> g \<in> borel_measurable (embed_measure M f) \<Longrightarrow> + nn_integral (embed_measure M f) g = nn_integral M (\<lambda>x. g (f x))" + apply (subst embed_measure_eq_distr', simp) + apply (subst nn_integral_distr) + apply (simp_all add: measurable_embed_measure2') + done + +lemma nn_integral_embed_measure: + "inj f \<Longrightarrow> g \<in> borel_measurable (embed_measure M f) \<Longrightarrow> + nn_integral (embed_measure M f) g = nn_integral M (\<lambda>x. g (f x))" + by(erule nn_integral_embed_measure'[OF subset_inj_on]) simp + +lemma emeasure_embed_measure': + assumes "inj_on f (space M)" "A \<in> sets (embed_measure M f)" + shows "emeasure (embed_measure M f) A = emeasure M (f -` A \<inter> space M)" + by (subst embed_measure_eq_distr'[OF assms(1)]) + (simp add: emeasure_distr[OF measurable_embed_measure2'[OF assms(1)] assms(2)]) + +lemma emeasure_embed_measure: + assumes "inj f" "A \<in> sets (embed_measure M f)" + shows "emeasure (embed_measure M f) A = emeasure M (f -` A \<inter> space M)" + using assms by (intro emeasure_embed_measure') (auto intro!: inj_onI dest: injD) + +lemma embed_measure_comp: + assumes [simp]: "inj f" "inj g" + shows "embed_measure (embed_measure M f) g = embed_measure M (g \<circ> f)" +proof- + have [simp]: "inj (\<lambda>x. g (f x))" by (subst o_def[symmetric]) (auto intro: inj_comp) + note measurable_embed_measure2[measurable] + have "embed_measure (embed_measure M f) g = + distr M (embed_measure (embed_measure M f) g) (g \<circ> f)" + by (subst (1 2) embed_measure_eq_distr) + (simp_all add: distr_distr sets_embed_measure cong: distr_cong) + also have "... = embed_measure M (g \<circ> f)" + by (subst (3) embed_measure_eq_distr, simp add: o_def, rule distr_cong) + (auto simp: sets_embed_measure o_def image_image[symmetric] + intro: inj_comp cong: distr_cong) + finally show ?thesis . +qed + +lemma sigma_finite_embed_measure: + assumes "sigma_finite_measure M" and inj: "inj f" + shows "sigma_finite_measure (embed_measure M f)" +proof - + from assms(1) interpret sigma_finite_measure M . + from sigma_finite_countable obtain A where + A_props: "countable A" "A \<subseteq> sets M" "\<Union>A = space M" "\<And>X. X\<in>A \<Longrightarrow> emeasure M X \<noteq> \<infinity>" by blast + from A_props have "countable (op ` f`A)" by auto + moreover + from inj and A_props have "op ` f`A \<subseteq> sets (embed_measure M f)" + by (auto simp: sets_embed_measure) + moreover + from A_props and inj have "\<Union>(op ` f`A) = space (embed_measure M f)" + by (auto simp: space_embed_measure intro!: imageI) + moreover + from A_props and inj have "\<forall>a\<in>op ` f ` A. emeasure (embed_measure M f) a \<noteq> \<infinity>" + by (intro ballI, subst emeasure_embed_measure) + (auto simp: inj_vimage_image_eq intro: in_sets_embed_measure) + ultimately show ?thesis by - (standard, blast) +qed + +lemma embed_measure_count_space': + "inj_on f A \<Longrightarrow> embed_measure (count_space A) f = count_space (f`A)" + apply (subst distr_bij_count_space[of f A "f`A", symmetric]) + apply (simp add: inj_on_def bij_betw_def) + apply (subst embed_measure_eq_distr') + apply simp + apply(auto 4 3 intro!: measure_eqI imageI simp add: sets_embed_measure' subset_image_iff) + apply (subst (1 2) emeasure_distr) + apply (auto simp: space_embed_measure sets_embed_measure') + done + +lemma embed_measure_count_space: + "inj f \<Longrightarrow> embed_measure (count_space A) f = count_space (f`A)" + by(rule embed_measure_count_space')(erule subset_inj_on, simp) + +lemma sets_embed_measure_alt: + "inj f \<Longrightarrow> sets (embed_measure M f) = (op`f) ` sets M" + by (auto simp: sets_embed_measure) + +lemma emeasure_embed_measure_image': + assumes "inj_on f (space M)" "X \<in> sets M" + shows "emeasure (embed_measure M f) (f`X) = emeasure M X" +proof- + from assms have "emeasure (embed_measure M f) (f`X) = emeasure M (f -` f ` X \<inter> space M)" + by (subst emeasure_embed_measure') (auto simp: sets_embed_measure') + also from assms have "f -` f ` X \<inter> space M = X" by (auto dest: inj_onD sets.sets_into_space) + finally show ?thesis . +qed + +lemma emeasure_embed_measure_image: + "inj f \<Longrightarrow> X \<in> sets M \<Longrightarrow> emeasure (embed_measure M f) (f`X) = emeasure M X" + by (simp_all add: emeasure_embed_measure in_sets_embed_measure inj_vimage_image_eq) + +lemma embed_measure_eq_iff: + assumes "inj f" + shows "embed_measure A f = embed_measure B f \<longleftrightarrow> A = B" (is "?M = ?N \<longleftrightarrow> _") +proof + from assms have I: "inj (op` f)" by (auto intro: injI dest: injD) + assume asm: "?M = ?N" + hence "sets (embed_measure A f) = sets (embed_measure B f)" by simp + with assms have "sets A = sets B" by (simp only: I inj_image_eq_iff sets_embed_measure_alt) + moreover { + fix X assume "X \<in> sets A" + from asm have "emeasure ?M (f`X) = emeasure ?N (f`X)" by simp + with \<open>X \<in> sets A\<close> and \<open>sets A = sets B\<close> and assms + have "emeasure A X = emeasure B X" by (simp add: emeasure_embed_measure_image) + } + ultimately show "A = B" by (rule measure_eqI) +qed simp + +lemma the_inv_into_in_Pi: "inj_on f A \<Longrightarrow> the_inv_into A f \<in> f ` A \<rightarrow> A" + by (auto simp: the_inv_into_f_f) + +lemma map_prod_image: "map_prod f g ` (A \<times> B) = (f`A) \<times> (g`B)" + using map_prod_surj_on[OF refl refl] . + +lemma map_prod_vimage: "map_prod f g -` (A \<times> B) = (f-`A) \<times> (g-`B)" + by auto + +lemma embed_measure_prod: + assumes f: "inj f" and g: "inj g" and [simp]: "sigma_finite_measure M" "sigma_finite_measure N" + shows "embed_measure M f \<Otimes>\<^sub>M embed_measure N g = embed_measure (M \<Otimes>\<^sub>M N) (\<lambda>(x, y). (f x, g y))" + (is "?L = _") + unfolding map_prod_def[symmetric] +proof (rule pair_measure_eqI) + have fg[simp]: "\<And>A. inj_on (map_prod f g) A" "\<And>A. inj_on f A" "\<And>A. inj_on g A" + using f g by (auto simp: inj_on_def) + + note complete_lattice_class.Sup_insert[simp del] ccSup_insert[simp del] SUP_insert[simp del] + ccSUP_insert[simp del] + show sets: "sets ?L = sets (embed_measure (M \<Otimes>\<^sub>M N) (map_prod f g))" + unfolding map_prod_def[symmetric] + apply (simp add: sets_pair_eq_sets_fst_snd sets_embed_eq_vimage_algebra + cong: vimage_algebra_cong) + apply (subst sets_vimage_Sup_eq[where Y="space (M \<Otimes>\<^sub>M N)"]) + apply (simp_all add: space_pair_measure[symmetric]) + apply (auto simp add: the_inv_into_f_f + simp del: map_prod_simp + del: prod_fun_imageE) [] + apply auto [] + apply (subst image_insert) + apply simp + apply (subst (1 2 3 4 ) vimage_algebra_vimage_algebra_eq) + apply (simp_all add: the_inv_into_in_Pi Pi_iff[of snd] Pi_iff[of fst] space_pair_measure) + apply (simp_all add: Pi_iff[of snd] Pi_iff[of fst] the_inv_into_in_Pi vimage_algebra_vimage_algebra_eq + space_pair_measure[symmetric] map_prod_image[symmetric]) + apply (intro arg_cong[where f=sets] arg_cong[where f=Sup] arg_cong2[where f=insert] vimage_algebra_cong) + apply (auto simp: map_prod_image the_inv_into_f_f + simp del: map_prod_simp del: prod_fun_imageE) + apply (simp_all add: the_inv_into_f_f space_pair_measure) + done + + note measurable_embed_measure2[measurable] + fix A B assume AB: "A \<in> sets (embed_measure M f)" "B \<in> sets (embed_measure N g)" + moreover have "f -` A \<times> g -` B \<inter> space (M \<Otimes>\<^sub>M N) = (f -` A \<inter> space M) \<times> (g -` B \<inter> space N)" + by (auto simp: space_pair_measure) + ultimately show "emeasure (embed_measure M f) A * emeasure (embed_measure N g) B = + emeasure (embed_measure (M \<Otimes>\<^sub>M N) (map_prod f g)) (A \<times> B)" + by (simp add: map_prod_vimage sets[symmetric] emeasure_embed_measure + sigma_finite_measure.emeasure_pair_measure_Times) +qed (insert assms, simp_all add: sigma_finite_embed_measure) + +lemma mono_embed_measure: + "space M = space M' \<Longrightarrow> sets M \<subseteq> sets M' \<Longrightarrow> sets (embed_measure M f) \<subseteq> sets (embed_measure M' f)" + unfolding embed_measure_def + apply (subst (1 2) sets_measure_of) + apply (blast dest: sets.sets_into_space) + apply (blast dest: sets.sets_into_space) + apply simp + apply (intro sigma_sets_mono') + apply safe + apply (simp add: subset_eq) + apply metis + done + +lemma density_embed_measure: + assumes inj: "inj f" and Mg[measurable]: "g \<in> borel_measurable (embed_measure M f)" + shows "density (embed_measure M f) g = embed_measure (density M (g \<circ> f)) f" (is "?M1 = ?M2") +proof (rule measure_eqI) + fix X assume X: "X \<in> sets ?M1" + from inj have Mf[measurable]: "f \<in> measurable M (embed_measure M f)" + by (rule measurable_embed_measure2) + from Mg and X have "emeasure ?M1 X = \<integral>\<^sup>+ x. g x * indicator X x \<partial>embed_measure M f" + by (subst emeasure_density) simp_all + also from X have "... = \<integral>\<^sup>+ x. g (f x) * indicator X (f x) \<partial>M" + by (subst embed_measure_eq_distr[OF inj], subst nn_integral_distr) auto + also have "... = \<integral>\<^sup>+ x. g (f x) * indicator (f -` X \<inter> space M) x \<partial>M" + by (intro nn_integral_cong) (auto split: split_indicator) + also from X have "... = emeasure (density M (g \<circ> f)) (f -` X \<inter> space M)" + by (subst emeasure_density) (simp_all add: measurable_comp[OF Mf Mg] measurable_sets[OF Mf]) + also from X and inj have "... = emeasure ?M2 X" + by (subst emeasure_embed_measure) (simp_all add: sets_embed_measure) + finally show "emeasure ?M1 X = emeasure ?M2 X" . +qed (simp_all add: sets_embed_measure inj) + +lemma density_embed_measure': + assumes inj: "inj f" and inv: "\<And>x. f' (f x) = x" and Mg[measurable]: "g \<in> borel_measurable M" + shows "density (embed_measure M f) (g \<circ> f') = embed_measure (density M g) f" +proof- + have "density (embed_measure M f) (g \<circ> f') = embed_measure (density M (g \<circ> f' \<circ> f)) f" + by (rule density_embed_measure[OF inj]) + (rule measurable_comp, rule measurable_embed_measure1, subst measurable_cong, + rule inv, rule measurable_ident_sets, simp, rule Mg) + also have "density M (g \<circ> f' \<circ> f) = density M g" + by (intro density_cong) (subst measurable_cong, simp add: o_def inv, simp_all add: Mg inv) + finally show ?thesis . +qed + +lemma inj_on_image_subset_iff: + assumes "inj_on f C" "A \<subseteq> C" "B \<subseteq> C" + shows "f ` A \<subseteq> f ` B \<longleftrightarrow> A \<subseteq> B" +proof (intro iffI subsetI) + fix x assume A: "f ` A \<subseteq> f ` B" and B: "x \<in> A" + from B have "f x \<in> f ` A" by blast + with A have "f x \<in> f ` B" by blast + then obtain y where "f x = f y" and "y \<in> B" by blast + with assms and B have "x = y" by (auto dest: inj_onD) + with \<open>y \<in> B\<close> show "x \<in> B" by simp +qed auto + + +lemma AE_embed_measure': + assumes inj: "inj_on f (space M)" + shows "(AE x in embed_measure M f. P x) \<longleftrightarrow> (AE x in M. P (f x))" +proof + let ?M = "embed_measure M f" + assume "AE x in ?M. P x" + then obtain A where A_props: "A \<in> sets ?M" "emeasure ?M A = 0" "{x\<in>space ?M. \<not>P x} \<subseteq> A" + by (force elim: AE_E) + then obtain A' where A'_props: "A = f ` A'" "A' \<in> sets M" by (auto simp: sets_embed_measure' inj) + moreover have B: "{x\<in>space ?M. \<not>P x} = f ` {x\<in>space M. \<not>P (f x)}" + by (auto simp: inj space_embed_measure) + from A_props(3) have "{x\<in>space M. \<not>P (f x)} \<subseteq> A'" + by (subst (asm) B, subst (asm) A'_props, subst (asm) inj_on_image_subset_iff[OF inj]) + (insert A'_props, auto dest: sets.sets_into_space) + moreover from A_props A'_props have "emeasure M A' = 0" + by (simp add: emeasure_embed_measure_image' inj) + ultimately show "AE x in M. P (f x)" by (intro AE_I) +next + let ?M = "embed_measure M f" + assume "AE x in M. P (f x)" + then obtain A where A_props: "A \<in> sets M" "emeasure M A = 0" "{x\<in>space M. \<not>P (f x)} \<subseteq> A" + by (force elim: AE_E) + hence "f`A \<in> sets ?M" "emeasure ?M (f`A) = 0" "{x\<in>space ?M. \<not>P x} \<subseteq> f`A" + by (auto simp: space_embed_measure emeasure_embed_measure_image' sets_embed_measure' inj) + thus "AE x in ?M. P x" by (intro AE_I) +qed + +lemma AE_embed_measure: + assumes inj: "inj f" + shows "(AE x in embed_measure M f. P x) \<longleftrightarrow> (AE x in M. P (f x))" + using assms by (intro AE_embed_measure') (auto intro!: inj_onI dest: injD) + +lemma nn_integral_monotone_convergence_SUP_countable: + fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> ennreal" + assumes nonempty: "Y \<noteq> {}" + and chain: "Complete_Partial_Order.chain op \<le> (f ` Y)" + and countable: "countable B" + shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space B) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space B))" + (is "?lhs = ?rhs") +proof - + let ?f = "(\<lambda>i x. f i (from_nat_into B x) * indicator (to_nat_on B ` B) x)" + have "?lhs = \<integral>\<^sup>+ x. (SUP i:Y. f i (from_nat_into B (to_nat_on B x))) \<partial>count_space B" + by(rule nn_integral_cong)(simp add: countable) + also have "\<dots> = \<integral>\<^sup>+ x. (SUP i:Y. f i (from_nat_into B x)) \<partial>count_space (to_nat_on B ` B)" + by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1) + also have "\<dots> = \<integral>\<^sup>+ x. (SUP i:Y. ?f i x) \<partial>count_space UNIV" + by(simp add: nn_integral_count_space_indicator ennreal_indicator[symmetric] SUP_mult_right_ennreal nonempty) + also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV)" + proof(rule nn_integral_monotone_convergence_SUP_nat) + show "Complete_Partial_Order.chain op \<le> (?f ` Y)" + by(rule chain_imageI[OF chain, unfolded image_image])(auto intro!: le_funI split: split_indicator dest: le_funD) + qed fact + also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B x) \<partial>count_space (to_nat_on B ` B))" + by(simp add: nn_integral_count_space_indicator) + also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B (to_nat_on B x)) \<partial>count_space B)" + by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1) + also have "\<dots> = ?rhs" + by(intro arg_cong2[where f="SUPREMUM"] ext nn_integral_cong_AE)(simp_all add: AE_count_space countable) + finally show ?thesis . +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Euclidean_Space.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,224 @@ +(* Title: HOL/Analysis/Euclidean_Space.thy + Author: Johannes Hölzl, TU München + Author: Brian Huffman, Portland State University +*) + +section \<open>Finite-Dimensional Inner Product Spaces\<close> + +theory Euclidean_Space +imports + L2_Norm + "~~/src/HOL/Library/Inner_Product" + "~~/src/HOL/Library/Product_Vector" +begin + +subsection \<open>Type class of Euclidean spaces\<close> + +class euclidean_space = real_inner + + fixes Basis :: "'a set" + assumes nonempty_Basis [simp]: "Basis \<noteq> {}" + assumes finite_Basis [simp]: "finite Basis" + assumes inner_Basis: + "\<lbrakk>u \<in> Basis; v \<in> Basis\<rbrakk> \<Longrightarrow> inner u v = (if u = v then 1 else 0)" + assumes euclidean_all_zero_iff: + "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> (x = 0)" + +syntax "_type_dimension" :: "type \<Rightarrow> nat" ("(1DIM/(1'(_')))") +translations "DIM('a)" \<rightharpoonup> "CONST card (CONST Basis :: 'a set)" +typed_print_translation \<open> + [(@{const_syntax card}, + fn ctxt => fn _ => fn [Const (@{const_syntax Basis}, Type (@{type_name set}, [T]))] => + Syntax.const @{syntax_const "_type_dimension"} $ Syntax_Phases.term_of_typ ctxt T)] +\<close> + +lemma (in euclidean_space) norm_Basis[simp]: "u \<in> Basis \<Longrightarrow> norm u = 1" + unfolding norm_eq_sqrt_inner by (simp add: inner_Basis) + +lemma (in euclidean_space) inner_same_Basis[simp]: "u \<in> Basis \<Longrightarrow> inner u u = 1" + by (simp add: inner_Basis) + +lemma (in euclidean_space) inner_not_same_Basis: "u \<in> Basis \<Longrightarrow> v \<in> Basis \<Longrightarrow> u \<noteq> v \<Longrightarrow> inner u v = 0" + by (simp add: inner_Basis) + +lemma (in euclidean_space) sgn_Basis: "u \<in> Basis \<Longrightarrow> sgn u = u" + unfolding sgn_div_norm by (simp add: scaleR_one) + +lemma (in euclidean_space) Basis_zero [simp]: "0 \<notin> Basis" +proof + assume "0 \<in> Basis" thus "False" + using inner_Basis [of 0 0] by simp +qed + +lemma (in euclidean_space) nonzero_Basis: "u \<in> Basis \<Longrightarrow> u \<noteq> 0" + by clarsimp + +lemma (in euclidean_space) SOME_Basis: "(SOME i. i \<in> Basis) \<in> Basis" + by (metis ex_in_conv nonempty_Basis someI_ex) + +lemma (in euclidean_space) inner_setsum_left_Basis[simp]: + "b \<in> Basis \<Longrightarrow> inner (\<Sum>i\<in>Basis. f i *\<^sub>R i) b = f b" + by (simp add: inner_setsum_left inner_Basis if_distrib comm_monoid_add_class.setsum.If_cases) + +lemma (in euclidean_space) euclidean_eqI: + assumes b: "\<And>b. b \<in> Basis \<Longrightarrow> inner x b = inner y b" shows "x = y" +proof - + from b have "\<forall>b\<in>Basis. inner (x - y) b = 0" + by (simp add: inner_diff_left) + then show "x = y" + by (simp add: euclidean_all_zero_iff) +qed + +lemma (in euclidean_space) euclidean_eq_iff: + "x = y \<longleftrightarrow> (\<forall>b\<in>Basis. inner x b = inner y b)" + by (auto intro: euclidean_eqI) + +lemma (in euclidean_space) euclidean_representation_setsum: + "(\<Sum>i\<in>Basis. f i *\<^sub>R i) = b \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)" + by (subst euclidean_eq_iff) simp + +lemma (in euclidean_space) euclidean_representation_setsum': + "b = (\<Sum>i\<in>Basis. f i *\<^sub>R i) \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)" + by (auto simp add: euclidean_representation_setsum[symmetric]) + +lemma (in euclidean_space) euclidean_representation: "(\<Sum>b\<in>Basis. inner x b *\<^sub>R b) = x" + unfolding euclidean_representation_setsum by simp + +lemma (in euclidean_space) choice_Basis_iff: + fixes P :: "'a \<Rightarrow> real \<Rightarrow> bool" + shows "(\<forall>i\<in>Basis. \<exists>x. P i x) \<longleftrightarrow> (\<exists>x. \<forall>i\<in>Basis. P i (inner x i))" + unfolding bchoice_iff +proof safe + fix f assume "\<forall>i\<in>Basis. P i (f i)" + then show "\<exists>x. \<forall>i\<in>Basis. P i (inner x i)" + by (auto intro!: exI[of _ "\<Sum>i\<in>Basis. f i *\<^sub>R i"]) +qed auto + +lemma (in euclidean_space) euclidean_representation_setsum_fun: + "(\<lambda>x. \<Sum>b\<in>Basis. inner (f x) b *\<^sub>R b) = f" + by (rule ext) (simp add: euclidean_representation_setsum) + +lemma euclidean_isCont: + assumes "\<And>b. b \<in> Basis \<Longrightarrow> isCont (\<lambda>x. (inner (f x) b) *\<^sub>R b) x" + shows "isCont f x" + apply (subst euclidean_representation_setsum_fun [symmetric]) + apply (rule isCont_setsum) + apply (blast intro: assms) + done + +lemma DIM_positive: "0 < DIM('a::euclidean_space)" + by (simp add: card_gt_0_iff) + +lemma DIM_ge_Suc0 [iff]: "Suc 0 \<le> card Basis" + by (meson DIM_positive Suc_leI) + + +lemma setsum_inner_Basis_scaleR [simp]: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_vector" + assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) *\<^sub>R f i) = f b" + by (simp add: comm_monoid_add_class.setsum.remove [OF finite_Basis assms] + assms inner_not_same_Basis comm_monoid_add_class.setsum.neutral) + +lemma setsum_inner_Basis_eq [simp]: + assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) * f i) = f b" + by (simp add: comm_monoid_add_class.setsum.remove [OF finite_Basis assms] + assms inner_not_same_Basis comm_monoid_add_class.setsum.neutral) + +subsection \<open>Subclass relationships\<close> + +instance euclidean_space \<subseteq> perfect_space +proof + fix x :: 'a show "\<not> open {x}" + proof + assume "open {x}" + then obtain e where "0 < e" and e: "\<forall>y. dist y x < e \<longrightarrow> y = x" + unfolding open_dist by fast + define y where "y = x + scaleR (e/2) (SOME b. b \<in> Basis)" + have [simp]: "(SOME b. b \<in> Basis) \<in> Basis" + by (rule someI_ex) (auto simp: ex_in_conv) + from \<open>0 < e\<close> have "y \<noteq> x" + unfolding y_def by (auto intro!: nonzero_Basis) + from \<open>0 < e\<close> have "dist y x < e" + unfolding y_def by (simp add: dist_norm) + from \<open>y \<noteq> x\<close> and \<open>dist y x < e\<close> show "False" + using e by simp + qed +qed + +subsection \<open>Class instances\<close> + +subsubsection \<open>Type @{typ real}\<close> + +instantiation real :: euclidean_space +begin + +definition + [simp]: "Basis = {1::real}" + +instance + by standard auto + +end + +lemma DIM_real[simp]: "DIM(real) = 1" + by simp + +subsubsection \<open>Type @{typ complex}\<close> + +instantiation complex :: euclidean_space +begin + +definition Basis_complex_def: "Basis = {1, \<i>}" + +instance + by standard (auto simp add: Basis_complex_def intro: complex_eqI split: if_split_asm) + +end + +lemma DIM_complex[simp]: "DIM(complex) = 2" + unfolding Basis_complex_def by simp + +subsubsection \<open>Type @{typ "'a \<times> 'b"}\<close> + +instantiation prod :: (euclidean_space, euclidean_space) euclidean_space +begin + +definition + "Basis = (\<lambda>u. (u, 0)) ` Basis \<union> (\<lambda>v. (0, v)) ` Basis" + +lemma setsum_Basis_prod_eq: + fixes f::"('a*'b)\<Rightarrow>('a*'b)" + shows "setsum f Basis = setsum (\<lambda>i. f (i, 0)) Basis + setsum (\<lambda>i. f (0, i)) Basis" +proof - + have "inj_on (\<lambda>u. (u::'a, 0::'b)) Basis" "inj_on (\<lambda>u. (0::'a, u::'b)) Basis" + by (auto intro!: inj_onI Pair_inject) + thus ?thesis + unfolding Basis_prod_def + by (subst setsum.union_disjoint) (auto simp: Basis_prod_def setsum.reindex) +qed + +instance proof + show "(Basis :: ('a \<times> 'b) set) \<noteq> {}" + unfolding Basis_prod_def by simp +next + show "finite (Basis :: ('a \<times> 'b) set)" + unfolding Basis_prod_def by simp +next + fix u v :: "'a \<times> 'b" + assume "u \<in> Basis" and "v \<in> Basis" + thus "inner u v = (if u = v then 1 else 0)" + unfolding Basis_prod_def inner_prod_def + by (auto simp add: inner_Basis split: if_split_asm) +next + fix x :: "'a \<times> 'b" + show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0" + unfolding Basis_prod_def ball_Un ball_simps + by (simp add: inner_prod_def prod_eq_iff euclidean_all_zero_iff) +qed + +lemma DIM_prod[simp]: "DIM('a \<times> 'b) = DIM('a) + DIM('b)" + unfolding Basis_prod_def + by (subst card_Un_disjoint) (auto intro!: card_image arg_cong2[where f="op +"] inj_onI) + +end + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Extended_Real_Limits.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,544 @@ +(* Title: HOL/Analysis/Extended_Real_Limits.thy + Author: Johannes Hölzl, TU München + Author: Robert Himmelmann, TU München + Author: Armin Heller, TU München + Author: Bogdan Grechuk, University of Edinburgh +*) + +section \<open>Limits on the Extended real number line\<close> + +theory Extended_Real_Limits +imports + Topology_Euclidean_Space + "~~/src/HOL/Library/Extended_Real" + "~~/src/HOL/Library/Extended_Nonnegative_Real" + "~~/src/HOL/Library/Indicator_Function" +begin + +lemma compact_UNIV: + "compact (UNIV :: 'a::{complete_linorder,linorder_topology,second_countable_topology} set)" + using compact_complete_linorder + by (auto simp: seq_compact_eq_compact[symmetric] seq_compact_def) + +lemma compact_eq_closed: + fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set" + shows "compact S \<longleftrightarrow> closed S" + using closed_Int_compact[of S, OF _ compact_UNIV] compact_imp_closed + by auto + +lemma closed_contains_Sup_cl: + fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set" + assumes "closed S" + and "S \<noteq> {}" + shows "Sup S \<in> S" +proof - + from compact_eq_closed[of S] compact_attains_sup[of S] assms + obtain s where S: "s \<in> S" "\<forall>t\<in>S. t \<le> s" + by auto + then have "Sup S = s" + by (auto intro!: Sup_eqI) + with S show ?thesis + by simp +qed + +lemma closed_contains_Inf_cl: + fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set" + assumes "closed S" + and "S \<noteq> {}" + shows "Inf S \<in> S" +proof - + from compact_eq_closed[of S] compact_attains_inf[of S] assms + obtain s where S: "s \<in> S" "\<forall>t\<in>S. s \<le> t" + by auto + then have "Inf S = s" + by (auto intro!: Inf_eqI) + with S show ?thesis + by simp +qed + +instance ereal :: second_countable_topology +proof (standard, intro exI conjI) + let ?B = "(\<Union>r\<in>\<rat>. {{..< r}, {r <..}} :: ereal set set)" + show "countable ?B" + by (auto intro: countable_rat) + show "open = generate_topology ?B" + proof (intro ext iffI) + fix S :: "ereal set" + assume "open S" + then show "generate_topology ?B S" + unfolding open_generated_order + proof induct + case (Basis b) + then obtain e where "b = {..<e} \<or> b = {e<..}" + by auto + moreover have "{..<e} = \<Union>{{..<x}|x. x \<in> \<rat> \<and> x < e}" "{e<..} = \<Union>{{x<..}|x. x \<in> \<rat> \<and> e < x}" + by (auto dest: ereal_dense3 + simp del: ex_simps + simp add: ex_simps[symmetric] conj_commute Rats_def image_iff) + ultimately show ?case + by (auto intro: generate_topology.intros) + qed (auto intro: generate_topology.intros) + next + fix S + assume "generate_topology ?B S" + then show "open S" + by induct auto + qed +qed + +text \<open>This is a copy from \<open>ereal :: second_countable_topology\<close>. Maybe find a common super class of +topological spaces where the rational numbers are densely embedded ?\<close> +instance ennreal :: second_countable_topology +proof (standard, intro exI conjI) + let ?B = "(\<Union>r\<in>\<rat>. {{..< r}, {r <..}} :: ennreal set set)" + show "countable ?B" + by (auto intro: countable_rat) + show "open = generate_topology ?B" + proof (intro ext iffI) + fix S :: "ennreal set" + assume "open S" + then show "generate_topology ?B S" + unfolding open_generated_order + proof induct + case (Basis b) + then obtain e where "b = {..<e} \<or> b = {e<..}" + by auto + moreover have "{..<e} = \<Union>{{..<x}|x. x \<in> \<rat> \<and> x < e}" "{e<..} = \<Union>{{x<..}|x. x \<in> \<rat> \<and> e < x}" + by (auto dest: ennreal_rat_dense + simp del: ex_simps + simp add: ex_simps[symmetric] conj_commute Rats_def image_iff) + ultimately show ?case + by (auto intro: generate_topology.intros) + qed (auto intro: generate_topology.intros) + next + fix S + assume "generate_topology ?B S" + then show "open S" + by induct auto + qed +qed + +lemma ereal_open_closed_aux: + fixes S :: "ereal set" + assumes "open S" + and "closed S" + and S: "(-\<infinity>) \<notin> S" + shows "S = {}" +proof (rule ccontr) + assume "\<not> ?thesis" + then have *: "Inf S \<in> S" + + by (metis assms(2) closed_contains_Inf_cl) + { + assume "Inf S = -\<infinity>" + then have False + using * assms(3) by auto + } + moreover + { + assume "Inf S = \<infinity>" + then have "S = {\<infinity>}" + by (metis Inf_eq_PInfty \<open>S \<noteq> {}\<close>) + then have False + by (metis assms(1) not_open_singleton) + } + moreover + { + assume fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>" + from ereal_open_cont_interval[OF assms(1) * fin] + obtain e where e: "e > 0" "{Inf S - e<..<Inf S + e} \<subseteq> S" . + then obtain b where b: "Inf S - e < b" "b < Inf S" + using fin ereal_between[of "Inf S" e] dense[of "Inf S - e"] + by auto + then have "b: {Inf S - e <..< Inf S + e}" + using e fin ereal_between[of "Inf S" e] + by auto + then have "b \<in> S" + using e by auto + then have False + using b by (metis complete_lattice_class.Inf_lower leD) + } + ultimately show False + by auto +qed + +lemma ereal_open_closed: + fixes S :: "ereal set" + shows "open S \<and> closed S \<longleftrightarrow> S = {} \<or> S = UNIV" +proof - + { + assume lhs: "open S \<and> closed S" + { + assume "-\<infinity> \<notin> S" + then have "S = {}" + using lhs ereal_open_closed_aux by auto + } + moreover + { + assume "-\<infinity> \<in> S" + then have "- S = {}" + using lhs ereal_open_closed_aux[of "-S"] by auto + } + ultimately have "S = {} \<or> S = UNIV" + by auto + } + then show ?thesis + by auto +qed + +lemma ereal_open_atLeast: + fixes x :: ereal + shows "open {x..} \<longleftrightarrow> x = -\<infinity>" +proof + assume "x = -\<infinity>" + then have "{x..} = UNIV" + by auto + then show "open {x..}" + by auto +next + assume "open {x..}" + then have "open {x..} \<and> closed {x..}" + by auto + then have "{x..} = UNIV" + unfolding ereal_open_closed by auto + then show "x = -\<infinity>" + by (simp add: bot_ereal_def atLeast_eq_UNIV_iff) +qed + +lemma mono_closed_real: + fixes S :: "real set" + assumes mono: "\<forall>y z. y \<in> S \<and> y \<le> z \<longrightarrow> z \<in> S" + and "closed S" + shows "S = {} \<or> S = UNIV \<or> (\<exists>a. S = {a..})" +proof - + { + assume "S \<noteq> {}" + { assume ex: "\<exists>B. \<forall>x\<in>S. B \<le> x" + then have *: "\<forall>x\<in>S. Inf S \<le> x" + using cInf_lower[of _ S] ex by (metis bdd_below_def) + then have "Inf S \<in> S" + apply (subst closed_contains_Inf) + using ex \<open>S \<noteq> {}\<close> \<open>closed S\<close> + apply auto + done + then have "\<forall>x. Inf S \<le> x \<longleftrightarrow> x \<in> S" + using mono[rule_format, of "Inf S"] * + by auto + then have "S = {Inf S ..}" + by auto + then have "\<exists>a. S = {a ..}" + by auto + } + moreover + { + assume "\<not> (\<exists>B. \<forall>x\<in>S. B \<le> x)" + then have nex: "\<forall>B. \<exists>x\<in>S. x < B" + by (simp add: not_le) + { + fix y + obtain x where "x\<in>S" and "x < y" + using nex by auto + then have "y \<in> S" + using mono[rule_format, of x y] by auto + } + then have "S = UNIV" + by auto + } + ultimately have "S = UNIV \<or> (\<exists>a. S = {a ..})" + by blast + } + then show ?thesis + by blast +qed + +lemma mono_closed_ereal: + fixes S :: "real set" + assumes mono: "\<forall>y z. y \<in> S \<and> y \<le> z \<longrightarrow> z \<in> S" + and "closed S" + shows "\<exists>a. S = {x. a \<le> ereal x}" +proof - + { + assume "S = {}" + then have ?thesis + apply (rule_tac x=PInfty in exI) + apply auto + done + } + moreover + { + assume "S = UNIV" + then have ?thesis + apply (rule_tac x="-\<infinity>" in exI) + apply auto + done + } + moreover + { + assume "\<exists>a. S = {a ..}" + then obtain a where "S = {a ..}" + by auto + then have ?thesis + apply (rule_tac x="ereal a" in exI) + apply auto + done + } + ultimately show ?thesis + using mono_closed_real[of S] assms by auto +qed + +lemma Liminf_within: + fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice" + shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S \<inter> ball x e - {x}). f y)" + unfolding Liminf_def eventually_at +proof (rule SUP_eq, simp_all add: Ball_def Bex_def, safe) + fix P d + assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y" + then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}" + by (auto simp: zero_less_dist_iff dist_commute) + then show "\<exists>r>0. INFIMUM (Collect P) f \<le> INFIMUM (S \<inter> ball x r - {x}) f" + by (intro exI[of _ d] INF_mono conjI \<open>0 < d\<close>) auto +next + fix d :: real + assume "0 < d" + then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> xa \<noteq> x \<and> dist xa x < d \<longrightarrow> P xa) \<and> + INFIMUM (S \<inter> ball x d - {x}) f \<le> INFIMUM (Collect P) f" + by (intro exI[of _ "\<lambda>y. y \<in> S \<inter> ball x d - {x}"]) + (auto intro!: INF_mono exI[of _ d] simp: dist_commute) +qed + +lemma Limsup_within: + fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice" + shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S \<inter> ball x e - {x}). f y)" + unfolding Limsup_def eventually_at +proof (rule INF_eq, simp_all add: Ball_def Bex_def, safe) + fix P d + assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y" + then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}" + by (auto simp: zero_less_dist_iff dist_commute) + then show "\<exists>r>0. SUPREMUM (S \<inter> ball x r - {x}) f \<le> SUPREMUM (Collect P) f" + by (intro exI[of _ d] SUP_mono conjI \<open>0 < d\<close>) auto +next + fix d :: real + assume "0 < d" + then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> xa \<noteq> x \<and> dist xa x < d \<longrightarrow> P xa) \<and> + SUPREMUM (Collect P) f \<le> SUPREMUM (S \<inter> ball x d - {x}) f" + by (intro exI[of _ "\<lambda>y. y \<in> S \<inter> ball x d - {x}"]) + (auto intro!: SUP_mono exI[of _ d] simp: dist_commute) +qed + +lemma Liminf_at: + fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice" + shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)" + using Liminf_within[of x UNIV f] by simp + +lemma Limsup_at: + fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice" + shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)" + using Limsup_within[of x UNIV f] by simp + +lemma min_Liminf_at: + fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_linorder" + shows "min (f x) (Liminf (at x) f) = (SUP e:{0<..}. INF y:ball x e. f y)" + unfolding inf_min[symmetric] Liminf_at + apply (subst inf_commute) + apply (subst SUP_inf) + apply (intro SUP_cong[OF refl]) + apply (cut_tac A="ball x xa - {x}" and B="{x}" and M=f in INF_union) + apply (drule sym) + apply auto + apply (metis INF_absorb centre_in_ball) + done + +subsection \<open>monoset\<close> + +definition (in order) mono_set: + "mono_set S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)" + +lemma (in order) mono_greaterThan [intro, simp]: "mono_set {B<..}" unfolding mono_set by auto +lemma (in order) mono_atLeast [intro, simp]: "mono_set {B..}" unfolding mono_set by auto +lemma (in order) mono_UNIV [intro, simp]: "mono_set UNIV" unfolding mono_set by auto +lemma (in order) mono_empty [intro, simp]: "mono_set {}" unfolding mono_set by auto + +lemma (in complete_linorder) mono_set_iff: + fixes S :: "'a set" + defines "a \<equiv> Inf S" + shows "mono_set S \<longleftrightarrow> S = {a <..} \<or> S = {a..}" (is "_ = ?c") +proof + assume "mono_set S" + then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S" + by (auto simp: mono_set) + show ?c + proof cases + assume "a \<in> S" + show ?c + using mono[OF _ \<open>a \<in> S\<close>] + by (auto intro: Inf_lower simp: a_def) + next + assume "a \<notin> S" + have "S = {a <..}" + proof safe + fix x assume "x \<in> S" + then have "a \<le> x" + unfolding a_def by (rule Inf_lower) + then show "a < x" + using \<open>x \<in> S\<close> \<open>a \<notin> S\<close> by (cases "a = x") auto + next + fix x assume "a < x" + then obtain y where "y < x" "y \<in> S" + unfolding a_def Inf_less_iff .. + with mono[of y x] show "x \<in> S" + by auto + qed + then show ?c .. + qed +qed auto + +lemma ereal_open_mono_set: + fixes S :: "ereal set" + shows "open S \<and> mono_set S \<longleftrightarrow> S = UNIV \<or> S = {Inf S <..}" + by (metis Inf_UNIV atLeast_eq_UNIV_iff ereal_open_atLeast + ereal_open_closed mono_set_iff open_ereal_greaterThan) + +lemma ereal_closed_mono_set: + fixes S :: "ereal set" + shows "closed S \<and> mono_set S \<longleftrightarrow> S = {} \<or> S = {Inf S ..}" + by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_ereal_atLeast + ereal_open_closed mono_empty mono_set_iff open_ereal_greaterThan) + +lemma ereal_Liminf_Sup_monoset: + fixes f :: "'a \<Rightarrow> ereal" + shows "Liminf net f = + Sup {l. \<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}" + (is "_ = Sup ?A") +proof (safe intro!: Liminf_eqI complete_lattice_class.Sup_upper complete_lattice_class.Sup_least) + fix P + assume P: "eventually P net" + fix S + assume S: "mono_set S" "INFIMUM (Collect P) f \<in> S" + { + fix x + assume "P x" + then have "INFIMUM (Collect P) f \<le> f x" + by (intro complete_lattice_class.INF_lower) simp + with S have "f x \<in> S" + by (simp add: mono_set) + } + with P show "eventually (\<lambda>x. f x \<in> S) net" + by (auto elim: eventually_mono) +next + fix y l + assume S: "\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net" + assume P: "\<forall>P. eventually P net \<longrightarrow> INFIMUM (Collect P) f \<le> y" + show "l \<le> y" + proof (rule dense_le) + fix B + assume "B < l" + then have "eventually (\<lambda>x. f x \<in> {B <..}) net" + by (intro S[rule_format]) auto + then have "INFIMUM {x. B < f x} f \<le> y" + using P by auto + moreover have "B \<le> INFIMUM {x. B < f x} f" + by (intro INF_greatest) auto + ultimately show "B \<le> y" + by simp + qed +qed + +lemma ereal_Limsup_Inf_monoset: + fixes f :: "'a \<Rightarrow> ereal" + shows "Limsup net f = + Inf {l. \<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}" + (is "_ = Inf ?A") +proof (safe intro!: Limsup_eqI complete_lattice_class.Inf_lower complete_lattice_class.Inf_greatest) + fix P + assume P: "eventually P net" + fix S + assume S: "mono_set (uminus`S)" "SUPREMUM (Collect P) f \<in> S" + { + fix x + assume "P x" + then have "f x \<le> SUPREMUM (Collect P) f" + by (intro complete_lattice_class.SUP_upper) simp + with S(1)[unfolded mono_set, rule_format, of "- SUPREMUM (Collect P) f" "- f x"] S(2) + have "f x \<in> S" + by (simp add: inj_image_mem_iff) } + with P show "eventually (\<lambda>x. f x \<in> S) net" + by (auto elim: eventually_mono) +next + fix y l + assume S: "\<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net" + assume P: "\<forall>P. eventually P net \<longrightarrow> y \<le> SUPREMUM (Collect P) f" + show "y \<le> l" + proof (rule dense_ge) + fix B + assume "l < B" + then have "eventually (\<lambda>x. f x \<in> {..< B}) net" + by (intro S[rule_format]) auto + then have "y \<le> SUPREMUM {x. f x < B} f" + using P by auto + moreover have "SUPREMUM {x. f x < B} f \<le> B" + by (intro SUP_least) auto + ultimately show "y \<le> B" + by simp + qed +qed + +lemma liminf_bounded_open: + fixes x :: "nat \<Rightarrow> ereal" + shows "x0 \<le> liminf x \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> x0 \<in> S \<longrightarrow> (\<exists>N. \<forall>n\<ge>N. x n \<in> S))" + (is "_ \<longleftrightarrow> ?P x0") +proof + assume "?P x0" + then show "x0 \<le> liminf x" + unfolding ereal_Liminf_Sup_monoset eventually_sequentially + by (intro complete_lattice_class.Sup_upper) auto +next + assume "x0 \<le> liminf x" + { + fix S :: "ereal set" + assume om: "open S" "mono_set S" "x0 \<in> S" + { + assume "S = UNIV" + then have "\<exists>N. \<forall>n\<ge>N. x n \<in> S" + by auto + } + moreover + { + assume "S \<noteq> UNIV" + then obtain B where B: "S = {B<..}" + using om ereal_open_mono_set by auto + then have "B < x0" + using om by auto + then have "\<exists>N. \<forall>n\<ge>N. x n \<in> S" + unfolding B + using \<open>x0 \<le> liminf x\<close> liminf_bounded_iff + by auto + } + ultimately have "\<exists>N. \<forall>n\<ge>N. x n \<in> S" + by auto + } + then show "?P x0" + by auto +qed + +subsection "Relate extended reals and the indicator function" + +lemma ereal_indicator_le_0: "(indicator S x::ereal) \<le> 0 \<longleftrightarrow> x \<notin> S" + by (auto split: split_indicator simp: one_ereal_def) + +lemma ereal_indicator: "ereal (indicator A x) = indicator A x" + by (auto simp: indicator_def one_ereal_def) + +lemma ereal_mult_indicator: "ereal (x * indicator A y) = ereal x * indicator A y" + by (simp split: split_indicator) + +lemma ereal_indicator_mult: "ereal (indicator A y * x) = indicator A y * ereal x" + by (simp split: split_indicator) + +lemma ereal_indicator_nonneg[simp, intro]: "0 \<le> (indicator A x ::ereal)" + unfolding indicator_def by auto + +lemma indicator_inter_arith_ereal: "indicator A x * indicator B x = (indicator (A \<inter> B) x :: ereal)" + by (simp split: split_indicator) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Fashoda_Theorem.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,1136 @@ +(* Author: John Harrison + Author: Robert Himmelmann, TU Muenchen (translation from HOL light) +*) + +section \<open>Fashoda meet theorem\<close> + +theory Fashoda_Theorem +imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space +begin + +subsection \<open>Bijections between intervals.\<close> + +definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space" + where "interval_bij = + (\<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i))" + +lemma interval_bij_affine: + "interval_bij (a,b) (u,v) = (\<lambda>x. (\<Sum>i\<in>Basis. ((v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (x\<bullet>i)) *\<^sub>R i) + + (\<Sum>i\<in>Basis. (u\<bullet>i - (v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (a\<bullet>i)) *\<^sub>R i))" + by (auto simp: setsum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff + field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum.cong) + +lemma continuous_interval_bij: + fixes a b :: "'a::euclidean_space" + shows "continuous (at x) (interval_bij (a, b) (u, v))" + by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros) + +lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))" + apply(rule continuous_at_imp_continuous_on) + apply (rule, rule continuous_interval_bij) + done + +lemma in_interval_interval_bij: + fixes a b u v x :: "'a::euclidean_space" + assumes "x \<in> cbox a b" + and "cbox u v \<noteq> {}" + shows "interval_bij (a, b) (u, v) x \<in> cbox u v" + apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong) + apply safe +proof - + fix i :: 'a + assume i: "i \<in> Basis" + have "cbox a b \<noteq> {}" + using assms by auto + with i have *: "a\<bullet>i \<le> b\<bullet>i" "u\<bullet>i \<le> v\<bullet>i" + using assms(2) by (auto simp add: box_eq_empty) + have x: "a\<bullet>i\<le>x\<bullet>i" "x\<bullet>i\<le>b\<bullet>i" + using assms(1)[unfolded mem_box] using i by auto + have "0 \<le> (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)" + using * x by auto + then show "u \<bullet> i \<le> u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)" + using * by auto + have "((x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (v \<bullet> i - u \<bullet> i) \<le> 1 * (v \<bullet> i - u \<bullet> i)" + apply (rule mult_right_mono) + unfolding divide_le_eq_1 + using * x + apply auto + done + then show "u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i) \<le> v \<bullet> i" + using * by auto +qed + +lemma interval_bij_bij: + "\<forall>(i::'a::euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow> + interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = x" + by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a]) + +lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i" + shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" + using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis) + + +subsection \<open>Fashoda meet theorem\<close> + +lemma infnorm_2: + fixes x :: "real^2" + shows "infnorm x = max \<bar>x$1\<bar> \<bar>x$2\<bar>" + unfolding infnorm_cart UNIV_2 by (rule cSup_eq) auto + +lemma infnorm_eq_1_2: + fixes x :: "real^2" + shows "infnorm x = 1 \<longleftrightarrow> + \<bar>x$1\<bar> \<le> 1 \<and> \<bar>x$2\<bar> \<le> 1 \<and> (x$1 = -1 \<or> x$1 = 1 \<or> x$2 = -1 \<or> x$2 = 1)" + unfolding infnorm_2 by auto + +lemma infnorm_eq_1_imp: + fixes x :: "real^2" + assumes "infnorm x = 1" + shows "\<bar>x$1\<bar> \<le> 1" and "\<bar>x$2\<bar> \<le> 1" + using assms unfolding infnorm_eq_1_2 by auto + +lemma fashoda_unit: + fixes f g :: "real \<Rightarrow> real^2" + assumes "f ` {-1 .. 1} \<subseteq> cbox (-1) 1" + and "g ` {-1 .. 1} \<subseteq> cbox (-1) 1" + and "continuous_on {-1 .. 1} f" + and "continuous_on {-1 .. 1} g" + and "f (- 1)$1 = - 1" + and "f 1$1 = 1" "g (- 1) $2 = -1" + and "g 1 $2 = 1" + shows "\<exists>s\<in>{-1 .. 1}. \<exists>t\<in>{-1 .. 1}. f s = g t" +proof (rule ccontr) + assume "\<not> ?thesis" + note as = this[unfolded bex_simps,rule_format] + define sqprojection + where [abs_def]: "sqprojection z = (inverse (infnorm z)) *\<^sub>R z" for z :: "real^2" + define negatex :: "real^2 \<Rightarrow> real^2" + where "negatex x = (vector [-(x$1), x$2])" for x + have lem1: "\<forall>z::real^2. infnorm (negatex z) = infnorm z" + unfolding negatex_def infnorm_2 vector_2 by auto + have lem2: "\<forall>z. z \<noteq> 0 \<longrightarrow> infnorm (sqprojection z) = 1" + unfolding sqprojection_def + unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR] + unfolding abs_inverse real_abs_infnorm + apply (subst infnorm_eq_0[symmetric]) + apply auto + done + let ?F = "\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w" + have *: "\<And>i. (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}" + apply (rule set_eqI) + unfolding image_iff Bex_def mem_interval_cart interval_cbox_cart + apply rule + defer + apply (rule_tac x="vec x" in exI) + apply auto + done + { + fix x + assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` (cbox (- 1) (1::real^2))" + then obtain w :: "real^2" where w: + "w \<in> cbox (- 1) 1" + "x = (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w" + unfolding image_iff .. + then have "x \<noteq> 0" + using as[of "w$1" "w$2"] + unfolding mem_interval_cart atLeastAtMost_iff + by auto + } note x0 = this + have 21: "\<And>i::2. i \<noteq> 1 \<Longrightarrow> i = 2" + using UNIV_2 by auto + have 1: "box (- 1) (1::real^2) \<noteq> {}" + unfolding interval_eq_empty_cart by auto + have 2: "continuous_on (cbox (- 1) 1) (negatex \<circ> sqprojection \<circ> ?F)" + apply (intro continuous_intros continuous_on_component) + unfolding * + apply (rule assms)+ + apply (subst sqprojection_def) + apply (intro continuous_intros) + apply (simp add: infnorm_eq_0 x0) + apply (rule linear_continuous_on) + proof - + show "bounded_linear negatex" + apply (rule bounded_linearI') + unfolding vec_eq_iff + proof (rule_tac[!] allI) + fix i :: 2 + fix x y :: "real^2" + fix c :: real + show "negatex (x + y) $ i = + (negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i" + apply - + apply (case_tac[!] "i\<noteq>1") + prefer 3 + apply (drule_tac[1-2] 21) + unfolding negatex_def + apply (auto simp add:vector_2) + done + qed + qed + have 3: "(negatex \<circ> sqprojection \<circ> ?F) ` cbox (-1) 1 \<subseteq> cbox (-1) 1" + unfolding subset_eq + proof (rule, goal_cases) + case (1 x) + then obtain y :: "real^2" where y: + "y \<in> cbox (- 1) 1" + "x = (negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) y" + unfolding image_iff .. + have "?F y \<noteq> 0" + apply (rule x0) + using y(1) + apply auto + done + then have *: "infnorm (sqprojection (?F y)) = 1" + unfolding y o_def + by - (rule lem2[rule_format]) + have "infnorm x = 1" + unfolding *[symmetric] y o_def + by (rule lem1[rule_format]) + then show "x \<in> cbox (-1) 1" + unfolding mem_interval_cart interval_cbox_cart infnorm_2 + apply - + apply rule + proof - + fix i + assume "max \<bar>x $ 1\<bar> \<bar>x $ 2\<bar> = 1" + then show "(- 1) $ i \<le> x $ i \<and> x $ i \<le> 1 $ i" + apply (cases "i = 1") + defer + apply (drule 21) + apply auto + done + qed + qed + obtain x :: "real^2" where x: + "x \<in> cbox (- 1) 1" + "(negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) x = x" + apply (rule brouwer_weak[of "cbox (- 1) (1::real^2)" "negatex \<circ> sqprojection \<circ> ?F"]) + apply (rule compact_cbox convex_box)+ + unfolding interior_cbox + apply (rule 1 2 3)+ + apply blast + done + have "?F x \<noteq> 0" + apply (rule x0) + using x(1) + apply auto + done + then have *: "infnorm (sqprojection (?F x)) = 1" + unfolding o_def + by (rule lem2[rule_format]) + have nx: "infnorm x = 1" + apply (subst x(2)[symmetric]) + unfolding *[symmetric] o_def + apply (rule lem1[rule_format]) + done + have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)" + and "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)" + apply - + apply (rule_tac[!] allI impI)+ + proof - + fix x :: "real^2" + fix i :: 2 + assume x: "x \<noteq> 0" + have "inverse (infnorm x) > 0" + using x[unfolded infnorm_pos_lt[symmetric]] by auto + then show "(0 < sqprojection x $ i) = (0 < x $ i)" + and "(sqprojection x $ i < 0) = (x $ i < 0)" + unfolding sqprojection_def vector_component_simps vector_scaleR_component real_scaleR_def + unfolding zero_less_mult_iff mult_less_0_iff + by (auto simp add: field_simps) + qed + note lem3 = this[rule_format] + have x1: "x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}" + using x(1) unfolding mem_interval_cart by auto + then have nz: "f (x $ 1) - g (x $ 2) \<noteq> 0" + unfolding right_minus_eq + apply - + apply (rule as) + apply auto + done + have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1" + using nx unfolding infnorm_eq_1_2 by auto + then show False + proof - + fix P Q R S + presume "P \<or> Q \<or> R \<or> S" + and "P \<Longrightarrow> False" + and "Q \<Longrightarrow> False" + and "R \<Longrightarrow> False" + and "S \<Longrightarrow> False" + then show False by auto + next + assume as: "x$1 = 1" + then have *: "f (x $ 1) $ 1 = 1" + using assms(6) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "g (x $ 2) \<in> cbox (-1) 1" + apply - + apply (rule assms(2)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=1 in allE) + apply auto + done + next + assume as: "x$1 = -1" + then have *: "f (x $ 1) $ 1 = - 1" + using assms(5) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "g (x $ 2) \<in> cbox (-1) 1" + apply - + apply (rule assms(2)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=1 in allE) + apply auto + done + next + assume as: "x$2 = 1" + then have *: "g (x $ 2) $ 2 = 1" + using assms(8) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "f (x $ 1) \<in> cbox (-1) 1" + apply - + apply (rule assms(1)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=2 in allE) + apply auto + done + next + assume as: "x$2 = -1" + then have *: "g (x $ 2) $ 2 = - 1" + using assms(7) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "f (x $ 1) \<in> cbox (-1) 1" + apply - + apply (rule assms(1)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=2 in allE) + apply auto + done + qed auto +qed + +lemma fashoda_unit_path: + fixes f g :: "real \<Rightarrow> real^2" + assumes "path f" + and "path g" + and "path_image f \<subseteq> cbox (-1) 1" + and "path_image g \<subseteq> cbox (-1) 1" + and "(pathstart f)$1 = -1" + and "(pathfinish f)$1 = 1" + and "(pathstart g)$2 = -1" + and "(pathfinish g)$2 = 1" + obtains z where "z \<in> path_image f" and "z \<in> path_image g" +proof - + note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def] + define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real + have isc: "iscale ` {- 1..1} \<subseteq> {0..1}" + unfolding iscale_def by auto + have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t" + proof (rule fashoda_unit) + show "(f \<circ> iscale) ` {- 1..1} \<subseteq> cbox (- 1) 1" "(g \<circ> iscale) ` {- 1..1} \<subseteq> cbox (- 1) 1" + using isc and assms(3-4) by (auto simp add: image_comp [symmetric]) + have *: "continuous_on {- 1..1} iscale" + unfolding iscale_def by (rule continuous_intros)+ + show "continuous_on {- 1..1} (f \<circ> iscale)" "continuous_on {- 1..1} (g \<circ> iscale)" + apply - + apply (rule_tac[!] continuous_on_compose[OF *]) + apply (rule_tac[!] continuous_on_subset[OF _ isc]) + apply (rule assms)+ + done + have *: "(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" + unfolding vec_eq_iff by auto + show "(f \<circ> iscale) (- 1) $ 1 = - 1" + and "(f \<circ> iscale) 1 $ 1 = 1" + and "(g \<circ> iscale) (- 1) $ 2 = -1" + and "(g \<circ> iscale) 1 $ 2 = 1" + unfolding o_def iscale_def + using assms + by (auto simp add: *) + qed + then obtain s t where st: + "s \<in> {- 1..1}" + "t \<in> {- 1..1}" + "(f \<circ> iscale) s = (g \<circ> iscale) t" + by auto + show thesis + apply (rule_tac z = "f (iscale s)" in that) + using st + unfolding o_def path_image_def image_iff + apply - + apply (rule_tac x="iscale s" in bexI) + prefer 3 + apply (rule_tac x="iscale t" in bexI) + using isc[unfolded subset_eq, rule_format] + apply auto + done +qed + +lemma fashoda: + fixes b :: "real^2" + assumes "path f" + and "path g" + and "path_image f \<subseteq> cbox a b" + and "path_image g \<subseteq> cbox a b" + and "(pathstart f)$1 = a$1" + and "(pathfinish f)$1 = b$1" + and "(pathstart g)$2 = a$2" + and "(pathfinish g)$2 = b$2" + obtains z where "z \<in> path_image f" and "z \<in> path_image g" +proof - + fix P Q S + presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" and "Q \<Longrightarrow> thesis" and "S \<Longrightarrow> thesis" + then show thesis + by auto +next + have "cbox a b \<noteq> {}" + using assms(3) using path_image_nonempty[of f] by auto + then have "a \<le> b" + unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less) + then show "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)" + unfolding less_eq_vec_def forall_2 by auto +next + assume as: "a$1 = b$1" + have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2" + apply (rule connected_ivt_component_cart) + apply (rule connected_path_image assms)+ + apply (rule pathstart_in_path_image) + apply (rule pathfinish_in_path_image) + unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"] + unfolding pathstart_def + apply (auto simp add: less_eq_vec_def mem_interval_cart) + done + then obtain z :: "real^2" where z: "z \<in> path_image g" "z $ 2 = pathstart f $ 2" .. + have "z \<in> cbox a b" + using z(1) assms(4) + unfolding path_image_def + by blast + then have "z = f 0" + unfolding vec_eq_iff forall_2 + unfolding z(2) pathstart_def + using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1] + unfolding mem_interval_cart + apply (erule_tac x=1 in allE) + using as + apply auto + done + then show thesis + apply - + apply (rule that[OF _ z(1)]) + unfolding path_image_def + apply auto + done +next + assume as: "a$2 = b$2" + have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1" + apply (rule connected_ivt_component_cart) + apply (rule connected_path_image assms)+ + apply (rule pathstart_in_path_image) + apply (rule pathfinish_in_path_image) + unfolding assms + using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"] + unfolding pathstart_def + apply (auto simp add: less_eq_vec_def mem_interval_cart) + done + then obtain z where z: "z \<in> path_image f" "z $ 1 = pathstart g $ 1" .. + have "z \<in> cbox a b" + using z(1) assms(3) + unfolding path_image_def + by blast + then have "z = g 0" + unfolding vec_eq_iff forall_2 + unfolding z(2) pathstart_def + using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2] + unfolding mem_interval_cart + apply (erule_tac x=2 in allE) + using as + apply auto + done + then show thesis + apply - + apply (rule that[OF z(1)]) + unfolding path_image_def + apply auto + done +next + assume as: "a $ 1 < b $ 1 \<and> a $ 2 < b $ 2" + have int_nem: "cbox (-1) (1::real^2) \<noteq> {}" + unfolding interval_eq_empty_cart by auto + obtain z :: "real^2" where z: + "z \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}" + "z \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}" + apply (rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \<circ> f" "interval_bij (a,b) (- 1,1) \<circ> g"]) + unfolding path_def path_image_def pathstart_def pathfinish_def + apply (rule_tac[1-2] continuous_on_compose) + apply (rule assms[unfolded path_def] continuous_on_interval_bij)+ + unfolding subset_eq + apply(rule_tac[1-2] ballI) + proof - + fix x + assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}" + then obtain y where y: + "y \<in> {0..1}" + "x = (interval_bij (a, b) (- 1, 1) \<circ> f) y" + unfolding image_iff .. + show "x \<in> cbox (- 1) 1" + unfolding y o_def + apply (rule in_interval_interval_bij) + using y(1) + using assms(3)[unfolded path_image_def subset_eq] int_nem + apply auto + done + next + fix x + assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}" + then obtain y where y: + "y \<in> {0..1}" + "x = (interval_bij (a, b) (- 1, 1) \<circ> g) y" + unfolding image_iff .. + show "x \<in> cbox (- 1) 1" + unfolding y o_def + apply (rule in_interval_interval_bij) + using y(1) + using assms(4)[unfolded path_image_def subset_eq] int_nem + apply auto + done + next + show "(interval_bij (a, b) (- 1, 1) \<circ> f) 0 $ 1 = -1" + and "(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1" + and "(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1" + and "(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1" + using assms as + by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def) + (simp_all add: inner_axis) + qed + from z(1) obtain zf where zf: + "zf \<in> {0..1}" + "z = (interval_bij (a, b) (- 1, 1) \<circ> f) zf" + unfolding image_iff .. + from z(2) obtain zg where zg: + "zg \<in> {0..1}" + "z = (interval_bij (a, b) (- 1, 1) \<circ> g) zg" + unfolding image_iff .. + have *: "\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i" + unfolding forall_2 + using as + by auto + show thesis + apply (rule_tac z="interval_bij (- 1,1) (a,b) z" in that) + apply (subst zf) + defer + apply (subst zg) + unfolding o_def interval_bij_bij_cart[OF *] path_image_def + using zf(1) zg(1) + apply auto + done +qed + + +subsection \<open>Some slightly ad hoc lemmas I use below\<close> + +lemma segment_vertical: + fixes a :: "real^2" + assumes "a$1 = b$1" + shows "x \<in> closed_segment a b \<longleftrightarrow> + x$1 = a$1 \<and> x$1 = b$1 \<and> (a$2 \<le> x$2 \<and> x$2 \<le> b$2 \<or> b$2 \<le> x$2 \<and> x$2 \<le> a$2)" + (is "_ = ?R") +proof - + let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1" + { + presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L" + then show ?thesis + unfolding closed_segment_def mem_Collect_eq + unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps + by blast + } + { + assume ?L + then obtain u where u: + "x $ 1 = (1 - u) * a $ 1 + u * b $ 1" + "x $ 2 = (1 - u) * a $ 2 + u * b $ 2" + "0 \<le> u" + "u \<le> 1" + by blast + { fix b a + assume "b + u * a > a + u * b" + then have "(1 - u) * b > (1 - u) * a" + by (auto simp add:field_simps) + then have "b \<ge> a" + apply (drule_tac mult_left_less_imp_less) + using u + apply auto + done + then have "u * a \<le> u * b" + apply - + apply (rule mult_left_mono[OF _ u(3)]) + using u(3-4) + apply (auto simp add: field_simps) + done + } note * = this + { + fix a b + assume "u * b > u * a" + then have "(1 - u) * a \<le> (1 - u) * b" + apply - + apply (rule mult_left_mono) + apply (drule mult_left_less_imp_less) + using u + apply auto + done + then have "a + u * b \<le> b + u * a" + by (auto simp add: field_simps) + } note ** = this + then show ?R + unfolding u assms + using u + by (auto simp add:field_simps not_le intro: * **) + } + { + assume ?R + then show ?L + proof (cases "x$2 = b$2") + case True + then show ?L + apply (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) + unfolding assms True + using \<open>?R\<close> + apply (auto simp add: field_simps) + done + next + case False + then show ?L + apply (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) + unfolding assms + using \<open>?R\<close> + apply (auto simp add: field_simps) + done + qed + } +qed + +lemma segment_horizontal: + fixes a :: "real^2" + assumes "a$2 = b$2" + shows "x \<in> closed_segment a b \<longleftrightarrow> + x$2 = a$2 \<and> x$2 = b$2 \<and> (a$1 \<le> x$1 \<and> x$1 \<le> b$1 \<or> b$1 \<le> x$1 \<and> x$1 \<le> a$1)" + (is "_ = ?R") +proof - + let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1" + { + presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L" + then show ?thesis + unfolding closed_segment_def mem_Collect_eq + unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps + by blast + } + { + assume ?L + then obtain u where u: + "x $ 1 = (1 - u) * a $ 1 + u * b $ 1" + "x $ 2 = (1 - u) * a $ 2 + u * b $ 2" + "0 \<le> u" + "u \<le> 1" + by blast + { + fix b a + assume "b + u * a > a + u * b" + then have "(1 - u) * b > (1 - u) * a" + by (auto simp add: field_simps) + then have "b \<ge> a" + apply (drule_tac mult_left_less_imp_less) + using u + apply auto + done + then have "u * a \<le> u * b" + apply - + apply (rule mult_left_mono[OF _ u(3)]) + using u(3-4) + apply (auto simp add: field_simps) + done + } note * = this + { + fix a b + assume "u * b > u * a" + then have "(1 - u) * a \<le> (1 - u) * b" + apply - + apply (rule mult_left_mono) + apply (drule mult_left_less_imp_less) + using u + apply auto + done + then have "a + u * b \<le> b + u * a" + by (auto simp add: field_simps) + } note ** = this + then show ?R + unfolding u assms + using u + by (auto simp add: field_simps not_le intro: * **) + } + { + assume ?R + then show ?L + proof (cases "x$1 = b$1") + case True + then show ?L + apply (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) + unfolding assms True + using \<open>?R\<close> + apply (auto simp add: field_simps) + done + next + case False + then show ?L + apply (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) + unfolding assms + using \<open>?R\<close> + apply (auto simp add: field_simps) + done + qed + } +qed + + +subsection \<open>Useful Fashoda corollary pointed out to me by Tom Hales\<close> + +lemma fashoda_interlace: + fixes a :: "real^2" + assumes "path f" + and "path g" + and "path_image f \<subseteq> cbox a b" + and "path_image g \<subseteq> cbox a b" + and "(pathstart f)$2 = a$2" + and "(pathfinish f)$2 = a$2" + and "(pathstart g)$2 = a$2" + and "(pathfinish g)$2 = a$2" + and "(pathstart f)$1 < (pathstart g)$1" + and "(pathstart g)$1 < (pathfinish f)$1" + and "(pathfinish f)$1 < (pathfinish g)$1" + obtains z where "z \<in> path_image f" and "z \<in> path_image g" +proof - + have "cbox a b \<noteq> {}" + using path_image_nonempty[of f] using assms(3) by auto + note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less] + have "pathstart f \<in> cbox a b" + and "pathfinish f \<in> cbox a b" + and "pathstart g \<in> cbox a b" + and "pathfinish g \<in> cbox a b" + using pathstart_in_path_image pathfinish_in_path_image + using assms(3-4) + by auto + note startfin = this[unfolded mem_interval_cart forall_2] + let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++ + linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++ + linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++ + linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" + let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++ + linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++ + linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++ + linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])" + let ?a = "vector[a$1 - 2, a$2 - 3]" + let ?b = "vector[b$1 + 2, b$2 + 3]" + have P1P2: "path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union> + path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union> + path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union> + path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))" + "path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union> + path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union> + path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union> + path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) + by(auto simp add: path_image_join path_linepath) + have abab: "cbox a b \<subseteq> cbox ?a ?b" + unfolding interval_cbox_cart[symmetric] + by (auto simp add:less_eq_vec_def forall_2 vector_2) + obtain z where + "z \<in> path_image + (linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++ + linepath (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f) +++ + f +++ + linepath (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) +++ + linepath (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]))" + "z \<in> path_image + (linepath (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g) +++ + g +++ + linepath (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1]) +++ + linepath (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1]) +++ + linepath (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]))" + apply (rule fashoda[of ?P1 ?P2 ?a ?b]) + unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 + proof - + show "path ?P1" and "path ?P2" + using assms by auto + have "path_image ?P1 \<subseteq> cbox ?a ?b" + unfolding P1P2 path_image_linepath + apply (rule Un_least)+ + defer 3 + apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) + unfolding mem_interval_cart forall_2 vector_2 + using ab startfin abab assms(3) + using assms(9-) + unfolding assms + apply (auto simp add: field_simps box_def) + done + then show "path_image ?P1 \<subseteq> cbox ?a ?b" . + have "path_image ?P2 \<subseteq> cbox ?a ?b" + unfolding P1P2 path_image_linepath + apply (rule Un_least)+ + defer 2 + apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) + unfolding mem_interval_cart forall_2 vector_2 + using ab startfin abab assms(4) + using assms(9-) + unfolding assms + apply (auto simp add: field_simps box_def) + done + then show "path_image ?P2 \<subseteq> cbox ?a ?b" . + show "a $ 1 - 2 = a $ 1 - 2" + and "b $ 1 + 2 = b $ 1 + 2" + and "pathstart g $ 2 - 3 = a $ 2 - 3" + and "b $ 2 + 3 = b $ 2 + 3" + by (auto simp add: assms) + qed + note z=this[unfolded P1P2 path_image_linepath] + show thesis + apply (rule that[of z]) + proof - + have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or> + z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or> + z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or> + z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow> + (((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or> + z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or> + z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or> + z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False" + proof (simp only: segment_vertical segment_horizontal vector_2, goal_cases) + case prems: 1 + have "pathfinish f \<in> cbox a b" + using assms(3) pathfinish_in_path_image[of f] by auto + then have "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False" + unfolding mem_interval_cart forall_2 by auto + then have "z$1 \<noteq> pathfinish f$1" + using prems(2) + using assms ab + by (auto simp add: field_simps) + moreover have "pathstart f \<in> cbox a b" + using assms(3) pathstart_in_path_image[of f] + by auto + then have "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False" + unfolding mem_interval_cart forall_2 + by auto + then have "z$1 \<noteq> pathstart f$1" + using prems(2) using assms ab + by (auto simp add: field_simps) + ultimately have *: "z$2 = a$2 - 2" + using prems(1) + by auto + have "z$1 \<noteq> pathfinish g$1" + using prems(2) + using assms ab + by (auto simp add: field_simps *) + moreover have "pathstart g \<in> cbox a b" + using assms(4) pathstart_in_path_image[of g] + by auto + note this[unfolded mem_interval_cart forall_2] + then have "z$1 \<noteq> pathstart g$1" + using prems(1) + using assms ab + by (auto simp add: field_simps *) + ultimately have "a $ 2 - 1 \<le> z $ 2 \<and> z $ 2 \<le> b $ 2 + 3 \<or> b $ 2 + 3 \<le> z $ 2 \<and> z $ 2 \<le> a $ 2 - 1" + using prems(2) + unfolding * assms + by (auto simp add: field_simps) + then show False + unfolding * using ab by auto + qed + then have "z \<in> path_image f \<or> z \<in> path_image g" + using z unfolding Un_iff by blast + then have z': "z \<in> cbox a b" + using assms(3-4) + by auto + have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow> + z = pathstart f \<or> z = pathfinish f" + unfolding vec_eq_iff forall_2 assms + by auto + with z' show "z \<in> path_image f" + using z(1) + unfolding Un_iff mem_interval_cart forall_2 + apply - + apply (simp only: segment_vertical segment_horizontal vector_2) + unfolding assms + apply auto + done + have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow> + z = pathstart g \<or> z = pathfinish g" + unfolding vec_eq_iff forall_2 assms + by auto + with z' show "z \<in> path_image g" + using z(2) + unfolding Un_iff mem_interval_cart forall_2 + apply (simp only: segment_vertical segment_horizontal vector_2) + unfolding assms + apply auto + done + qed +qed + +(** The Following still needs to be translated. Maybe I will do that later. + +(* ------------------------------------------------------------------------- *) +(* Complement in dimension N >= 2 of set homeomorphic to any interval in *) +(* any dimension is (path-)connected. This naively generalizes the argument *) +(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer *) +(* fixed point theorem", American Mathematical Monthly 1984. *) +(* ------------------------------------------------------------------------- *) + +let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove + (`!p:real^M->real^N a b. + ~(interval[a,b] = {}) /\ + p continuous_on interval[a,b] /\ + (!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y) + ==> ?f. f continuous_on (:real^N) /\ + IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\ + (!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`, + REPEAT STRIP_TAC THEN + FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN + DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN + SUBGOAL_THEN `(q:real^N->real^M) continuous_on + (IMAGE p (interval[a:real^M,b]))` + ASSUME_TAC THENL + [MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]; + ALL_TAC] THEN + MP_TAC(ISPECL [`q:real^N->real^M`; + `IMAGE (p:real^M->real^N) + (interval[a,b])`; + `a:real^M`; `b:real^M`] + TIETZE_CLOSED_INTERVAL) THEN + ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE; + COMPACT_IMP_CLOSED] THEN + ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN + DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN + EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN + REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN + CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN + MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN + FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ] + CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);; + +let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove + (`!s:real^N->bool a b:real^M. + s homeomorphic (interval[a,b]) + ==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`, + REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN + REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN + MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN + DISCH_TAC THEN + SUBGOAL_THEN + `!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ + (p:real^M->real^N) x = p y ==> x = y` + ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN + FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN + DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN + ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN + ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV; + NOT_BOUNDED_UNIV] THEN + ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN + X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN + SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN + SUBGOAL_THEN `bounded((path_component s c) UNION + (IMAGE (p:real^M->real^N) (interval[a,b])))` + MP_TAC THENL + [ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED; + COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; + ALL_TAC] THEN + DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN + REWRITE_TAC[UNION_SUBSET] THEN + DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN + MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`] + RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN + ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN + DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN + DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC + (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN + REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN + ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN + SUBGOAL_THEN + `(q:real^N->real^N) continuous_on + (closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))` + MP_TAC THENL + [EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN + REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN + REPEAT CONJ_TAC THENL + [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN + ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; + COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; + ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; + ALL_TAC] THEN + X_GEN_TAC `z:real^N` THEN + REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN + STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN + MP_TAC(ISPECL + [`path_component s (z:real^N)`; `path_component s (c:real^N)`] + OPEN_INTER_CLOSURE_EQ_EMPTY) THEN + ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL + [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN + ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; + COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; + REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN + DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN + GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN + REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]]; + ALL_TAC] THEN + SUBGOAL_THEN + `closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) = + (:real^N)` + SUBST1_TAC THENL + [MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN + REWRITE_TAC[CLOSURE_SUBSET]; + DISCH_TAC] THEN + MP_TAC(ISPECL + [`(\x. &2 % c - x) o + (\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`; + `cball(c:real^N,B)`] + BROUWER) THEN + REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN + ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN + SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL + [X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN + REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN + ASM SET_TAC[PATH_COMPONENT_REFL_EQ]; + ALL_TAC] THEN + REPEAT CONJ_TAC THENL + [MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN + SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN + MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL + [ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN + MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN + MATCH_MP_TAC CONTINUOUS_ON_MUL THEN + SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN + REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN + MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN + MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN + ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN + SUBGOAL_THEN + `(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)` + SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN + MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN + ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST; + CONTINUOUS_ON_LIFT_NORM]; + REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN + X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN + REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN + REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN + ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN + ASM_REAL_ARITH_TAC; + REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN + REWRITE_TAC[IN_CBALL; o_THM; dist] THEN + X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN + REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN + ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL + [MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN + REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN + ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN + ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN + UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN + REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB]; + EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN + REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN + ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN + SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL + [ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN + ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);; + +let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove + (`!s:real^N->bool a b:real^M. + 2 <= dimindex(:N) /\ s homeomorphic interval[a,b] + ==> path_connected((:real^N) DIFF s)`, + REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN + FIRST_ASSUM(MP_TAC o MATCH_MP + UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN + ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN + ABBREV_TAC `t = (:real^N) DIFF s` THEN + DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN + STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN + REWRITE_TAC[COMPACT_INTERVAL] THEN + DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN + REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN + X_GEN_TAC `B:real` THEN STRIP_TAC THEN + SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\ + (?v:real^N. v IN path_component t y /\ B < norm(v))` + STRIP_ASSUME_TAC THENL + [ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN + MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN + CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN + MATCH_MP_TAC PATH_COMPONENT_SYM THEN + MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN + CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN + MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN + EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL + [EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE + `s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN + ASM_REWRITE_TAC[SUBSET; IN_CBALL_0]; + MP_TAC(ISPEC `cball(vec 0:real^N,B)` + PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN + ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN + REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN + DISCH_THEN MATCH_MP_TAC THEN + ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);; + +(* ------------------------------------------------------------------------- *) +(* In particular, apply all these to the special case of an arc. *) +(* ------------------------------------------------------------------------- *) + +let RETRACTION_ARC = prove + (`!p. arc p + ==> ?f. f continuous_on (:real^N) /\ + IMAGE f (:real^N) SUBSET path_image p /\ + (!x. x IN path_image p ==> f x = x)`, + REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN + MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN + ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);; + +let PATH_CONNECTED_ARC_COMPLEMENT = prove + (`!p. 2 <= dimindex(:N) /\ arc p + ==> path_connected((:real^N) DIFF path_image p)`, + REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN + MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`] + PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN + ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN + ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN + MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN + EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);; + +let CONNECTED_ARC_COMPLEMENT = prove + (`!p. 2 <= dimindex(:N) /\ arc p + ==> connected((:real^N) DIFF path_image p)`, + SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,558 @@ +(* Title: HOL/Analysis/Finite_Cartesian_Product.thy + Author: Amine Chaieb, University of Cambridge +*) + +section \<open>Definition of finite Cartesian product types.\<close> + +theory Finite_Cartesian_Product +imports + Euclidean_Space + L2_Norm + "~~/src/HOL/Library/Numeral_Type" +begin + +subsection \<open>Finite Cartesian products, with indexing and lambdas.\<close> + +typedef ('a, 'b) vec = "UNIV :: (('b::finite) \<Rightarrow> 'a) set" + morphisms vec_nth vec_lambda .. + +notation + vec_nth (infixl "$" 90) and + vec_lambda (binder "\<chi>" 10) + +(* + Translate "'b ^ 'n" into "'b ^ ('n :: finite)". When 'n has already more than + the finite type class write "vec 'b 'n" +*) + +syntax "_finite_vec" :: "type \<Rightarrow> type \<Rightarrow> type" ("(_ ^/ _)" [15, 16] 15) + +parse_translation \<open> + let + fun vec t u = Syntax.const @{type_syntax vec} $ t $ u; + fun finite_vec_tr [t, u] = + (case Term_Position.strip_positions u of + v as Free (x, _) => + if Lexicon.is_tid x then + vec t (Syntax.const @{syntax_const "_ofsort"} $ v $ + Syntax.const @{class_syntax finite}) + else vec t u + | _ => vec t u) + in + [(@{syntax_const "_finite_vec"}, K finite_vec_tr)] + end +\<close> + +lemma vec_eq_iff: "(x = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)" + by (simp add: vec_nth_inject [symmetric] fun_eq_iff) + +lemma vec_lambda_beta [simp]: "vec_lambda g $ i = g i" + by (simp add: vec_lambda_inverse) + +lemma vec_lambda_unique: "(\<forall>i. f$i = g i) \<longleftrightarrow> vec_lambda g = f" + by (auto simp add: vec_eq_iff) + +lemma vec_lambda_eta: "(\<chi> i. (g$i)) = g" + by (simp add: vec_eq_iff) + + +subsection \<open>Group operations and class instances\<close> + +instantiation vec :: (zero, finite) zero +begin + definition "0 \<equiv> (\<chi> i. 0)" + instance .. +end + +instantiation vec :: (plus, finite) plus +begin + definition "op + \<equiv> (\<lambda> x y. (\<chi> i. x$i + y$i))" + instance .. +end + +instantiation vec :: (minus, finite) minus +begin + definition "op - \<equiv> (\<lambda> x y. (\<chi> i. x$i - y$i))" + instance .. +end + +instantiation vec :: (uminus, finite) uminus +begin + definition "uminus \<equiv> (\<lambda> x. (\<chi> i. - (x$i)))" + instance .. +end + +lemma zero_index [simp]: "0 $ i = 0" + unfolding zero_vec_def by simp + +lemma vector_add_component [simp]: "(x + y)$i = x$i + y$i" + unfolding plus_vec_def by simp + +lemma vector_minus_component [simp]: "(x - y)$i = x$i - y$i" + unfolding minus_vec_def by simp + +lemma vector_uminus_component [simp]: "(- x)$i = - (x$i)" + unfolding uminus_vec_def by simp + +instance vec :: (semigroup_add, finite) semigroup_add + by standard (simp add: vec_eq_iff add.assoc) + +instance vec :: (ab_semigroup_add, finite) ab_semigroup_add + by standard (simp add: vec_eq_iff add.commute) + +instance vec :: (monoid_add, finite) monoid_add + by standard (simp_all add: vec_eq_iff) + +instance vec :: (comm_monoid_add, finite) comm_monoid_add + by standard (simp add: vec_eq_iff) + +instance vec :: (cancel_semigroup_add, finite) cancel_semigroup_add + by standard (simp_all add: vec_eq_iff) + +instance vec :: (cancel_ab_semigroup_add, finite) cancel_ab_semigroup_add + by standard (simp_all add: vec_eq_iff diff_diff_eq) + +instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add .. + +instance vec :: (group_add, finite) group_add + by standard (simp_all add: vec_eq_iff) + +instance vec :: (ab_group_add, finite) ab_group_add + by standard (simp_all add: vec_eq_iff) + + +subsection \<open>Real vector space\<close> + +instantiation vec :: (real_vector, finite) real_vector +begin + +definition "scaleR \<equiv> (\<lambda> r x. (\<chi> i. scaleR r (x$i)))" + +lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)" + unfolding scaleR_vec_def by simp + +instance + by standard (simp_all add: vec_eq_iff scaleR_left_distrib scaleR_right_distrib) + +end + + +subsection \<open>Topological space\<close> + +instantiation vec :: (topological_space, finite) topological_space +begin + +definition [code del]: + "open (S :: ('a ^ 'b) set) \<longleftrightarrow> + (\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and> + (\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))" + +instance proof + show "open (UNIV :: ('a ^ 'b) set)" + unfolding open_vec_def by auto +next + fix S T :: "('a ^ 'b) set" + assume "open S" "open T" thus "open (S \<inter> T)" + unfolding open_vec_def + apply clarify + apply (drule (1) bspec)+ + apply (clarify, rename_tac Sa Ta) + apply (rule_tac x="\<lambda>i. Sa i \<inter> Ta i" in exI) + apply (simp add: open_Int) + done +next + fix K :: "('a ^ 'b) set set" + assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)" + unfolding open_vec_def + apply clarify + apply (drule (1) bspec) + apply (drule (1) bspec) + apply clarify + apply (rule_tac x=A in exI) + apply fast + done +qed + +end + +lemma open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}" + unfolding open_vec_def by auto + +lemma open_vimage_vec_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)" + unfolding open_vec_def + apply clarify + apply (rule_tac x="\<lambda>k. if k = i then S else UNIV" in exI, simp) + done + +lemma closed_vimage_vec_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)" + unfolding closed_open vimage_Compl [symmetric] + by (rule open_vimage_vec_nth) + +lemma closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}" +proof - + have "{x. \<forall>i. x $ i \<in> S i} = (\<Inter>i. (\<lambda>x. x $ i) -` S i)" by auto + thus "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}" + by (simp add: closed_INT closed_vimage_vec_nth) +qed + +lemma tendsto_vec_nth [tendsto_intros]: + assumes "((\<lambda>x. f x) \<longlongrightarrow> a) net" + shows "((\<lambda>x. f x $ i) \<longlongrightarrow> a $ i) net" +proof (rule topological_tendstoI) + fix S assume "open S" "a $ i \<in> S" + then have "open ((\<lambda>y. y $ i) -` S)" "a \<in> ((\<lambda>y. y $ i) -` S)" + by (simp_all add: open_vimage_vec_nth) + with assms have "eventually (\<lambda>x. f x \<in> (\<lambda>y. y $ i) -` S) net" + by (rule topological_tendstoD) + then show "eventually (\<lambda>x. f x $ i \<in> S) net" + by simp +qed + +lemma isCont_vec_nth [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x $ i) a" + unfolding isCont_def by (rule tendsto_vec_nth) + +lemma vec_tendstoI: + assumes "\<And>i. ((\<lambda>x. f x $ i) \<longlongrightarrow> a $ i) net" + shows "((\<lambda>x. f x) \<longlongrightarrow> a) net" +proof (rule topological_tendstoI) + fix S assume "open S" and "a \<in> S" + then obtain A where A: "\<And>i. open (A i)" "\<And>i. a $ i \<in> A i" + and S: "\<And>y. \<forall>i. y $ i \<in> A i \<Longrightarrow> y \<in> S" + unfolding open_vec_def by metis + have "\<And>i. eventually (\<lambda>x. f x $ i \<in> A i) net" + using assms A by (rule topological_tendstoD) + hence "eventually (\<lambda>x. \<forall>i. f x $ i \<in> A i) net" + by (rule eventually_all_finite) + thus "eventually (\<lambda>x. f x \<in> S) net" + by (rule eventually_mono, simp add: S) +qed + +lemma tendsto_vec_lambda [tendsto_intros]: + assumes "\<And>i. ((\<lambda>x. f x i) \<longlongrightarrow> a i) net" + shows "((\<lambda>x. \<chi> i. f x i) \<longlongrightarrow> (\<chi> i. a i)) net" + using assms by (simp add: vec_tendstoI) + +lemma open_image_vec_nth: assumes "open S" shows "open ((\<lambda>x. x $ i) ` S)" +proof (rule openI) + fix a assume "a \<in> (\<lambda>x. x $ i) ` S" + then obtain z where "a = z $ i" and "z \<in> S" .. + then obtain A where A: "\<forall>i. open (A i) \<and> z $ i \<in> A i" + and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" + using \<open>open S\<close> unfolding open_vec_def by auto + hence "A i \<subseteq> (\<lambda>x. x $ i) ` S" + by (clarsimp, rule_tac x="\<chi> j. if j = i then x else z $ j" in image_eqI, + simp_all) + hence "open (A i) \<and> a \<in> A i \<and> A i \<subseteq> (\<lambda>x. x $ i) ` S" + using A \<open>a = z $ i\<close> by simp + then show "\<exists>T. open T \<and> a \<in> T \<and> T \<subseteq> (\<lambda>x. x $ i) ` S" by - (rule exI) +qed + +instance vec :: (perfect_space, finite) perfect_space +proof + fix x :: "'a ^ 'b" show "\<not> open {x}" + proof + assume "open {x}" + hence "\<forall>i. open ((\<lambda>x. x $ i) ` {x})" by (fast intro: open_image_vec_nth) + hence "\<forall>i. open {x $ i}" by simp + thus "False" by (simp add: not_open_singleton) + qed +qed + + +subsection \<open>Metric space\<close> +(* TODO: Product of uniform spaces and compatibility with metric_spaces! *) + +instantiation vec :: (metric_space, finite) dist +begin + +definition + "dist x y = setL2 (\<lambda>i. dist (x$i) (y$i)) UNIV" + +instance .. +end + +instantiation vec :: (metric_space, finite) uniformity_dist +begin + +definition [code del]: + "(uniformity :: (('a, 'b) vec \<times> ('a, 'b) vec) filter) = + (INF e:{0 <..}. principal {(x, y). dist x y < e})" + +instance + by standard (rule uniformity_vec_def) +end + +declare uniformity_Abort[where 'a="'a :: metric_space ^ 'b :: finite", code] + +instantiation vec :: (metric_space, finite) metric_space +begin + +lemma dist_vec_nth_le: "dist (x $ i) (y $ i) \<le> dist x y" + unfolding dist_vec_def by (rule member_le_setL2) simp_all + +instance proof + fix x y :: "'a ^ 'b" + show "dist x y = 0 \<longleftrightarrow> x = y" + unfolding dist_vec_def + by (simp add: setL2_eq_0_iff vec_eq_iff) +next + fix x y z :: "'a ^ 'b" + show "dist x y \<le> dist x z + dist y z" + unfolding dist_vec_def + apply (rule order_trans [OF _ setL2_triangle_ineq]) + apply (simp add: setL2_mono dist_triangle2) + done +next + fix S :: "('a ^ 'b) set" + have *: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" + proof + assume "open S" show "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" + proof + fix x assume "x \<in> S" + obtain A where A: "\<forall>i. open (A i)" "\<forall>i. x $ i \<in> A i" + and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" + using \<open>open S\<close> and \<open>x \<in> S\<close> unfolding open_vec_def by metis + have "\<forall>i\<in>UNIV. \<exists>r>0. \<forall>y. dist y (x $ i) < r \<longrightarrow> y \<in> A i" + using A unfolding open_dist by simp + hence "\<exists>r. \<forall>i\<in>UNIV. 0 < r i \<and> (\<forall>y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i)" + by (rule finite_set_choice [OF finite]) + then obtain r where r1: "\<forall>i. 0 < r i" + and r2: "\<forall>i y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i" by fast + have "0 < Min (range r) \<and> (\<forall>y. dist y x < Min (range r) \<longrightarrow> y \<in> S)" + by (simp add: r1 r2 S le_less_trans [OF dist_vec_nth_le]) + thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" .. + qed + next + assume *: "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" show "open S" + proof (unfold open_vec_def, rule) + fix x assume "x \<in> S" + then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S" + using * by fast + define r where [abs_def]: "r i = e / sqrt (of_nat CARD('b))" for i :: 'b + from \<open>0 < e\<close> have r: "\<forall>i. 0 < r i" + unfolding r_def by simp_all + from \<open>0 < e\<close> have e: "e = setL2 r UNIV" + unfolding r_def by (simp add: setL2_constant) + define A where "A i = {y. dist (x $ i) y < r i}" for i + have "\<forall>i. open (A i) \<and> x $ i \<in> A i" + unfolding A_def by (simp add: open_ball r) + moreover have "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" + by (simp add: A_def S dist_vec_def e setL2_strict_mono dist_commute) + ultimately show "\<exists>A. (\<forall>i. open (A i) \<and> x $ i \<in> A i) \<and> + (\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S)" by metis + qed + qed + show "open S = (\<forall>x\<in>S. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> S)" + unfolding * eventually_uniformity_metric + by (simp del: split_paired_All add: dist_vec_def dist_commute) +qed + +end + +lemma Cauchy_vec_nth: + "Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)" + unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_vec_nth_le]) + +lemma vec_CauchyI: + fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n" + assumes X: "\<And>i. Cauchy (\<lambda>n. X n $ i)" + shows "Cauchy (\<lambda>n. X n)" +proof (rule metric_CauchyI) + fix r :: real assume "0 < r" + hence "0 < r / of_nat CARD('n)" (is "0 < ?s") by simp + define N where "N i = (LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s)" for i + define M where "M = Max (range N)" + have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s" + using X \<open>0 < ?s\<close> by (rule metric_CauchyD) + hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s" + unfolding N_def by (rule LeastI_ex) + hence M: "\<And>i. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m $ i) (X n $ i) < ?s" + unfolding M_def by simp + { + fix m n :: nat + assume "M \<le> m" "M \<le> n" + have "dist (X m) (X n) = setL2 (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV" + unfolding dist_vec_def .. + also have "\<dots> \<le> setsum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV" + by (rule setL2_le_setsum [OF zero_le_dist]) + also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV" + by (rule setsum_strict_mono, simp_all add: M \<open>M \<le> m\<close> \<open>M \<le> n\<close>) + also have "\<dots> = r" + by simp + finally have "dist (X m) (X n) < r" . + } + hence "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" + by simp + then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" .. +qed + +instance vec :: (complete_space, finite) complete_space +proof + fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X" + have "\<And>i. (\<lambda>n. X n $ i) \<longlonglongrightarrow> lim (\<lambda>n. X n $ i)" + using Cauchy_vec_nth [OF \<open>Cauchy X\<close>] + by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) + hence "X \<longlonglongrightarrow> vec_lambda (\<lambda>i. lim (\<lambda>n. X n $ i))" + by (simp add: vec_tendstoI) + then show "convergent X" + by (rule convergentI) +qed + + +subsection \<open>Normed vector space\<close> + +instantiation vec :: (real_normed_vector, finite) real_normed_vector +begin + +definition "norm x = setL2 (\<lambda>i. norm (x$i)) UNIV" + +definition "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" + +instance proof + fix a :: real and x y :: "'a ^ 'b" + show "norm x = 0 \<longleftrightarrow> x = 0" + unfolding norm_vec_def + by (simp add: setL2_eq_0_iff vec_eq_iff) + show "norm (x + y) \<le> norm x + norm y" + unfolding norm_vec_def + apply (rule order_trans [OF _ setL2_triangle_ineq]) + apply (simp add: setL2_mono norm_triangle_ineq) + done + show "norm (scaleR a x) = \<bar>a\<bar> * norm x" + unfolding norm_vec_def + by (simp add: setL2_right_distrib) + show "sgn x = scaleR (inverse (norm x)) x" + by (rule sgn_vec_def) + show "dist x y = norm (x - y)" + unfolding dist_vec_def norm_vec_def + by (simp add: dist_norm) +qed + +end + +lemma norm_nth_le: "norm (x $ i) \<le> norm x" +unfolding norm_vec_def +by (rule member_le_setL2) simp_all + +lemma bounded_linear_vec_nth: "bounded_linear (\<lambda>x. x $ i)" +apply standard +apply (rule vector_add_component) +apply (rule vector_scaleR_component) +apply (rule_tac x="1" in exI, simp add: norm_nth_le) +done + +instance vec :: (banach, finite) banach .. + + +subsection \<open>Inner product space\<close> + +instantiation vec :: (real_inner, finite) real_inner +begin + +definition "inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) UNIV" + +instance proof + fix r :: real and x y z :: "'a ^ 'b" + show "inner x y = inner y x" + unfolding inner_vec_def + by (simp add: inner_commute) + show "inner (x + y) z = inner x z + inner y z" + unfolding inner_vec_def + by (simp add: inner_add_left setsum.distrib) + show "inner (scaleR r x) y = r * inner x y" + unfolding inner_vec_def + by (simp add: setsum_right_distrib) + show "0 \<le> inner x x" + unfolding inner_vec_def + by (simp add: setsum_nonneg) + show "inner x x = 0 \<longleftrightarrow> x = 0" + unfolding inner_vec_def + by (simp add: vec_eq_iff setsum_nonneg_eq_0_iff) + show "norm x = sqrt (inner x x)" + unfolding inner_vec_def norm_vec_def setL2_def + by (simp add: power2_norm_eq_inner) +qed + +end + + +subsection \<open>Euclidean space\<close> + +text \<open>Vectors pointing along a single axis.\<close> + +definition "axis k x = (\<chi> i. if i = k then x else 0)" + +lemma axis_nth [simp]: "axis i x $ i = x" + unfolding axis_def by simp + +lemma axis_eq_axis: "axis i x = axis j y \<longleftrightarrow> x = y \<and> i = j \<or> x = 0 \<and> y = 0" + unfolding axis_def vec_eq_iff by auto + +lemma inner_axis_axis: + "inner (axis i x) (axis j y) = (if i = j then inner x y else 0)" + unfolding inner_vec_def + apply (cases "i = j") + apply clarsimp + apply (subst setsum.remove [of _ j], simp_all) + apply (rule setsum.neutral, simp add: axis_def) + apply (rule setsum.neutral, simp add: axis_def) + done + +lemma setsum_single: + assumes "finite A" and "k \<in> A" and "f k = y" + assumes "\<And>i. i \<in> A \<Longrightarrow> i \<noteq> k \<Longrightarrow> f i = 0" + shows "(\<Sum>i\<in>A. f i) = y" + apply (subst setsum.remove [OF assms(1,2)]) + apply (simp add: setsum.neutral assms(3,4)) + done + +lemma inner_axis: "inner x (axis i y) = inner (x $ i) y" + unfolding inner_vec_def + apply (rule_tac k=i in setsum_single) + apply simp_all + apply (simp add: axis_def) + done + +instantiation vec :: (euclidean_space, finite) euclidean_space +begin + +definition "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})" + +instance proof + show "(Basis :: ('a ^ 'b) set) \<noteq> {}" + unfolding Basis_vec_def by simp +next + show "finite (Basis :: ('a ^ 'b) set)" + unfolding Basis_vec_def by simp +next + fix u v :: "'a ^ 'b" + assume "u \<in> Basis" and "v \<in> Basis" + thus "inner u v = (if u = v then 1 else 0)" + unfolding Basis_vec_def + by (auto simp add: inner_axis_axis axis_eq_axis inner_Basis) +next + fix x :: "'a ^ 'b" + show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0" + unfolding Basis_vec_def + by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff) +qed + +lemma DIM_cart[simp]: "DIM('a^'b) = CARD('b) * DIM('a)" + apply (simp add: Basis_vec_def) + apply (subst card_UN_disjoint) + apply simp + apply simp + apply (auto simp: axis_eq_axis) [1] + apply (subst card_UN_disjoint) + apply (auto simp: axis_eq_axis) + done + +end + +lemma cart_eq_inner_axis: "a $ i = inner a (axis i 1)" + by (simp add: inner_axis) + +lemma axis_in_Basis: "a \<in> Basis \<Longrightarrow> axis i a \<in> Basis" + by (auto simp add: Basis_vec_def axis_eq_axis) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Finite_Product_Measure.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,1199 @@ +(* Title: HOL/Analysis/Finite_Product_Measure.thy + Author: Johannes Hölzl, TU München +*) + +section \<open>Finite product measures\<close> + +theory Finite_Product_Measure +imports Binary_Product_Measure +begin + +lemma PiE_choice: "(\<exists>f\<in>PiE I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)" + by (auto simp: Bex_def PiE_iff Ball_def dest!: choice_iff'[THEN iffD1]) + (force intro: exI[of _ "restrict f I" for f]) + +lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)" + by auto + +subsubsection \<open>More about Function restricted by @{const extensional}\<close> + +definition + "merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)" + +lemma merge_apply[simp]: + "I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i" + "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i" + "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i" + "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i" + "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I J (x, y) i = undefined" + unfolding merge_def by auto + +lemma merge_commute: + "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) = merge J I (y, x)" + by (force simp: merge_def) + +lemma Pi_cancel_merge_range[simp]: + "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A" + "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A" + "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A" + "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A" + by (auto simp: Pi_def) + +lemma Pi_cancel_merge[simp]: + "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" + "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" + "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B" + "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B" + by (auto simp: Pi_def) + +lemma extensional_merge[simp]: "merge I J (x, y) \<in> extensional (I \<union> J)" + by (auto simp: extensional_def) + +lemma restrict_merge[simp]: + "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I" + "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J" + "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I" + "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J" + by (auto simp: restrict_def) + +lemma split_merge: "P (merge I J (x,y) i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)" + unfolding merge_def by auto + +lemma PiE_cancel_merge[simp]: + "I \<inter> J = {} \<Longrightarrow> + merge I J (x, y) \<in> PiE (I \<union> J) B \<longleftrightarrow> x \<in> Pi I B \<and> y \<in> Pi J B" + by (auto simp: PiE_def restrict_Pi_cancel) + +lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)" + unfolding merge_def by (auto simp: fun_eq_iff) + +lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I J (x, y) \<in> extensional K" + unfolding merge_def extensional_def by auto + +lemma merge_restrict[simp]: + "merge I J (restrict x I, y) = merge I J (x, y)" + "merge I J (x, restrict y J) = merge I J (x, y)" + unfolding merge_def by auto + +lemma merge_x_x_eq_restrict[simp]: + "merge I J (x, x) = restrict x (I \<union> J)" + unfolding merge_def by auto + +lemma injective_vimage_restrict: + assumes J: "J \<subseteq> I" + and sets: "A \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" and ne: "(\<Pi>\<^sub>E i\<in>I. S i) \<noteq> {}" + and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)" + shows "A = B" +proof (intro set_eqI) + fix x + from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto + have "J \<inter> (I - J) = {}" by auto + show "x \<in> A \<longleftrightarrow> x \<in> B" + proof cases + assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)" + have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)" + using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S] + by (auto simp del: PiE_cancel_merge simp add: Un_absorb1) + then show "x \<in> A \<longleftrightarrow> x \<in> B" + using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S] + by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq) + qed (insert sets, auto) +qed + +lemma restrict_vimage: + "I \<inter> J = {} \<Longrightarrow> + (\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^sub>E I E \<times> Pi\<^sub>E J F) = Pi (I \<union> J) (merge I J (E, F))" + by (auto simp: restrict_Pi_cancel PiE_def) + +lemma merge_vimage: + "I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E" + by (auto simp: restrict_Pi_cancel PiE_def) + +subsection \<open>Finite product spaces\<close> + +subsubsection \<open>Products\<close> + +definition prod_emb where + "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))" + +lemma prod_emb_iff: + "f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))" + unfolding prod_emb_def PiE_def by auto + +lemma + shows prod_emb_empty[simp]: "prod_emb M L K {} = {}" + and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B" + and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B" + and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))" + and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))" + and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B" + by (auto simp: prod_emb_def) + +lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> + prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))" + by (force simp: prod_emb_def PiE_iff if_split_mem2) + +lemma prod_emb_PiE_same_index[simp]: + "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E" + by (auto simp: prod_emb_def PiE_iff) + +lemma prod_emb_trans[simp]: + "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X" + by (auto simp add: Int_absorb1 prod_emb_def PiE_def) + +lemma prod_emb_Pi: + assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K" + shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))" + using assms sets.space_closed + by (auto simp: prod_emb_def PiE_iff split: if_split_asm) blast+ + +lemma prod_emb_id: + "B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B" + by (auto simp: prod_emb_def subset_eq extensional_restrict) + +lemma prod_emb_mono: + "F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G" + by (auto simp: prod_emb_def) + +definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where + "PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i)) + {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))} + (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) + (\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" + +definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where + "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) ` + {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}" + +abbreviation + "Pi\<^sub>M I M \<equiv> PiM I M" + +syntax + "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>M _\<in>_./ _)" 10) +translations + "\<Pi>\<^sub>M x\<in>I. M" == "CONST PiM I (%x. M)" + +lemma extend_measure_cong: + assumes "\<Omega> = \<Omega>'" "I = I'" "G = G'" "\<And>i. i \<in> I' \<Longrightarrow> \<mu> i = \<mu>' i" + shows "extend_measure \<Omega> I G \<mu> = extend_measure \<Omega>' I' G' \<mu>'" + unfolding extend_measure_def by (auto simp add: assms) + +lemma Pi_cong_sets: + "\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N" + unfolding Pi_def by auto + +lemma PiM_cong: + assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x" + shows "PiM I M = PiM J N" + unfolding PiM_def +proof (rule extend_measure_cong, goal_cases) + case 1 + show ?case using assms + by (subst assms(1), intro PiE_cong[of J "\<lambda>i. space (M i)" "\<lambda>i. space (N i)"]) simp_all +next + case 2 + have "\<And>K. K \<subseteq> J \<Longrightarrow> (\<Pi> j\<in>K. sets (M j)) = (\<Pi> j\<in>K. sets (N j))" + using assms by (intro Pi_cong_sets) auto + thus ?case by (auto simp: assms) +next + case 3 + show ?case using assms + by (intro ext) (auto simp: prod_emb_def dest: PiE_mem) +next + case (4 x) + thus ?case using assms + by (auto intro!: setprod.cong split: if_split_asm) +qed + + +lemma prod_algebra_sets_into_space: + "prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))" + by (auto simp: prod_emb_def prod_algebra_def) + +lemma prod_algebra_eq_finite: + assumes I: "finite I" + shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R") +proof (intro iffI set_eqI) + fix A assume "A \<in> ?L" + then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)" + and A: "A = prod_emb I M J (PIE j:J. E j)" + by (auto simp: prod_algebra_def) + let ?A = "\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i)" + have A: "A = ?A" + unfolding A using J by (intro prod_emb_PiE sets.sets_into_space) auto + show "A \<in> ?R" unfolding A using J sets.top + by (intro CollectI exI[of _ "\<lambda>i. if i \<in> J then E i else space (M i)"]) simp +next + fix A assume "A \<in> ?R" + then obtain X where A: "A = (\<Pi>\<^sub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto + then have A: "A = prod_emb I M I (\<Pi>\<^sub>E i\<in>I. X i)" + by (simp add: prod_emb_PiE_same_index[OF sets.sets_into_space] Pi_iff) + from X I show "A \<in> ?L" unfolding A + by (auto simp: prod_algebra_def) +qed + +lemma prod_algebraI: + "finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)) + \<Longrightarrow> prod_emb I M J (PIE j:J. E j) \<in> prod_algebra I M" + by (auto simp: prod_algebra_def) + +lemma prod_algebraI_finite: + "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M" + using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp + +lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}" +proof (safe intro!: Int_stableI) + fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)" + then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))" + by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int) +qed + +lemma prod_algebraE: + assumes A: "A \<in> prod_algebra I M" + obtains J E where "A = prod_emb I M J (PIE j:J. E j)" + "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)" + using A by (auto simp: prod_algebra_def) + +lemma prod_algebraE_all: + assumes A: "A \<in> prod_algebra I M" + obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))" +proof - + from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)" + and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))" + by (auto simp: prod_algebra_def) + from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)" + using sets.sets_into_space by auto + then have "A = (\<Pi>\<^sub>E i\<in>I. if i\<in>J then E i else space (M i))" + using A J by (auto simp: prod_emb_PiE) + moreover have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))" + using sets.top E by auto + ultimately show ?thesis using that by auto +qed + +lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)" +proof (unfold Int_stable_def, safe) + fix A assume "A \<in> prod_algebra I M" + from prod_algebraE[OF this] guess J E . note A = this + fix B assume "B \<in> prod_algebra I M" + from prod_algebraE[OF this] guess K F . note B = this + have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter> + (if i \<in> K then F i else space (M i)))" + unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4) + B(5)[THEN sets.sets_into_space] + apply (subst (1 2 3) prod_emb_PiE) + apply (simp_all add: subset_eq PiE_Int) + apply blast + apply (intro PiE_cong) + apply auto + done + also have "\<dots> \<in> prod_algebra I M" + using A B by (auto intro!: prod_algebraI) + finally show "A \<inter> B \<in> prod_algebra I M" . +qed + +lemma prod_algebra_mono: + assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)" + assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)" + shows "prod_algebra I E \<subseteq> prod_algebra I F" +proof + fix A assume "A \<in> prod_algebra I E" + then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" + and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)" + and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)" + by (auto simp: prod_algebra_def) + moreover + from space have "(\<Pi>\<^sub>E i\<in>I. space (E i)) = (\<Pi>\<^sub>E i\<in>I. space (F i))" + by (rule PiE_cong) + with A have "A = prod_emb I F J (\<Pi>\<^sub>E i\<in>J. G i)" + by (simp add: prod_emb_def) + moreover + from sets G J have "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (F i)" + by auto + ultimately show "A \<in> prod_algebra I F" + apply (simp add: prod_algebra_def image_iff) + apply (intro exI[of _ J] exI[of _ G] conjI) + apply auto + done +qed + +lemma prod_algebra_cong: + assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))" + shows "prod_algebra I M = prod_algebra J N" +proof - + have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)" + using sets_eq_imp_space_eq[OF sets] by auto + with sets show ?thesis unfolding \<open>I = J\<close> + by (intro antisym prod_algebra_mono) auto +qed + +lemma space_in_prod_algebra: + "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M" +proof cases + assume "I = {}" then show ?thesis + by (auto simp add: prod_algebra_def image_iff prod_emb_def) +next + assume "I \<noteq> {}" + then obtain i where "i \<in> I" by auto + then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))" + by (auto simp: prod_emb_def) + also have "\<dots> \<in> prod_algebra I M" + using \<open>i \<in> I\<close> by (intro prod_algebraI) auto + finally show ?thesis . +qed + +lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" + using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp + +lemma prod_emb_subset_PiM[simp]: "prod_emb I M K X \<subseteq> space (PiM I M)" + by (auto simp: prod_emb_def space_PiM) + +lemma space_PiM_empty_iff[simp]: "space (PiM I M) = {} \<longleftrightarrow> (\<exists>i\<in>I. space (M i) = {})" + by (auto simp: space_PiM PiE_eq_empty_iff) + +lemma undefined_in_PiM_empty[simp]: "(\<lambda>x. undefined) \<in> space (PiM {} M)" + by (auto simp: space_PiM) + +lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)" + using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp + +lemma sets_PiM_single: "sets (PiM I M) = + sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}" + (is "_ = sigma_sets ?\<Omega> ?R") + unfolding sets_PiM +proof (rule sigma_sets_eqI) + interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto + fix A assume "A \<in> prod_algebra I M" + from prod_algebraE[OF this] guess J X . note X = this + show "A \<in> sigma_sets ?\<Omega> ?R" + proof cases + assume "I = {}" + with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def) + with \<open>I = {}\<close> show ?thesis by (auto intro!: sigma_sets_top) + next + assume "I \<noteq> {}" + with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})" + by (auto simp: prod_emb_def) + also have "\<dots> \<in> sigma_sets ?\<Omega> ?R" + using X \<open>I \<noteq> {}\<close> by (intro R.finite_INT sigma_sets.Basic) auto + finally show "A \<in> sigma_sets ?\<Omega> ?R" . + qed +next + fix A assume "A \<in> ?R" + then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)" + by auto + then have "A = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. B)" + by (auto simp: prod_emb_def) + also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)" + using A by (intro sigma_sets.Basic prod_algebraI) auto + finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" . +qed + +lemma sets_PiM_eq_proj: + "I \<noteq> {} \<Longrightarrow> sets (PiM I M) = sets (SUP i:I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i))" + apply (simp add: sets_PiM_single) + apply (subst sets_Sup_eq[where X="\<Pi>\<^sub>E i\<in>I. space (M i)"]) + apply auto [] + apply auto [] + apply simp + apply (subst SUP_cong[OF refl]) + apply (rule sets_vimage_algebra2) + apply auto [] + apply (auto intro!: arg_cong2[where f=sigma_sets]) + done + +lemma + shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}" + and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\<lambda>k. undefined} }" + by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq) + +lemma sets_PiM_sigma: + assumes \<Omega>_cover: "\<And>i. i \<in> I \<Longrightarrow> \<exists>S\<subseteq>E i. countable S \<and> \<Omega> i = \<Union>S" + assumes E: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (\<Omega> i)" + assumes J: "\<And>j. j \<in> J \<Longrightarrow> finite j" "\<Union>J = I" + defines "P \<equiv> {{f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i} | A j. j \<in> J \<and> A \<in> Pi j E}" + shows "sets (\<Pi>\<^sub>M i\<in>I. sigma (\<Omega> i) (E i)) = sets (sigma (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P)" +proof cases + assume "I = {}" + with \<open>\<Union>J = I\<close> have "P = {{\<lambda>_. undefined}} \<or> P = {}" + by (auto simp: P_def) + with \<open>I = {}\<close> show ?thesis + by (auto simp add: sets_PiM_empty sigma_sets_empty_eq) +next + let ?F = "\<lambda>i. {(\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega> |A. A \<in> E i}" + assume "I \<noteq> {}" + then have "sets (Pi\<^sub>M I (\<lambda>i. sigma (\<Omega> i) (E i))) = + sets (SUP i:I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<lambda>x. x i) (sigma (\<Omega> i) (E i)))" + by (subst sets_PiM_eq_proj) (auto simp: space_measure_of_conv) + also have "\<dots> = sets (SUP i:I. sigma (Pi\<^sub>E I \<Omega>) (?F i))" + using E by (intro sets_SUP_cong arg_cong[where f=sets] vimage_algebra_sigma) auto + also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i))" + using \<open>I \<noteq> {}\<close> by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto + also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) P)" + proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) + show "(\<Union>i\<in>I. ?F i) \<subseteq> Pow (Pi\<^sub>E I \<Omega>)" "P \<subseteq> Pow (Pi\<^sub>E I \<Omega>)" + by (auto simp: P_def) + next + interpret P: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P" + by (auto intro!: sigma_algebra_sigma_sets simp: P_def) + + fix Z assume "Z \<in> (\<Union>i\<in>I. ?F i)" + then obtain i A where i: "i \<in> I" "A \<in> E i" and Z_def: "Z = (\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega>" + by auto + from \<open>i \<in> I\<close> J obtain j where j: "i \<in> j" "j \<in> J" "j \<subseteq> I" "finite j" + by auto + obtain S where S: "\<And>i. i \<in> j \<Longrightarrow> S i \<subseteq> E i" "\<And>i. i \<in> j \<Longrightarrow> countable (S i)" + "\<And>i. i \<in> j \<Longrightarrow> \<Omega> i = \<Union>(S i)" + by (metis subset_eq \<Omega>_cover \<open>j \<subseteq> I\<close>) + define A' where "A' n = n(i := A)" for n + then have A'_i: "\<And>n. A' n i = A" + by simp + { fix n assume "n \<in> Pi\<^sub>E (j - {i}) S" + then have "A' n \<in> Pi j E" + unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def \<open>A \<in> E i\<close> ) + with \<open>j \<in> J\<close> have "{f \<in> Pi\<^sub>E I \<Omega>. \<forall>i\<in>j. f i \<in> A' n i} \<in> P" + by (auto simp: P_def) } + note A'_in_P = this + + { fix x assume "x i \<in> A" "x \<in> Pi\<^sub>E I \<Omega>" + with S(3) \<open>j \<subseteq> I\<close> have "\<forall>i\<in>j. \<exists>s\<in>S i. x i \<in> s" + by (auto simp: PiE_def Pi_def) + then obtain s where s: "\<And>i. i \<in> j \<Longrightarrow> s i \<in> S i" "\<And>i. i \<in> j \<Longrightarrow> x i \<in> s i" + by metis + with \<open>x i \<in> A\<close> have "\<exists>n\<in>PiE (j-{i}) S. \<forall>i\<in>j. x i \<in> A' n i" + by (intro bexI[of _ "restrict (s(i := A)) (j-{i})"]) (auto simp: A'_def split: if_splits) } + then have "Z = (\<Union>n\<in>PiE (j-{i}) S. {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A' n i})" + unfolding Z_def + by (auto simp add: set_eq_iff ball_conj_distrib \<open>i\<in>j\<close> A'_i dest: bspec[OF _ \<open>i\<in>j\<close>] + cong: conj_cong) + also have "\<dots> \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P" + using \<open>finite j\<close> S(2) + by (intro P.countable_UN' countable_PiE) (simp_all add: image_subset_iff A'_in_P) + finally show "Z \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P" . + next + interpret F: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<Union>i\<in>I. ?F i)" + by (auto intro!: sigma_algebra_sigma_sets) + + fix b assume "b \<in> P" + then obtain A j where b: "b = {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i}" "j \<in> J" "A \<in> Pi j E" + by (auto simp: P_def) + show "b \<in> sigma_sets (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i)" + proof cases + assume "j = {}" + with b have "b = (\<Pi>\<^sub>E i\<in>I. \<Omega> i)" + by auto + then show ?thesis + by blast + next + assume "j \<noteq> {}" + with J b(2,3) have eq: "b = (\<Inter>i\<in>j. ((\<lambda>x. x i) -` A i \<inter> Pi\<^sub>E I \<Omega>))" + unfolding b(1) + by (auto simp: PiE_def Pi_def) + show ?thesis + unfolding eq using \<open>A \<in> Pi j E\<close> \<open>j \<in> J\<close> J(2) + by (intro F.finite_INT J \<open>j \<in> J\<close> \<open>j \<noteq> {}\<close> sigma_sets.Basic) blast + qed + qed + finally show "?thesis" . +qed + +lemma sets_PiM_in_sets: + assumes space: "space N = (\<Pi>\<^sub>E i\<in>I. space (M i))" + assumes sets: "\<And>i A. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {x\<in>space N. x i \<in> A} \<in> sets N" + shows "sets (\<Pi>\<^sub>M i \<in> I. M i) \<subseteq> sets N" + unfolding sets_PiM_single space[symmetric] + by (intro sets.sigma_sets_subset subsetI) (auto intro: sets) + +lemma sets_PiM_cong[measurable_cong]: + assumes "I = J" "\<And>i. i \<in> J \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (PiM I M) = sets (PiM J N)" + using assms sets_eq_imp_space_eq[OF assms(2)] by (simp add: sets_PiM_single cong: PiE_cong conj_cong) + +lemma sets_PiM_I: + assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)" + shows "prod_emb I M J (PIE j:J. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)" +proof cases + assume "J = {}" + then have "prod_emb I M J (PIE j:J. E j) = (PIE j:I. space (M j))" + by (auto simp: prod_emb_def) + then show ?thesis + by (auto simp add: sets_PiM intro!: sigma_sets_top) +next + assume "J \<noteq> {}" with assms show ?thesis + by (force simp add: sets_PiM prod_algebra_def) +qed + +lemma measurable_PiM: + assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" + assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow> + f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N" + shows "f \<in> measurable N (PiM I M)" + using sets_PiM prod_algebra_sets_into_space space +proof (rule measurable_sigma_sets) + fix A assume "A \<in> prod_algebra I M" + from prod_algebraE[OF this] guess J X . + with sets[of J X] show "f -` A \<inter> space N \<in> sets N" by auto +qed + +lemma measurable_PiM_Collect: + assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" + assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow> + {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N" + shows "f \<in> measurable N (PiM I M)" + using sets_PiM prod_algebra_sets_into_space space +proof (rule measurable_sigma_sets) + fix A assume "A \<in> prod_algebra I M" + from prod_algebraE[OF this] guess J X . note X = this + then have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}" + using space by (auto simp: prod_emb_def del: PiE_I) + also have "\<dots> \<in> sets N" using X(3,2,4,5) by (rule sets) + finally show "f -` A \<inter> space N \<in> sets N" . +qed + +lemma measurable_PiM_single: + assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" + assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N" + shows "f \<in> measurable N (PiM I M)" + using sets_PiM_single +proof (rule measurable_sigma_sets) + fix A assume "A \<in> {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}" + then obtain B i where "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)" + by auto + with space have "f -` A \<inter> space N = {\<omega> \<in> space N. f \<omega> i \<in> B}" by auto + also have "\<dots> \<in> sets N" using B by (rule sets) + finally show "f -` A \<inter> space N \<in> sets N" . +qed (auto simp: space) + +lemma measurable_PiM_single': + assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)" + and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" + shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^sub>M I M)" +proof (rule measurable_PiM_single) + fix A i assume A: "i \<in> I" "A \<in> sets (M i)" + then have "{\<omega> \<in> space N. f i \<omega> \<in> A} = f i -` A \<inter> space N" + by auto + then show "{\<omega> \<in> space N. f i \<omega> \<in> A} \<in> sets N" + using A f by (auto intro!: measurable_sets) +qed fact + +lemma sets_PiM_I_finite[measurable]: + assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))" + shows "(PIE j:I. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)" + using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] \<open>finite I\<close> sets by auto + +lemma measurable_component_singleton[measurable (raw)]: + assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (M i)" +proof (unfold measurable_def, intro CollectI conjI ballI) + fix A assume "A \<in> sets (M i)" + then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) = prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. A)" + using sets.sets_into_space \<open>i \<in> I\<close> + by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: if_split_asm) + then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)" + using \<open>A \<in> sets (M i)\<close> \<open>i \<in> I\<close> by (auto intro!: sets_PiM_I) +qed (insert \<open>i \<in> I\<close>, auto simp: space_PiM) + +lemma measurable_component_singleton'[measurable_dest]: + assumes f: "f \<in> measurable N (Pi\<^sub>M I M)" + assumes g: "g \<in> measurable L N" + assumes i: "i \<in> I" + shows "(\<lambda>x. (f (g x)) i) \<in> measurable L (M i)" + using measurable_compose[OF measurable_compose[OF g f] measurable_component_singleton, OF i] . + +lemma measurable_PiM_component_rev: + "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<lambda>x. f (x i)) \<in> measurable (PiM I M) N" + by simp + +lemma measurable_case_nat[measurable (raw)]: + assumes [measurable (raw)]: "i = 0 \<Longrightarrow> f \<in> measurable M N" + "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N" + shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N" + by (cases i) simp_all + +lemma measurable_case_nat'[measurable (raw)]: + assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" + shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" + using fg[THEN measurable_space] + by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split) + +lemma measurable_add_dim[measurable]: + "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M)" + (is "?f \<in> measurable ?P ?I") +proof (rule measurable_PiM_single) + fix j A assume j: "j \<in> insert i I" and A: "A \<in> sets (M j)" + have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} = + (if j = i then space (Pi\<^sub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)" + using sets.sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM) + also have "\<dots> \<in> sets ?P" + using A j + by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) + finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" . +qed (auto simp: space_pair_measure space_PiM PiE_def) + +lemma measurable_fun_upd: + assumes I: "I = J \<union> {i}" + assumes f[measurable]: "f \<in> measurable N (PiM J M)" + assumes h[measurable]: "h \<in> measurable N (M i)" + shows "(\<lambda>x. (f x) (i := h x)) \<in> measurable N (PiM I M)" +proof (intro measurable_PiM_single') + fix j assume "j \<in> I" then show "(\<lambda>\<omega>. ((f \<omega>)(i := h \<omega>)) j) \<in> measurable N (M j)" + unfolding I by (cases "j = i") auto +next + show "(\<lambda>x. (f x)(i := h x)) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))" + using I f[THEN measurable_space] h[THEN measurable_space] + by (auto simp: space_PiM PiE_iff extensional_def) +qed + +lemma measurable_component_update: + "x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)" + by simp + +lemma measurable_merge[measurable]: + "merge I J \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M)" + (is "?f \<in> measurable ?P ?U") +proof (rule measurable_PiM_single) + fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J" + then have "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} = + (if i \<in> I then ((\<lambda>x. x i) \<circ> fst) -` A \<inter> space ?P else ((\<lambda>x. x i) \<circ> snd) -` A \<inter> space ?P)" + by (auto simp: merge_def) + also have "\<dots> \<in> sets ?P" + using A + by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) + finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" . +qed (auto simp: space_pair_measure space_PiM PiE_iff merge_def extensional_def) + +lemma measurable_restrict[measurable (raw)]: + assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)" + shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^sub>M I M)" +proof (rule measurable_PiM_single) + fix A i assume A: "i \<in> I" "A \<in> sets (M i)" + then have "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} = X i -` A \<inter> space N" + by auto + then show "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} \<in> sets N" + using A X by (auto intro!: measurable_sets) +qed (insert X, auto simp add: PiE_def dest: measurable_space) + +lemma measurable_abs_UNIV: + "(\<And>n. (\<lambda>\<omega>. f n \<omega>) \<in> measurable M (N n)) \<Longrightarrow> (\<lambda>\<omega> n. f n \<omega>) \<in> measurable M (PiM UNIV N)" + by (intro measurable_PiM_single) (auto dest: measurable_space) + +lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)" + by (intro measurable_restrict measurable_component_singleton) auto + +lemma measurable_restrict_subset': + assumes "J \<subseteq> L" "\<And>x. x \<in> J \<Longrightarrow> sets (M x) = sets (N x)" + shows "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)" +proof- + from assms(1) have "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)" + by (rule measurable_restrict_subset) + also from assms(2) have "measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M) = measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)" + by (intro sets_PiM_cong measurable_cong_sets) simp_all + finally show ?thesis . +qed + +lemma measurable_prod_emb[intro, simp]: + "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)" + unfolding prod_emb_def space_PiM[symmetric] + by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton) + +lemma merge_in_prod_emb: + assumes "y \<in> space (PiM I M)" "x \<in> X" and X: "X \<in> sets (Pi\<^sub>M J M)" and "J \<subseteq> I" + shows "merge J I (x, y) \<in> prod_emb I M J X" + using assms sets.sets_into_space[OF X] + by (simp add: merge_def prod_emb_def subset_eq space_PiM PiE_def extensional_restrict Pi_iff + cong: if_cong restrict_cong) + (simp add: extensional_def) + +lemma prod_emb_eq_emptyD: + assumes J: "J \<subseteq> I" and ne: "space (PiM I M) \<noteq> {}" and X: "X \<in> sets (Pi\<^sub>M J M)" + and *: "prod_emb I M J X = {}" + shows "X = {}" +proof safe + fix x assume "x \<in> X" + obtain \<omega> where "\<omega> \<in> space (PiM I M)" + using ne by blast + from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto +qed + +lemma sets_in_Pi_aux: + "finite I \<Longrightarrow> (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow> + {x\<in>space (PiM I M). x \<in> Pi I F} \<in> sets (PiM I M)" + by (simp add: subset_eq Pi_iff) + +lemma sets_in_Pi[measurable (raw)]: + "finite I \<Longrightarrow> f \<in> measurable N (PiM I M) \<Longrightarrow> + (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow> + Measurable.pred N (\<lambda>x. f x \<in> Pi I F)" + unfolding pred_def + by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_Pi_aux]) auto + +lemma sets_in_extensional_aux: + "{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)" +proof - + have "{x\<in>space (PiM I M). x \<in> extensional I} = space (PiM I M)" + by (auto simp add: extensional_def space_PiM) + then show ?thesis by simp +qed + +lemma sets_in_extensional[measurable (raw)]: + "f \<in> measurable N (PiM I M) \<Longrightarrow> Measurable.pred N (\<lambda>x. f x \<in> extensional I)" + unfolding pred_def + by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto + +lemma sets_PiM_I_countable: + assumes I: "countable I" and E: "\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i)" shows "Pi\<^sub>E I E \<in> sets (Pi\<^sub>M I M)" +proof cases + assume "I \<noteq> {}" + then have "PiE I E = (\<Inter>i\<in>I. prod_emb I M {i} (PiE {i} E))" + using E[THEN sets.sets_into_space] by (auto simp: PiE_iff prod_emb_def fun_eq_iff) + also have "\<dots> \<in> sets (PiM I M)" + using I \<open>I \<noteq> {}\<close> by (safe intro!: sets.countable_INT' measurable_prod_emb sets_PiM_I_finite E) + finally show ?thesis . +qed (simp add: sets_PiM_empty) + +lemma sets_PiM_D_countable: + assumes A: "A \<in> PiM I M" + shows "\<exists>J\<subseteq>I. \<exists>X\<in>PiM J M. countable J \<and> A = prod_emb I M J X" + using A[unfolded sets_PiM_single] +proof induction + case (Basic A) + then obtain i X where *: "i \<in> I" "X \<in> sets (M i)" and "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> X}" + by auto + then have A: "A = prod_emb I M {i} (\<Pi>\<^sub>E _\<in>{i}. X)" + by (auto simp: prod_emb_def) + then show ?case + by (intro exI[of _ "{i}"] conjI bexI[of _ "\<Pi>\<^sub>E _\<in>{i}. X"]) + (auto intro: countable_finite * sets_PiM_I_finite) +next + case Empty then show ?case + by (intro exI[of _ "{}"] conjI bexI[of _ "{}"]) auto +next + case (Compl A) + then obtain J X where "J \<subseteq> I" "X \<in> sets (Pi\<^sub>M J M)" "countable J" "A = prod_emb I M J X" + by auto + then show ?case + by (intro exI[of _ J] bexI[of _ "space (PiM J M) - X"] conjI) + (auto simp add: space_PiM prod_emb_PiE intro!: sets_PiM_I_countable) +next + case (Union K) + obtain J X where J: "\<And>i. J i \<subseteq> I" "\<And>i. countable (J i)" and X: "\<And>i. X i \<in> sets (Pi\<^sub>M (J i) M)" + and K: "\<And>i. K i = prod_emb I M (J i) (X i)" + by (metis Union.IH) + show ?case + proof (intro exI[of _ "\<Union>i. J i"] bexI[of _ "\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i)"] conjI) + show "(\<Union>i. J i) \<subseteq> I" "countable (\<Union>i. J i)" using J by auto + with J show "UNION UNIV K = prod_emb I M (\<Union>i. J i) (\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i))" + by (simp add: K[abs_def] SUP_upper) + qed(auto intro: X) +qed + +lemma measure_eqI_PiM_finite: + assumes [simp]: "finite I" "sets P = PiM I M" "sets Q = PiM I M" + assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = Q (Pi\<^sub>E I A)" + assumes A: "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = space (PiM I M)" "\<And>i::nat. P (A i) \<noteq> \<infinity>" + shows "P = Q" +proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space]) + show "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. P (A i) \<noteq> \<infinity>" + unfolding space_PiM[symmetric] by fact+ + fix X assume "X \<in> prod_algebra I M" + then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)" + and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)" + by (force elim!: prod_algebraE) + then show "emeasure P X = emeasure Q X" + unfolding X by (subst (1 2) prod_emb_Pi) (auto simp: eq) +qed (simp_all add: sets_PiM) + +lemma measure_eqI_PiM_infinite: + assumes [simp]: "sets P = PiM I M" "sets Q = PiM I M" + assumes eq: "\<And>A J. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> + P (prod_emb I M J (Pi\<^sub>E J A)) = Q (prod_emb I M J (Pi\<^sub>E J A))" + assumes A: "finite_measure P" + shows "P = Q" +proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space]) + interpret finite_measure P by fact + define i where "i = (SOME i. i \<in> I)" + have i: "I \<noteq> {} \<Longrightarrow> i \<in> I" + unfolding i_def by (rule someI_ex) auto + define A where "A n = + (if I = {} then prod_emb I M {} (\<Pi>\<^sub>E i\<in>{}. {}) else prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)))" + for n :: nat + then show "range A \<subseteq> prod_algebra I M" + using prod_algebraI[of "{}" I "\<lambda>i. space (M i)" M] by (auto intro!: prod_algebraI i) + have "\<And>i. A i = space (PiM I M)" + by (auto simp: prod_emb_def space_PiM PiE_iff A_def i ex_in_conv[symmetric] exI) + then show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. emeasure P (A i) \<noteq> \<infinity>" + by (auto simp: space_PiM) +next + fix X assume X: "X \<in> prod_algebra I M" + then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)" + and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)" + by (force elim!: prod_algebraE) + then show "emeasure P X = emeasure Q X" + by (auto intro!: eq) +qed (auto simp: sets_PiM) + +locale product_sigma_finite = + fixes M :: "'i \<Rightarrow> 'a measure" + assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)" + +sublocale product_sigma_finite \<subseteq> M?: sigma_finite_measure "M i" for i + by (rule sigma_finite_measures) + +locale finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" + + fixes I :: "'i set" + assumes finite_index: "finite I" + +lemma (in finite_product_sigma_finite) sigma_finite_pairs: + "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set. + (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and> + (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k) \<and> + (\<Union>k. \<Pi>\<^sub>E i\<in>I. F i k) = space (PiM I M)" +proof - + have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)" + using M.sigma_finite_incseq by metis + from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" .. + then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. emeasure (M i) (F i k) \<noteq> \<infinity>" + by auto + let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k" + note space_PiM[simp] + show ?thesis + proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) + fix i show "range (F i) \<subseteq> sets (M i)" by fact + next + fix i k show "emeasure (M i) (F i k) \<noteq> \<infinity>" by fact + next + fix x assume "x \<in> (\<Union>i. ?F i)" with F(1) show "x \<in> space (PiM I M)" + by (auto simp: PiE_def dest!: sets.sets_into_space) + next + fix f assume "f \<in> space (PiM I M)" + with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F + show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def PiE_def) + next + fix i show "?F i \<subseteq> ?F (Suc i)" + using \<open>\<And>i. incseq (F i)\<close>[THEN incseq_SucD] by auto + qed +qed + +lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1" +proof - + let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ennreal)" + have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1" + proof (subst emeasure_extend_measure_Pair[OF PiM_def]) + show "positive (PiM {} M) ?\<mu>" + by (auto simp: positive_def) + show "countably_additive (PiM {} M) ?\<mu>" + by (rule sets.countably_additiveI_finite) + (auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: ) + qed (auto simp: prod_emb_def) + also have "(prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = {\<lambda>_. undefined}" + by (auto simp: prod_emb_def) + finally show ?thesis + by simp +qed + +lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}" + by (rule measure_eqI) (auto simp add: sets_PiM_empty) + +lemma (in product_sigma_finite) emeasure_PiM: + "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))" +proof (induct I arbitrary: A rule: finite_induct) + case (insert i I) + interpret finite_product_sigma_finite M I by standard fact + have "finite (insert i I)" using \<open>finite I\<close> by auto + interpret I': finite_product_sigma_finite M "insert i I" by standard fact + let ?h = "(\<lambda>(f, y). f(i := y))" + + let ?P = "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h" + let ?\<mu> = "emeasure ?P" + let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}" + let ?f = "\<lambda>J E j. if j \<in> J then emeasure (M j) (E j) else emeasure (M j) (space (M j))" + + have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) = + (\<Prod>i\<in>insert i I. emeasure (M i) (A i))" + proof (subst emeasure_extend_measure_Pair[OF PiM_def]) + fix J E assume "(J \<noteq> {} \<or> insert i I = {}) \<and> finite J \<and> J \<subseteq> insert i I \<and> E \<in> (\<Pi> j\<in>J. sets (M j))" + then have J: "J \<noteq> {}" "finite J" "J \<subseteq> insert i I" and E: "\<forall>j\<in>J. E j \<in> sets (M j)" by auto + let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)" + let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^sub>E j\<in>J-{i}. E j)" + have "?\<mu> ?p = + emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i))" + by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+ + also have "?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))" + using J E[rule_format, THEN sets.sets_into_space] + by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: if_split_asm) + also have "emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) = + emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))" + using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto + also have "?p' = (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))" + using J E[rule_format, THEN sets.sets_into_space] + by (auto simp: prod_emb_iff PiE_def Pi_iff split: if_split_asm) blast+ + also have "emeasure (Pi\<^sub>M I M) (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) = + (\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))" + using E by (subst insert) (auto intro!: setprod.cong) + also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) * + emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)" + using insert by (auto simp: mult.commute intro!: arg_cong2[where f="op *"] setprod.cong) + also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)" + using insert(1,2) J E by (intro setprod.mono_neutral_right) auto + finally show "?\<mu> ?p = \<dots>" . + + show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \<in> Pow (\<Pi>\<^sub>E i\<in>insert i I. space (M i))" + using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def) + next + show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" + using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all + next + show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and> + insert i I \<subseteq> insert i I \<and> A \<in> (\<Pi> j\<in>insert i I. sets (M j))" + using insert by auto + qed (auto intro!: setprod.cong) + with insert show ?case + by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space) +qed simp + +lemma (in product_sigma_finite) PiM_eqI: + assumes I[simp]: "finite I" and P: "sets P = PiM I M" + assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))" + shows "P = PiM I M" +proof - + interpret finite_product_sigma_finite M I + proof qed fact + from sigma_finite_pairs guess C .. note C = this + show ?thesis + proof (rule measure_eqI_PiM_finite[OF I refl P, symmetric]) + show "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>M I M) (Pi\<^sub>E I A) = P (Pi\<^sub>E I A)" for A + by (simp add: eq emeasure_PiM) + define A where "A n = (\<Pi>\<^sub>E i\<in>I. C i n)" for n + with C show "range A \<subseteq> prod_algebra I M" "\<And>i. emeasure (Pi\<^sub>M I M) (A i) \<noteq> \<infinity>" "(\<Union>i. A i) = space (PiM I M)" + by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_setprod_eq_top) + qed +qed + +lemma (in product_sigma_finite) sigma_finite: + assumes "finite I" + shows "sigma_finite_measure (PiM I M)" +proof + interpret finite_product_sigma_finite M I by standard fact + + obtain F where F: "\<And>j. countable (F j)" "\<And>j f. f \<in> F j \<Longrightarrow> f \<in> sets (M j)" + "\<And>j f. f \<in> F j \<Longrightarrow> emeasure (M j) f \<noteq> \<infinity>" and + in_space: "\<And>j. space (M j) = (\<Union>F j)" + using sigma_finite_countable by (metis subset_eq) + moreover have "(\<Union>(PiE I ` PiE I F)) = space (Pi\<^sub>M I M)" + using in_space by (auto simp: space_PiM PiE_iff intro!: PiE_choice[THEN iffD2]) + ultimately show "\<exists>A. countable A \<and> A \<subseteq> sets (Pi\<^sub>M I M) \<and> \<Union>A = space (Pi\<^sub>M I M) \<and> (\<forall>a\<in>A. emeasure (Pi\<^sub>M I M) a \<noteq> \<infinity>)" + by (intro exI[of _ "PiE I ` PiE I F"]) + (auto intro!: countable_PiE sets_PiM_I_finite + simp: PiE_iff emeasure_PiM finite_index ennreal_setprod_eq_top) +qed + +sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M" + using sigma_finite[OF finite_index] . + +lemma (in finite_product_sigma_finite) measure_times: + "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))" + using emeasure_PiM[OF finite_index] by auto + +lemma (in product_sigma_finite) nn_integral_empty: + "0 \<le> f (\<lambda>k. undefined) \<Longrightarrow> integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)" + by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2) + +lemma (in product_sigma_finite) distr_merge: + assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J" + shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M" + (is "?D = ?P") +proof (rule PiM_eqI) + interpret I: finite_product_sigma_finite M I by standard fact + interpret J: finite_product_sigma_finite M J by standard fact + fix A assume A: "\<And>i. i \<in> I \<union> J \<Longrightarrow> A i \<in> sets (M i)" + have *: "(merge I J -` Pi\<^sub>E (I \<union> J) A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)) = PiE I A \<times> PiE J A" + using A[THEN sets.sets_into_space] by (auto simp: space_PiM space_pair_measure) + from A fin show "emeasure (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) (Pi\<^sub>E (I \<union> J) A) = + (\<Prod>i\<in>I \<union> J. emeasure (M i) (A i))" + by (subst emeasure_distr) + (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times setprod.union_disjoint) +qed (insert fin, simp_all) + +lemma (in product_sigma_finite) product_nn_integral_fold: + assumes IJ: "I \<inter> J = {}" "finite I" "finite J" + and f[measurable]: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)" + shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f = + (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))" +proof - + interpret I: finite_product_sigma_finite M I by standard fact + interpret J: finite_product_sigma_finite M J by standard fact + interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard + have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)" + using measurable_comp[OF measurable_merge f] by (simp add: comp_def) + show ?thesis + apply (subst distr_merge[OF IJ, symmetric]) + apply (subst nn_integral_distr[OF measurable_merge]) + apply measurable [] + apply (subst J.nn_integral_fst[symmetric, OF P_borel]) + apply simp + done +qed + +lemma (in product_sigma_finite) distr_singleton: + "distr (Pi\<^sub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _") +proof (intro measure_eqI[symmetric]) + interpret I: finite_product_sigma_finite M "{i}" by standard simp + fix A assume A: "A \<in> sets (M i)" + then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M {i} M) = (\<Pi>\<^sub>E i\<in>{i}. A)" + using sets.sets_into_space by (auto simp: space_PiM) + then show "emeasure (M i) A = emeasure ?D A" + using A I.measure_times[of "\<lambda>_. A"] + by (simp add: emeasure_distr measurable_component_singleton) +qed simp + +lemma (in product_sigma_finite) product_nn_integral_singleton: + assumes f: "f \<in> borel_measurable (M i)" + shows "integral\<^sup>N (Pi\<^sub>M {i} M) (\<lambda>x. f (x i)) = integral\<^sup>N (M i) f" +proof - + interpret I: finite_product_sigma_finite M "{i}" by standard simp + from f show ?thesis + apply (subst distr_singleton[symmetric]) + apply (subst nn_integral_distr[OF measurable_component_singleton]) + apply simp_all + done +qed + +lemma (in product_sigma_finite) product_nn_integral_insert: + assumes I[simp]: "finite I" "i \<notin> I" + and f: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)" + shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^sub>M I M))" +proof - + interpret I: finite_product_sigma_finite M I by standard auto + interpret i: finite_product_sigma_finite M "{i}" by standard auto + have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I" + using f by auto + show ?thesis + unfolding product_nn_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f] + proof (rule nn_integral_cong, subst product_nn_integral_singleton[symmetric]) + fix x assume x: "x \<in> space (Pi\<^sub>M I M)" + let ?f = "\<lambda>y. f (x(i := y))" + show "?f \<in> borel_measurable (M i)" + using measurable_comp[OF measurable_component_update f, OF x \<open>i \<notin> I\<close>] + unfolding comp_def . + show "(\<integral>\<^sup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^sub>M {i} M) = (\<integral>\<^sup>+ y. f (x(i := y i)) \<partial>Pi\<^sub>M {i} M)" + using x + by (auto intro!: nn_integral_cong arg_cong[where f=f] + simp add: space_PiM extensional_def PiE_def) + qed +qed + +lemma (in product_sigma_finite) product_nn_integral_insert_rev: + assumes I[simp]: "finite I" "i \<notin> I" + and [measurable]: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)" + shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x(i := y)) \<partial>(Pi\<^sub>M I M)) \<partial>(M i))" + apply (subst product_nn_integral_insert[OF assms]) + apply (rule pair_sigma_finite.Fubini') + apply intro_locales [] + apply (rule sigma_finite[OF I(1)]) + apply measurable + done + +lemma (in product_sigma_finite) product_nn_integral_setprod: + assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)" + shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>N (M i) (f i))" +using assms proof (induction I) + case (insert i I) + note insert.prems[measurable] + note \<open>finite I\<close>[intro, simp] + interpret I: finite_product_sigma_finite M I by standard auto + have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" + using insert by (auto intro!: setprod.cong) + have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)" + using sets.sets_into_space insert + by (intro borel_measurable_setprod_ennreal + measurable_comp[OF measurable_component_singleton, unfolded comp_def]) + auto + then show ?case + apply (simp add: product_nn_integral_insert[OF insert(1,2)]) + apply (simp add: insert(2-) * nn_integral_multc) + apply (subst nn_integral_cmult) + apply (auto simp add: insert(2-)) + done +qed (simp add: space_PiM) + +lemma (in product_sigma_finite) product_nn_integral_pair: + assumes [measurable]: "case_prod f \<in> borel_measurable (M x \<Otimes>\<^sub>M M y)" + assumes xy: "x \<noteq> y" + shows "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {x, y} M) = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))" +proof- + interpret psm: pair_sigma_finite "M x" "M y" + unfolding pair_sigma_finite_def using sigma_finite_measures by simp_all + have "{x, y} = {y, x}" by auto + also have "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {y, x} M) = (\<integral>\<^sup>+y. \<integral>\<^sup>+\<sigma>. f (\<sigma> x) y \<partial>PiM {x} M \<partial>M y)" + using xy by (subst product_nn_integral_insert_rev) simp_all + also have "... = (\<integral>\<^sup>+y. \<integral>\<^sup>+x. f x y \<partial>M x \<partial>M y)" + by (intro nn_integral_cong, subst product_nn_integral_singleton) simp_all + also have "... = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))" + by (subst psm.nn_integral_snd[symmetric]) simp_all + finally show ?thesis . +qed + +lemma (in product_sigma_finite) distr_component: + "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P") +proof (intro PiM_eqI) + fix A assume A: "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)" + then have "(\<lambda>x. \<lambda>i\<in>{i}. x) -` Pi\<^sub>E {i} A \<inter> space (M i) = A i" + by (auto dest: sets.sets_into_space) + with A show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x)) (Pi\<^sub>E {i} A) = (\<Prod>i\<in>{i}. emeasure (M i) (A i))" + by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict) +qed simp_all + +lemma (in product_sigma_finite) + assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)" + shows emeasure_fold_integral: + "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" (is ?I) + and emeasure_fold_measurable: + "(\<lambda>x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M))) \<in> borel_measurable (Pi\<^sub>M I M)" (is ?B) +proof - + interpret I: finite_product_sigma_finite M I by standard fact + interpret J: finite_product_sigma_finite M J by standard fact + interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" .. + have merge: "merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) \<in> sets (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)" + by (intro measurable_sets[OF _ A] measurable_merge assms) + + show ?I + apply (subst distr_merge[symmetric, OF IJ]) + apply (subst emeasure_distr[OF measurable_merge A]) + apply (subst J.emeasure_pair_measure_alt[OF merge]) + apply (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure) + done + + show ?B + using IJ.measurable_emeasure_Pair1[OF merge] + by (simp add: vimage_comp comp_def space_pair_measure cong: measurable_cong) +qed + +lemma sets_Collect_single: + "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)" + by simp + +lemma pair_measure_eq_distr_PiM: + fixes M1 :: "'a measure" and M2 :: "'a measure" + assumes "sigma_finite_measure M1" "sigma_finite_measure M2" + shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))" + (is "?P = ?D") +proof (rule pair_measure_eqI[OF assms]) + interpret B: product_sigma_finite "case_bool M1 M2" + unfolding product_sigma_finite_def using assms by (auto split: bool.split) + let ?B = "Pi\<^sub>M UNIV (case_bool M1 M2)" + + have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)" + by auto + fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2" + have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (case_bool M1 M2 i) (case_bool A B i))" + by (simp add: UNIV_bool ac_simps) + also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (case_bool A B))" + using A B by (subst B.emeasure_PiM) (auto split: bool.split) + also have "Pi\<^sub>E UNIV (case_bool A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B" + using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space] + by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split) + finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)" + using A B + measurable_component_singleton[of True UNIV "case_bool M1 M2"] + measurable_component_singleton[of False UNIV "case_bool M1 M2"] + by (subst emeasure_distr) (auto simp: measurable_pair_iff) +qed simp + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Gamma_Function.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,2969 @@ +(* Title: HOL/Analysis/Gamma.thy + Author: Manuel Eberl, TU München +*) + +section \<open>The Gamma Function\<close> + +theory Gamma_Function +imports + Complex_Transcendental + Summation_Tests + Harmonic_Numbers + "~~/src/HOL/Library/Nonpos_Ints" + "~~/src/HOL/Library/Periodic_Fun" +begin + +text \<open> + Several equivalent definitions of the Gamma function and its + most important properties. Also contains the definition and some properties + of the log-Gamma function and the Digamma function and the other Polygamma functions. + + Based on the Gamma function, we also prove the Weierstraß product form of the + sin function and, based on this, the solution of the Basel problem (the + sum over all @{term "1 / (n::nat)^2"}. +\<close> + +lemma pochhammer_eq_0_imp_nonpos_Int: + "pochhammer (x::'a::field_char_0) n = 0 \<Longrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0" + by (auto simp: pochhammer_eq_0_iff) + +lemma closed_nonpos_Ints [simp]: "closed (\<int>\<^sub>\<le>\<^sub>0 :: 'a :: real_normed_algebra_1 set)" +proof - + have "\<int>\<^sub>\<le>\<^sub>0 = (of_int ` {n. n \<le> 0} :: 'a set)" + by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int) + also have "closed \<dots>" by (rule closed_of_int_image) + finally show ?thesis . +qed + +lemma plus_one_in_nonpos_Ints_imp: "z + 1 \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0" + using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all + +lemma of_int_in_nonpos_Ints_iff: + "(of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> 0" + by (auto simp: nonpos_Ints_def) + +lemma one_plus_of_int_in_nonpos_Ints_iff: + "(1 + of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> -1" +proof - + have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp + also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n + 1 \<le> 0" by (subst of_int_in_nonpos_Ints_iff) simp_all + also have "\<dots> \<longleftrightarrow> n \<le> -1" by presburger + finally show ?thesis . +qed + +lemma one_minus_of_nat_in_nonpos_Ints_iff: + "(1 - of_nat n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0" +proof - + have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp + also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger + finally show ?thesis . +qed + +lemma fraction_not_in_ints: + assumes "\<not>(n dvd m)" "n \<noteq> 0" + shows "of_int m / of_int n \<notin> (\<int> :: 'a :: {division_ring,ring_char_0} set)" +proof + assume "of_int m / (of_int n :: 'a) \<in> \<int>" + then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases) + with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: divide_simps) + hence "m = k * n" by (subst (asm) of_int_eq_iff) + hence "n dvd m" by simp + with assms(1) show False by contradiction +qed + +lemma fraction_not_in_nats: + assumes "\<not>n dvd m" "n \<noteq> 0" + shows "of_int m / of_int n \<notin> (\<nat> :: 'a :: {division_ring,ring_char_0} set)" +proof + assume "of_int m / of_int n \<in> (\<nat> :: 'a set)" + also note Nats_subset_Ints + finally have "of_int m / of_int n \<in> (\<int> :: 'a set)" . + moreover have "of_int m / of_int n \<notin> (\<int> :: 'a set)" + using assms by (intro fraction_not_in_ints) + ultimately show False by contradiction +qed + +lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \<notin> \<int> \<Longrightarrow> z \<notin> \<int>\<^sub>\<le>\<^sub>0" + by (auto simp: Ints_def nonpos_Ints_def) + +lemma double_in_nonpos_Ints_imp: + assumes "2 * (z :: 'a :: field_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0" + shows "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<or> z + 1/2 \<in> \<int>\<^sub>\<le>\<^sub>0" +proof- + from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases') + thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps) +qed + + +lemma sin_series: "(\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" +proof - + from sin_converges[of z] have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z" . + also have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z \<longleftrightarrow> + (\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" + by (subst sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric]) + (auto simp: sin_coeff_def subseq_def ac_simps elim!: oddE) + finally show ?thesis . +qed + +lemma cos_series: "(\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" +proof - + from cos_converges[of z] have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z" . + also have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z \<longleftrightarrow> + (\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" + by (subst sums_mono_reindex[of "\<lambda>n. 2*n", symmetric]) + (auto simp: cos_coeff_def subseq_def ac_simps elim!: evenE) + finally show ?thesis . +qed + +lemma sin_z_over_z_series: + fixes z :: "'a :: {real_normed_field,banach}" + assumes "z \<noteq> 0" + shows "(\<lambda>n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)" +proof - + from sin_series[of z] have "(\<lambda>n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z" + by (simp add: field_simps scaleR_conv_of_real) + from sums_mult[OF this, of "inverse z"] and assms show ?thesis + by (simp add: field_simps) +qed + +lemma sin_z_over_z_series': + fixes z :: "'a :: {real_normed_field,banach}" + assumes "z \<noteq> 0" + shows "(\<lambda>n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)" +proof - + from sums_split_initial_segment[OF sin_converges[of z], of 1] + have "(\<lambda>n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp + from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps) +qed + +lemma has_field_derivative_sin_z_over_z: + fixes A :: "'a :: {real_normed_field,banach} set" + shows "((\<lambda>z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)" + (is "(?f has_field_derivative ?f') _") +proof (rule has_field_derivative_at_within) + have "((\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n) + has_field_derivative (\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)" + proof (rule termdiffs_strong) + from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1] + show "summable (\<lambda>n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def) + qed simp + also have "(\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n) = ?f" + proof + fix z + show "(\<Sum>n. of_real (sin_coeff (n+1)) * z^n) = ?f z" + by (cases "z = 0") (insert sin_z_over_z_series'[of z], + simp_all add: scaleR_conv_of_real sums_iff powser_zero sin_coeff_def) + qed + also have "(\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) = + diffs (\<lambda>n. of_real (sin_coeff (Suc n))) 0" by (simp add: powser_zero) + also have "\<dots> = 0" by (simp add: sin_coeff_def diffs_def) + finally show "((\<lambda>z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" . +qed + +lemma round_Re_minimises_norm: + "norm ((z::complex) - of_int m) \<ge> norm (z - of_int (round (Re z)))" +proof - + let ?n = "round (Re z)" + have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)" + by (simp add: cmod_def) + also have "\<bar>Re z - of_int ?n\<bar> \<le> \<bar>Re z - of_int m\<bar>" by (rule round_diff_minimal) + hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \<le> sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)" + by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff) + also have "\<dots> = norm (z - of_int m)" by (simp add: cmod_def) + finally show ?thesis . +qed + +lemma Re_pos_in_ball: + assumes "Re z > 0" "t \<in> ball z (Re z/2)" + shows "Re t > 0" +proof - + have "Re (z - t) \<le> norm (z - t)" by (rule complex_Re_le_cmod) + also from assms have "\<dots> < Re z / 2" by (simp add: dist_complex_def) + finally show "Re t > 0" using assms by simp +qed + +lemma no_nonpos_Int_in_ball_complex: + assumes "Re z > 0" "t \<in> ball z (Re z/2)" + shows "t \<notin> \<int>\<^sub>\<le>\<^sub>0" + using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases) + +lemma no_nonpos_Int_in_ball: + assumes "t \<in> ball z (dist z (round (Re z)))" + shows "t \<notin> \<int>\<^sub>\<le>\<^sub>0" +proof + assume "t \<in> \<int>\<^sub>\<le>\<^sub>0" + then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases) + have "dist z (of_int n) \<le> dist z t + dist t (of_int n)" by (rule dist_triangle) + also from assms have "dist z t < dist z (round (Re z))" by simp + also have "\<dots> \<le> dist z (of_int n)" + using round_Re_minimises_norm[of z] by (simp add: dist_complex_def) + finally have "dist t (of_int n) > 0" by simp + with \<open>t = of_int n\<close> show False by simp +qed + +lemma no_nonpos_Int_in_ball': + assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \<notin> \<int>\<^sub>\<le>\<^sub>0" + obtains d where "d > 0" "\<And>t. t \<in> ball z d \<Longrightarrow> t \<notin> \<int>\<^sub>\<le>\<^sub>0" +proof (rule that) + from assms show "setdist {z} \<int>\<^sub>\<le>\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto +next + fix t assume "t \<in> ball z (setdist {z} \<int>\<^sub>\<le>\<^sub>0)" + thus "t \<notin> \<int>\<^sub>\<le>\<^sub>0" using setdist_le_dist[of z "{z}" t "\<int>\<^sub>\<le>\<^sub>0"] by force +qed + +lemma no_nonpos_Real_in_ball: + assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0" and t: "t \<in> ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" + shows "t \<notin> \<real>\<^sub>\<le>\<^sub>0" +using z +proof (cases "Im z = 0") + assume A: "Im z = 0" + with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff) + with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff) +next + assume A: "Im z \<noteq> 0" + have "abs (Im z) - abs (Im t) \<le> abs (Im z - Im t)" by linarith + also have "\<dots> = abs (Im (z - t))" by simp + also have "\<dots> \<le> norm (z - t)" by (rule abs_Im_le_cmod) + also from A t have "\<dots> \<le> abs (Im z) / 2" by (simp add: dist_complex_def) + finally have "abs (Im t) > 0" using A by simp + thus ?thesis by (force simp add: complex_nonpos_Reals_iff) +qed + + +subsection \<open>Definitions\<close> + +text \<open> + We define the Gamma function by first defining its multiplicative inverse @{term "Gamma_inv"}. + This is more convenient because @{term "Gamma_inv"} is entire, which makes proofs of its + properties more convenient because one does not have to watch out for discontinuities. + (e.g. @{term "Gamma_inv"} fulfils @{term "rGamma z = z * rGamma (z + 1)"} everywhere, + whereas @{term "Gamma"} does not fulfil the analogous equation on the non-positive integers) + + We define the Gamma function (resp. its inverse) in the Euler form. This form has the advantage + that it is a relatively simple limit that converges everywhere. The limit at the poles is 0 + (due to division by 0). The functional equation @{term "Gamma (z + 1) = z * Gamma z"} follows + immediately from the definition. +\<close> + +definition Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where + "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)" + +definition Gamma_series' :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where + "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n" + +definition rGamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where + "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))" + +lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)" + and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)" + unfolding Gamma_series_def rGamma_series_def by simp_all + +lemma rGamma_series_minus_of_nat: + "eventually (\<lambda>n. rGamma_series (- of_nat k) n = 0) sequentially" + using eventually_ge_at_top[of k] + by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff) + +lemma Gamma_series_minus_of_nat: + "eventually (\<lambda>n. Gamma_series (- of_nat k) n = 0) sequentially" + using eventually_ge_at_top[of k] + by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff) + +lemma Gamma_series'_minus_of_nat: + "eventually (\<lambda>n. Gamma_series' (- of_nat k) n = 0) sequentially" + using eventually_gt_at_top[of k] + by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff) + +lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma_series z \<longlonglongrightarrow> 0" + by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp) + +lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series z \<longlonglongrightarrow> 0" + by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp) + +lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series' z \<longlonglongrightarrow> 0" + by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp) + +lemma Gamma_series_Gamma_series': + assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(\<lambda>n. Gamma_series' z n / Gamma_series z n) \<longlonglongrightarrow> 1" +proof (rule Lim_transform_eventually) + from eventually_gt_at_top[of "0::nat"] + show "eventually (\<lambda>n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially" + proof eventually_elim + fix n :: nat assume n: "n > 0" + from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n" + by (cases n, simp) + (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec' + dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp) + also from n have "\<dots> = z / of_nat n + 1" by (simp add: divide_simps) + finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" .. + qed + have "(\<lambda>x. z / of_nat x) \<longlonglongrightarrow> 0" + by (rule tendsto_norm_zero_cancel) + (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n], + simp add: norm_divide inverse_eq_divide) + from tendsto_add[OF this tendsto_const[of 1]] show "(\<lambda>n. z / of_nat n + 1) \<longlonglongrightarrow> 1" by simp +qed + + +subsection \<open>Convergence of the Euler series form\<close> + +text \<open> + We now show that the series that defines the Gamma function in the Euler form converges + and that the function defined by it is continuous on the complex halfspace with positive + real part. + + We do this by showing that the logarithm of the Euler series is continuous and converges + locally uniformly, which means that the log-Gamma function defined by its limit is also + continuous. + + This will later allow us to lift holomorphicity and continuity from the log-Gamma + function to the inverse of the Gamma function, and from that to the Gamma function itself. +\<close> + +definition ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where + "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\<Sum>k=1..n. ln (z / of_nat k + 1))" + +definition ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where + "ln_Gamma_series' z n = + - euler_mascheroni*z - ln z + (\<Sum>k=1..n. z / of_nat n - ln (z / of_nat k + 1))" + +definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> 'a" where + "ln_Gamma z = lim (ln_Gamma_series z)" + + +text \<open> + We now show that the log-Gamma series converges locally uniformly for all complex numbers except + the non-positive integers. We do this by proving that the series is locally Cauchy, adapting this + proof: + http://math.stackexchange.com/questions/887158/convergence-of-gammaz-lim-n-to-infty-fracnzn-prod-m-0nzm +\<close> + +context +begin + +private lemma ln_Gamma_series_complex_converges_aux: + fixes z :: complex and k :: nat + assumes z: "z \<noteq> 0" and k: "of_nat k \<ge> 2*norm z" "k \<ge> 2" + shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \<le> 2*(norm z + norm z^2) / of_nat k^2" +proof - + let ?k = "of_nat k :: complex" and ?z = "norm z" + have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)" + by (simp add: algebra_simps) + also have "norm ... \<le> ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)" + by (subst norm_mult [symmetric], rule norm_triangle_ineq) + also have "norm (Ln (1 + -1/?k) - (-1/?k)) \<le> (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))" + using k by (intro Ln_approx_linear) (simp add: norm_divide) + hence "?z * norm (ln (1-1/?k) + 1/?k) \<le> ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))" + by (intro mult_left_mono) simp_all + also have "... \<le> (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k + by (simp add: field_simps power2_eq_square norm_divide) + also have "... \<le> (?z * 2) / of_nat k^2" using k + by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) + also have "norm (ln (1+z/?k) - z/?k) \<le> norm (z/?k)^2 / (1 - norm (z/?k))" using k + by (intro Ln_approx_linear) (simp add: norm_divide) + hence "norm (ln (1+z/?k) - z/?k) \<le> ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)" + by (simp add: field_simps norm_divide) + also have "... \<le> (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k + by (simp add: field_simps power2_eq_square) + also have "... \<le> (?z^2 * 2) / of_nat k^2" using k + by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) + also note add_divide_distrib [symmetric] + finally show ?thesis by (simp only: distrib_left mult.commute) +qed + +lemma ln_Gamma_series_complex_converges: + assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" + assumes d: "d > 0" "\<And>n. n \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> norm (z - of_int n) > d" + shows "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n :: complex)" +proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI') + fix e :: real assume e: "e > 0" + define e'' where "e'' = (SUP t:ball z d. norm t + norm t^2)" + define e' where "e' = e / (2*e'')" + have "bounded ((\<lambda>t. norm t + norm t^2) ` cball z d)" + by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) + hence "bounded ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto + hence bdd: "bdd_above ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above) + + with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def + by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z]) + have e'': "norm t + norm t^2 \<le> e''" if "t \<in> ball z d" for t unfolding e''_def using that + by (rule cSUP_upper[OF _ bdd]) + from e z e''_pos have e': "e' > 0" unfolding e'_def + by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps) + + have "summable (\<lambda>k. inverse ((real_of_nat k)^2))" + by (rule inverse_power_summable) simp + from summable_partial_sum_bound[OF this e'] guess M . note M = this + + define N where "N = max 2 (max (nat \<lceil>2 * (norm z + d)\<rceil>) M)" + { + from d have "\<lceil>2 * (cmod z + d)\<rceil> \<ge> \<lceil>0::real\<rceil>" + by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all + hence "2 * (norm z + d) \<le> of_nat (nat \<lceil>2 * (norm z + d)\<rceil>)" unfolding N_def + by (simp_all add: le_of_int_ceiling) + also have "... \<le> of_nat N" unfolding N_def + by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1) + finally have "of_nat N \<ge> 2 * (norm z + d)" . + moreover have "N \<ge> 2" "N \<ge> M" unfolding N_def by simp_all + moreover have "(\<Sum>k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \<ge> N" for m n + using M[OF order.trans[OF \<open>N \<ge> M\<close> that]] unfolding real_norm_def + by (subst (asm) abs_of_nonneg) (auto intro: setsum_nonneg simp: divide_simps) + moreover note calculation + } note N = this + + show "\<exists>M. \<forall>t\<in>ball z d. \<forall>m\<ge>M. \<forall>n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e" + unfolding dist_complex_def + proof (intro exI[of _ N] ballI allI impI) + fix t m n assume t: "t \<in> ball z d" and mn: "m \<ge> N" "n > m" + from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def) + also have "dist z 0 - dist z t \<le> dist 0 t" using dist_triangle[of 0 z t] + by (simp add: dist_commute) + finally have t_nz: "t \<noteq> 0" by auto + + have "norm t \<le> norm z + norm (t - z)" by (rule norm_triangle_sub) + also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute) + also have "2 * (norm z + d) \<le> of_nat N" by (rule N) + also have "N \<le> m" by (rule mn) + finally have norm_t: "2 * norm t < of_nat m" by simp + + have "ln_Gamma_series t m - ln_Gamma_series t n = + (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) + + ((\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)))" + by (simp add: ln_Gamma_series_def algebra_simps) + also have "(\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)) = + (\<Sum>k\<in>{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn + by (simp add: setsum_diff) + also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce + also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) = + (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn + by (subst setsum_telescope'' [symmetric]) simp_all + also have "... = (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N + by (intro setsum_cong_Suc) + (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat) + also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \<in> {Suc m..n}" for k + using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps) + hence "(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) = + (\<Sum>k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N + by (intro setsum.cong) simp_all + also note setsum.distrib [symmetric] + also have "norm (\<Sum>k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \<le> + (\<Sum>k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t + by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all + also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)" + by (simp add: setsum_right_distrib) + also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz + by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all + also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')" + by (simp add: e'_def field_simps power2_eq_square) + also from e''[OF t] e''_pos e + have "\<dots> \<le> e * 1" by (intro mult_left_mono) (simp_all add: field_simps) + finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp + qed +qed + +end + +lemma ln_Gamma_series_complex_converges': + assumes z: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "\<exists>d>0. uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)" +proof - + define d' where "d' = Re z" + define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)" + have "of_int (round d') \<in> \<int>\<^sub>\<le>\<^sub>0" if "d' \<le> 0" using that + by (intro nonpos_Ints_of_int) (simp_all add: round_def) + with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less) + + have "d < cmod (z - of_int n)" if "n \<in> \<int>\<^sub>\<le>\<^sub>0" for n + proof (cases "Re z > 0") + case True + from nonpos_Ints_nonpos[OF that] have n: "n \<le> 0" by simp + from True have "d = Re z/2" by (simp add: d_def d'_def) + also from n True have "\<dots> < Re (z - of_int n)" by simp + also have "\<dots> \<le> norm (z - of_int n)" by (rule complex_Re_le_cmod) + finally show ?thesis . + next + case False + with assms nonpos_Ints_of_int[of "round (Re z)"] + have "z \<noteq> of_int (round d')" by (auto simp: not_less) + with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def) + also have "\<dots> \<le> norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm) + finally show ?thesis . + qed + hence conv: "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)" + by (intro ln_Gamma_series_complex_converges d_pos z) simp_all + from d_pos conv show ?thesis by blast +qed + +lemma ln_Gamma_series_complex_converges'': "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> convergent (ln_Gamma_series z)" + by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent) + +lemma ln_Gamma_complex_LIMSEQ: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma_series z \<longlonglongrightarrow> ln_Gamma z" + using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def) + +lemma exp_ln_Gamma_series_complex: + assumes "n > 0" "z \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "exp (ln_Gamma_series z n :: complex) = Gamma_series z n" +proof - + from assms obtain m where m: "n = Suc m" by (cases n) blast + from assms have "z \<noteq> 0" by (intro notI) auto + with assms have "exp (ln_Gamma_series z n) = + (of_nat n) powr z / (z * (\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))))" + unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_setsum) + also from assms have "(\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))) = (\<Prod>k=1..n. z / of_nat k + 1)" + by (intro setprod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp) + also have "... = (\<Prod>k=1..n. z + k) / fact n" + by (simp add: fact_setprod) + (subst setprod_dividef [symmetric], simp_all add: field_simps) + also from m have "z * ... = (\<Prod>k=0..n. z + k) / fact n" + by (simp add: setprod.atLeast0_atMost_Suc_shift setprod.atLeast_Suc_atMost_Suc_shift) + also have "(\<Prod>k=0..n. z + k) = pochhammer z (Suc n)" + unfolding pochhammer_setprod + by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) + also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n" + unfolding Gamma_series_def using assms by (simp add: divide_simps powr_def Ln_of_nat) + finally show ?thesis . +qed + + +lemma ln_Gamma_series'_aux: + assumes "(z::complex) \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(\<lambda>k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums + (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s") +unfolding sums_def +proof (rule Lim_transform) + show "(\<lambda>n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \<longlonglongrightarrow> ?s" + (is "?g \<longlonglongrightarrow> _") + by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms) + + have A: "eventually (\<lambda>n. (\<Sum>k<n. ?f k) - ?g n = 0) sequentially" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + have "(\<Sum>k<n. ?f k) = (\<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k))" + by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric], + subst atLeastLessThanSuc_atLeastAtMost) simp_all + also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))" + by (simp add: harm_def setsum_subtractf setsum_right_distrib divide_inverse) + also from n have "\<dots> - ?g n = 0" + by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat) + finally show "(\<Sum>k<n. ?f k) - ?g n = 0" . + qed + show "(\<lambda>n. (\<Sum>k<n. ?f k) - ?g n) \<longlonglongrightarrow> 0" by (subst tendsto_cong[OF A]) simp_all +qed + + +lemma uniformly_summable_deriv_ln_Gamma: + assumes z: "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0" and d: "d > 0" "d \<le> norm z/2" + shows "uniformly_convergent_on (ball z d) + (\<lambda>k z. \<Sum>i<k. inverse (of_nat (Suc i)) - inverse (z + of_nat (Suc i)))" + (is "uniformly_convergent_on _ (\<lambda>k z. \<Sum>i<k. ?f i z)") +proof (rule weierstrass_m_test'_ev) + { + fix t assume t: "t \<in> ball z d" + have "norm z = norm (t + (z - t))" by simp + have "norm (t + (z - t)) \<le> norm t + norm (z - t)" by (rule norm_triangle_ineq) + also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm) + finally have A: "norm t > norm z / 2" using z by (simp add: field_simps) + + have "norm t = norm (z + (t - z))" by simp + also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq) + also from t d have "norm (t - z) \<le> norm z / 2" by (simp add: dist_norm norm_minus_commute) + also from z have "\<dots> < norm z" by simp + finally have B: "norm t < 2 * norm z" by simp + note A B + } note ball = this + + show "eventually (\<lambda>n. \<forall>t\<in>ball z d. norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially" + using eventually_gt_at_top apply eventually_elim + proof safe + fix t :: 'a assume t: "t \<in> ball z d" + from z ball[OF t] have t_nz: "t \<noteq> 0" by auto + fix n :: nat assume n: "n > nat \<lceil>4 * norm z\<rceil>" + from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp + also from n have "\<dots> < of_nat n" by linarith + finally have n: "of_nat n > 2 * norm t" . + hence "of_nat n > norm t" by simp + hence t': "t \<noteq> -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc) + + with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))" + by (simp add: divide_simps eq_neg_iff_add_eq_0 del: of_nat_Suc) + also have "norm \<dots> = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))" + by (simp add: norm_divide norm_mult divide_simps add_ac del: of_nat_Suc) + also { + from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \<le> of_nat (Suc n) / (2 * norm t)" + by (intro divide_left_mono mult_pos_pos) simp_all + also have "\<dots> < norm (of_nat (Suc n) / t) - norm (1 :: 'a)" + using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc) + also have "\<dots> \<le> norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq) + finally have "inverse (norm (of_nat (Suc n)/t + 1)) \<le> 4 * norm z / of_nat (Suc n)" + using z by (simp add: divide_simps norm_divide mult_ac del: of_nat_Suc) + } + also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) = + 4 * norm z * inverse (of_nat (Suc n)^2)" + by (simp add: divide_simps power2_eq_square del: of_nat_Suc) + finally show "norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)" + by (simp del: of_nat_Suc) + qed +next + show "summable (\<lambda>n. 4 * norm z * inverse ((of_nat (Suc n))^2))" + by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable) +qed + +lemma summable_deriv_ln_Gamma: + "z \<noteq> (0 :: 'a :: {real_normed_field,banach}) \<Longrightarrow> + summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))" + unfolding summable_iff_convergent + by (rule uniformly_convergent_imp_convergent, + rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all + + +definition Polygamma :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where + "Polygamma n z = (if n = 0 then + (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else + (-1)^Suc n * fact n * (\<Sum>k. inverse ((z + of_nat k)^Suc n)))" + +abbreviation Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where + "Digamma \<equiv> Polygamma 0" + +lemma Digamma_def: + "Digamma z = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni" + by (simp add: Polygamma_def) + + +lemma summable_Digamma: + assumes "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0" + shows "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" +proof - + have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums + (0 - inverse (z + of_nat 0))" + by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] + tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) + from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]] + show "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp +qed + +lemma summable_offset: + assumes "summable (\<lambda>n. f (n + k) :: 'a :: real_normed_vector)" + shows "summable f" +proof - + from assms have "convergent (\<lambda>m. \<Sum>n<m. f (n + k))" by (simp add: summable_iff_convergent) + hence "convergent (\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)))" + by (intro convergent_add convergent_const) + also have "(\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k))) = (\<lambda>m. \<Sum>n<m+k. f n)" + proof + fix m :: nat + have "{..<m+k} = {..<k} \<union> {k..<m+k}" by auto + also have "(\<Sum>n\<in>\<dots>. f n) = (\<Sum>n<k. f n) + (\<Sum>n=k..<m+k. f n)" + by (rule setsum.union_disjoint) auto + also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))" + by (intro setsum.reindex_cong[of "\<lambda>n. n + k"]) + (simp, subst image_add_atLeastLessThan, auto) + finally show "(\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)) = (\<Sum>n<m+k. f n)" by (simp add: atLeast0LessThan) + qed + finally have "(\<lambda>a. setsum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. setsum f {..<m + k})" + by (auto simp: convergent_LIMSEQ_iff dest: LIMSEQ_offset) + thus ?thesis by (auto simp: summable_iff_convergent convergent_def) +qed + +lemma Polygamma_converges: + fixes z :: "'a :: {real_normed_field,banach}" + assumes z: "z \<noteq> 0" and n: "n \<ge> 2" + shows "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)^n))" +proof (rule weierstrass_m_test'_ev) + define e where "e = (1 + d / norm z)" + define m where "m = nat \<lceil>norm z * e\<rceil>" + { + fix t assume t: "t \<in> ball z d" + have "norm t = norm (z + (t - z))" by simp + also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq) + also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute) + finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def) + } note ball = this + + show "eventually (\<lambda>k. \<forall>t\<in>ball z d. norm (inverse ((t + of_nat k)^n)) \<le> + inverse (of_nat (k - m)^n)) sequentially" + using eventually_gt_at_top[of m] apply eventually_elim + proof (intro ballI) + fix k :: nat and t :: 'a assume k: "k > m" and t: "t \<in> ball z d" + from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff) + also have "\<dots> \<le> norm (of_nat k :: 'a) - norm z * e" + unfolding m_def by (subst norm_of_nat) linarith + also from ball[OF t] have "\<dots> \<le> norm (of_nat k :: 'a) - norm t" by simp + also have "\<dots> \<le> norm (of_nat k + t)" by (rule norm_diff_ineq) + finally have "inverse ((norm (t + of_nat k))^n) \<le> inverse (real_of_nat (k - m)^n)" using k n + by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc) + thus "norm (inverse ((t + of_nat k)^n)) \<le> inverse (of_nat (k - m)^n)" + by (simp add: norm_inverse norm_power power_inverse) + qed + + have "summable (\<lambda>k. inverse ((real_of_nat k)^n))" + using inverse_power_summable[of n] n by simp + hence "summable (\<lambda>k. inverse ((real_of_nat (k + m - m))^n))" by simp + thus "summable (\<lambda>k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset) +qed + +lemma Polygamma_converges': + fixes z :: "'a :: {real_normed_field,banach}" + assumes z: "z \<noteq> 0" and n: "n \<ge> 2" + shows "summable (\<lambda>k. inverse ((z + of_nat k)^n))" + using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z] + by (simp add: summable_iff_convergent) + +lemma has_field_derivative_ln_Gamma_complex [derivative_intros]: + fixes z :: complex + assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0" + shows "(ln_Gamma has_field_derivative Digamma z) (at z)" +proof - + have not_nonpos_Int [simp]: "t \<notin> \<int>\<^sub>\<le>\<^sub>0" if "Re t > 0" for t + using that by (auto elim!: nonpos_Ints_cases') + from z have z': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" and z'': "z \<noteq> 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I + by blast+ + let ?f' = "\<lambda>z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))" + let ?f = "\<lambda>z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\<lambda>z. \<Sum>n. ?f' z n" + define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" + from z have d: "d > 0" "norm z/2 \<ge> d" by (auto simp add: complex_nonpos_Reals_iff d_def) + have ball: "Im t = 0 \<longrightarrow> Re t > 0" if "dist z t < d" for t + using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff) + have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums + (0 - inverse (z + of_nat 0))" + by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] + tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) + + have "((\<lambda>z. \<Sum>n. ?f z n) has_field_derivative ?F' z) (at z)" + using d z ln_Gamma_series'_aux[OF z'] + apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma) + apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball + simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff + simp del: of_nat_Suc) + apply (auto simp add: complex_nonpos_Reals_iff) + done + with z have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative + ?F' z - euler_mascheroni - inverse z) (at z)" + by (force intro!: derivative_eq_intros simp: Digamma_def) + also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp + also from sums have "-inverse z = (\<Sum>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))" + by (simp add: sums_iff) + also from sums summable_deriv_ln_Gamma[OF z''] + have "?F' z + \<dots> = (\<Sum>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" + by (subst suminf_add) (simp_all add: add_ac sums_iff) + also have "\<dots> - euler_mascheroni = Digamma z" by (simp add: Digamma_def) + finally have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) + has_field_derivative Digamma z) (at z)" . + moreover from eventually_nhds_ball[OF d(1), of z] + have "eventually (\<lambda>z. ln_Gamma z = (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)" + proof eventually_elim + fix t assume "t \<in> ball z d" + hence "t \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases) + from ln_Gamma_series'_aux[OF this] + show "ln_Gamma t = (\<Sum>k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff) + qed + ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) +qed + +declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros] + + +lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni" + by (simp add: Digamma_def) + +lemma Digamma_plus1: + assumes "z \<noteq> 0" + shows "Digamma (z+1) = Digamma z + 1/z" +proof - + have sums: "(\<lambda>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) + sums (inverse (z + of_nat 0) - 0)" + by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]] + tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) + have "Digamma (z+1) = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) - + euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac) + also have "suminf ?f = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + + (\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))" + using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff) + also have "(\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z" + using sums by (simp add: sums_iff inverse_eq_divide) + finally show ?thesis by (simp add: Digamma_def[of z]) +qed + +lemma Polygamma_plus1: + assumes "z \<noteq> 0" + shows "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" +proof (cases "n = 0") + assume n: "n \<noteq> 0" + let ?f = "\<lambda>k. inverse ((z + of_nat k) ^ Suc n)" + have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\<Sum>k. ?f (k+1))" + using n by (simp add: Polygamma_def add_ac) + also have "(\<Sum>k. ?f (k+1)) + (\<Sum>k<1. ?f k) = (\<Sum>k. ?f k)" + using Polygamma_converges'[OF assms, of "Suc n"] n + by (subst suminf_split_initial_segment [symmetric]) simp_all + hence "(\<Sum>k. ?f (k+1)) = (\<Sum>k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps) + also have "(-1) ^ Suc n * fact n * ((\<Sum>k. ?f k) - inverse (z ^ Suc n)) = + Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n + by (simp add: inverse_eq_divide algebra_simps Polygamma_def) + finally show ?thesis . +qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide) + +lemma Digamma_of_nat: + "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni" +proof (induction n) + case (Suc n) + have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp + also have "\<dots> = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))" + by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc) + also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc) + also have "\<dots> + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni" + by (simp add: harm_Suc) + finally show ?case . +qed (simp add: harm_def) + +lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni" + by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl) + +lemma Polygamma_of_real: "x \<noteq> 0 \<Longrightarrow> Polygamma n (of_real x) = of_real (Polygamma n x)" + unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"] + by (simp_all add: suminf_of_real) + +lemma Polygamma_Real: "z \<in> \<real> \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Polygamma n z \<in> \<real>" + by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all + +lemma Digamma_half_integer: + "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) = + (\<Sum>k<n. 2 / (of_nat (2*k+1))) - euler_mascheroni - of_real (2 * ln 2)" +proof (induction n) + case 0 + have "Digamma (1/2 :: 'a) = of_real (Digamma (1/2))" by (simp add: Polygamma_of_real [symmetric]) + also have "Digamma (1/2::real) = + (\<Sum>k. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni" + by (simp add: Digamma_def add_ac) + also have "(\<Sum>k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) = + (\<Sum>k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))" + by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide) + also have "\<dots> = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums'] + by (subst suminf_mult) (simp_all add: algebra_simps sums_iff) + finally show ?case by simp +next + case (Suc n) + have nz: "2 * of_nat n + (1:: 'a) \<noteq> 0" + using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac) + hence nz': "of_nat n + (1/2::'a) \<noteq> 0" by (simp add: field_simps) + have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp + also from nz' have "\<dots> = Digamma (of_nat n + 1 / 2) + 1 / (of_nat n + 1 / 2)" + by (rule Digamma_plus1) + also from nz nz' have "1 / (of_nat n + 1 / 2 :: 'a) = 2 / (2 * of_nat n + 1)" + by (subst divide_eq_eq) simp_all + also note Suc + finally show ?case by (simp add: add_ac) +qed + +lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)" + using Digamma_half_integer[of 0] by simp + +lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0" +proof - + have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp + also have "\<dots> = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp + also note euler_mascheroni_less_13_over_22 + also note ln2_le_25_over_36 + finally show ?thesis by simp +qed + + +lemma has_field_derivative_Polygamma [derivative_intros]: + fixes z :: "'a :: {real_normed_field,euclidean_space}" + assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)" +proof (rule has_field_derivative_at_within, cases "n = 0") + assume n: "n = 0" + let ?f = "\<lambda>k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)" + let ?F = "\<lambda>z. \<Sum>k. ?f k z" and ?f' = "\<lambda>k z. inverse ((z + of_nat k)\<^sup>2)" + from no_nonpos_Int_in_ball'[OF z] guess d . note d = this + from z have summable: "summable (\<lambda>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))" + by (intro summable_Digamma) force + from z have conv: "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)\<^sup>2))" + by (intro Polygamma_converges) auto + with d have "summable (\<lambda>k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent + by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent ) + + have "(?F has_field_derivative (\<Sum>k. ?f' k z)) (at z)" + proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) + fix k :: nat and t :: 'a assume t: "t \<in> ball z d" + from t d(2)[of t] show "((\<lambda>z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)" + by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc + dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases) + qed (insert d(1) summable conv, (assumption|simp)+) + with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" + unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n + by (force simp: power2_eq_square intro!: derivative_eq_intros) +next + assume n: "n \<noteq> 0" + from z have z': "z \<noteq> 0" by auto + from no_nonpos_Int_in_ball'[OF z] guess d . note d = this + define n' where "n' = Suc n" + from n have n': "n' \<ge> 2" by (simp add: n'_def) + have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative + (\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)" + proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) + fix k :: nat and t :: 'a assume t: "t \<in> ball z d" + with d have t': "t \<notin> \<int>\<^sub>\<le>\<^sub>0" "t \<noteq> 0" by auto + show "((\<lambda>a. inverse ((a + of_nat k) ^ n')) has_field_derivative + - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t' + by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp) + next + have "uniformly_convergent_on (ball z d) + (\<lambda>k z. (- of_nat n' :: 'a) * (\<Sum>i<k. inverse ((z + of_nat i) ^ (n'+1))))" + using z' n by (intro uniformly_convergent_mult Polygamma_converges) (simp_all add: n'_def) + thus "uniformly_convergent_on (ball z d) + (\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))" + by (subst (asm) setsum_right_distrib) simp + qed (insert Polygamma_converges'[OF z' n'] d, simp_all) + also have "(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) = + (- of_nat n') * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))" + using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all + finally have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative + - of_nat n' * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" . + from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"] + show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" + unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps) +qed + +declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros] + +lemma isCont_Polygamma [continuous_intros]: + fixes f :: "_ \<Rightarrow> 'a :: {real_normed_field,euclidean_space}" + shows "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Polygamma n (f x)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Polygamma]]) + +lemma continuous_on_Polygamma: + "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A (Polygamma n :: _ \<Rightarrow> 'a :: {real_normed_field,euclidean_space})" + by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast + +lemma isCont_ln_Gamma_complex [continuous_intros]: + fixes f :: "'a::t2_space \<Rightarrow> complex" + shows "isCont f z \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>z. ln_Gamma (f z)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]]) + +lemma continuous_on_ln_Gamma_complex [continuous_intros]: + fixes A :: "complex set" + shows "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A ln_Gamma" + by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident]) + fastforce + +text \<open> + We define a type class that captures all the fundamental properties of the inverse of the Gamma function + and defines the Gamma function upon that. This allows us to instantiate the type class both for + the reals and for the complex numbers with a minimal amount of proof duplication. +\<close> + +class Gamma = real_normed_field + complete_space + + fixes rGamma :: "'a \<Rightarrow> 'a" + assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \<longleftrightarrow> (\<exists>n. z = - of_nat n)" + assumes differentiable_rGamma_aux1: + "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow> + let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + \<longlonglongrightarrow> d) - scaleR euler_mascheroni 1 + in filterlim (\<lambda>y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R + norm (y - z)) (nhds 0) (at z)" + assumes differentiable_rGamma_aux2: + "let z = - of_nat n + in filterlim (\<lambda>y. (rGamma y - rGamma z - (-1)^n * (setprod of_nat {1..n}) * (y - z)) /\<^sub>R + norm (y - z)) (nhds 0) (at z)" + assumes rGamma_series_aux: "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow> + let fact' = (\<lambda>n. setprod of_nat {1..n}); + exp = (\<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x^k /\<^sub>R fact k) \<longlonglongrightarrow> e); + pochhammer' = (\<lambda>a n. (\<Prod>n = 0..n. a + of_nat n)) + in filterlim (\<lambda>n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1)))) + (nhds (rGamma z)) sequentially" +begin +subclass banach .. +end + +definition "Gamma z = inverse (rGamma z)" + + +subsection \<open>Basic properties\<close> + +lemma Gamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z = 0" + and rGamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z = 0" + using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') + +lemma Gamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z \<noteq> 0" + and rGamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z \<noteq> 0" + using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') + +lemma Gamma_eq_zero_iff: "Gamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0" + and rGamma_eq_zero_iff: "rGamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0" + using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') + +lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)" + unfolding Gamma_def by simp + +lemma rGamma_series_LIMSEQ [tendsto_intros]: + "rGamma_series z \<longlonglongrightarrow> rGamma z" +proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") + case False + hence "z \<noteq> - of_nat n" for n by auto + from rGamma_series_aux[OF this] show ?thesis + by (simp add: rGamma_series_def[abs_def] fact_setprod pochhammer_Suc_setprod + exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost) +qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ) + +lemma Gamma_series_LIMSEQ [tendsto_intros]: + "Gamma_series z \<longlonglongrightarrow> Gamma z" +proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") + case False + hence "(\<lambda>n. inverse (rGamma_series z n)) \<longlonglongrightarrow> inverse (rGamma z)" + by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff) + also have "(\<lambda>n. inverse (rGamma_series z n)) = Gamma_series z" + by (simp add: rGamma_series_def Gamma_series_def[abs_def]) + finally show ?thesis by (simp add: Gamma_def) +qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ) + +lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)" + using Gamma_series_LIMSEQ[of z] by (simp add: limI) + +lemma rGamma_1 [simp]: "rGamma 1 = 1" +proof - + have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" + using eventually_gt_at_top[of "0::nat"] + by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact + divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) + have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) + moreover have "rGamma_series 1 \<longlonglongrightarrow> rGamma 1" by (rule tendsto_intros) + ultimately show ?thesis by (intro LIMSEQ_unique) +qed + +lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z" +proof - + let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1" + have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * + pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" + by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) + also from n have "\<dots> = ?f n * rGamma_series z n" + by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac) + finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. + qed + moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z" + by (intro tendsto_intros lim_inverse_n) + hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp + ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z" + by (rule Lim_transform_eventually) + moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)" + by (intro tendsto_intros) + ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast +qed + + +lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)" +proof (induction n arbitrary: z) + case (Suc n z) + have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH) + also note rGamma_plus1 [symmetric] + finally show ?case by (simp add: add_ac pochhammer_rec') +qed simp_all + +lemma Gamma_plus1: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma (z + 1) = z * Gamma z" + using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff) + +lemma pochhammer_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> pochhammer z n = Gamma (z + of_nat n) / Gamma z" + using pochhammer_rGamma[of z] + by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps) + +lemma Gamma_0 [simp]: "Gamma 0 = 0" + and rGamma_0 [simp]: "rGamma 0 = 0" + and Gamma_neg_1 [simp]: "Gamma (- 1) = 0" + and rGamma_neg_1 [simp]: "rGamma (- 1) = 0" + and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0" + and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0" + and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0" + and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0" + by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff) + +lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp + +lemma Gamma_fact: "Gamma (1 + of_nat n) = fact n" + by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff + of_nat_Suc [symmetric] del: of_nat_Suc) + +lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)" + by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, + subst of_nat_Suc, subst Gamma_fact) (rule refl) + +lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)" +proof (cases "n > 0") + case True + hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all + with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp +qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int) + +lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)" + by (simp add: Gamma_of_int rGamma_inverse_Gamma) + +lemma Gamma_seriesI: + assumes "(\<lambda>n. g n / Gamma_series z n) \<longlonglongrightarrow> 1" + shows "g \<longlonglongrightarrow> Gamma z" +proof (rule Lim_transform_eventually) + have "1/2 > (0::real)" by simp + from tendstoD[OF assms, OF this] + show "eventually (\<lambda>n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially" + by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) + from assms have "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> 1 * Gamma z" + by (intro tendsto_intros) + thus "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> Gamma z" by simp +qed + +lemma Gamma_seriesI': + assumes "f \<longlonglongrightarrow> rGamma z" + assumes "(\<lambda>n. g n * f n) \<longlonglongrightarrow> 1" + assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "g \<longlonglongrightarrow> Gamma z" +proof (rule Lim_transform_eventually) + have "1/2 > (0::real)" by simp + from tendstoD[OF assms(2), OF this] show "eventually (\<lambda>n. g n * f n / f n = g n) sequentially" + by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) + from assms have "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> 1 / rGamma z" + by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff) + thus "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> Gamma z" by (simp add: Gamma_def divide_inverse) +qed + +lemma Gamma_series'_LIMSEQ: "Gamma_series' z \<longlonglongrightarrow> Gamma z" + by (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series'] + Gamma_series'_nonpos_Ints_LIMSEQ[of z]) + + +subsection \<open>Differentiability\<close> + +lemma has_field_derivative_rGamma_no_nonpos_int: + assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)" +proof (rule has_field_derivative_at_within) + from assms have "z \<noteq> - of_nat n" for n by auto + from differentiable_rGamma_aux1[OF this] + show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)" + unfolding Digamma_def suminf_def sums_def[abs_def] + has_field_derivative_def has_derivative_def netlimit_at + by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric]) +qed + +lemma has_field_derivative_rGamma_nonpos_int: + "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)" + apply (rule has_field_derivative_at_within) + using differentiable_rGamma_aux2[of n] + unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at + by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_setprod) simp + +lemma has_field_derivative_rGamma [derivative_intros]: + "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>norm z\<rfloor>) * fact (nat \<lfloor>norm z\<rfloor>) + else -rGamma z * Digamma z)) (at z within A)" +using has_field_derivative_rGamma_no_nonpos_int[of z A] + has_field_derivative_rGamma_nonpos_int[of "nat \<lfloor>norm z\<rfloor>" A] + by (auto elim!: nonpos_Ints_cases') + +declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros] +declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros] +declare has_field_derivative_rGamma_nonpos_int [derivative_intros] +declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros] +declare has_field_derivative_rGamma [derivative_intros] + + +lemma has_field_derivative_Gamma [derivative_intros]: + "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)" + unfolding Gamma_def [abs_def] + by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff) + +declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros] + +(* TODO: Hide ugly facts properly *) +hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2 + differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux + + + +(* TODO: differentiable etc. *) + + +subsection \<open>Continuity\<close> + +lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma" + by (rule DERIV_continuous_on has_field_derivative_rGamma)+ + +lemma continuous_on_Gamma [continuous_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A Gamma" + by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast + +lemma isCont_rGamma [continuous_intros]: + "isCont f z \<Longrightarrow> isCont (\<lambda>x. rGamma (f x)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_rGamma]]) + +lemma isCont_Gamma [continuous_intros]: + "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Gamma (f x)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]]) + + + +text \<open>The complex Gamma function\<close> + +instantiation complex :: Gamma +begin + +definition rGamma_complex :: "complex \<Rightarrow> complex" where + "rGamma_complex z = lim (rGamma_series z)" + +lemma rGamma_series_complex_converges: + "convergent (rGamma_series (z :: complex))" (is "?thesis1") + and rGamma_complex_altdef: + "rGamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2") +proof - + have "?thesis1 \<and> ?thesis2" + proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") + case False + have "rGamma_series z \<longlonglongrightarrow> exp (- ln_Gamma z)" + proof (rule Lim_transform_eventually) + from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE) + from this(1) uniformly_convergent_imp_convergent[OF this(2), of z] + have "ln_Gamma_series z \<longlonglongrightarrow> lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff) + thus "(\<lambda>n. exp (-ln_Gamma_series z n)) \<longlonglongrightarrow> exp (- ln_Gamma z)" + unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus) + from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False + show "eventually (\<lambda>n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially" + by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def) + qed + with False show ?thesis + by (auto simp: convergent_def rGamma_complex_def intro!: limI) + next + case True + then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases') + also have "rGamma_series \<dots> \<longlonglongrightarrow> 0" + by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const) + finally show ?thesis using True + by (auto simp: rGamma_complex_def convergent_def intro!: limI) + qed + thus "?thesis1" "?thesis2" by blast+ +qed + +context +begin + +(* TODO: duplication *) +private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)" +proof - + let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1" + have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * + pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" + by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) + also from n have "\<dots> = ?f n * rGamma_series z n" + by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac) + finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. + qed + moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z" + using rGamma_series_complex_converges + by (intro tendsto_intros lim_inverse_n) + (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def) + hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp + ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z" + by (rule Lim_transform_eventually) + moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)" + using rGamma_series_complex_converges + by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff) + ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast +qed + +private lemma has_field_derivative_rGamma_complex_no_nonpos_Int: + assumes "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" +proof - + have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z + proof (subst DERIV_cong_ev[OF refl _ refl]) + from that have "eventually (\<lambda>t. t \<in> ball z (Re z/2)) (nhds z)" + by (intro eventually_nhds_in_nhd) simp_all + thus "eventually (\<lambda>t. rGamma t = exp (- ln_Gamma t)) (nhds z)" + using no_nonpos_Int_in_ball_complex[OF that] + by (auto elim!: eventually_mono simp: rGamma_complex_altdef) + next + have "z \<notin> \<real>\<^sub>\<le>\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff) + with that show "((\<lambda>t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)" + by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef) + qed + + from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" + proof (induction "nat \<lfloor>1 - Re z\<rfloor>" arbitrary: z) + case (Suc n z) + from Suc.prems have z: "z \<noteq> 0" by auto + from Suc.hyps have "n = nat \<lfloor>- Re z\<rfloor>" by linarith + hence A: "n = nat \<lfloor>1 - Re (z + 1)\<rfloor>" by simp + from Suc.prems have B: "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp) + + have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1)) z) has_field_derivative + -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)" + by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps) + also have "(\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) = rGamma" + by (simp add: rGamma_complex_plus1) + also from z have "Digamma (z + 1) * z - 1 = z * Digamma z" + by (subst Digamma_plus1) (simp_all add: field_simps) + also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z" + by (simp add: rGamma_complex_plus1[of z, symmetric]) + finally show ?case . + qed (intro diff, simp) +qed + +private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1" +proof - + have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" + using eventually_gt_at_top[of "0::nat"] + by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact + divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) + have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) + thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI) +qed + +private lemma has_field_derivative_rGamma_complex_nonpos_Int: + "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))" +proof (induction n) + case 0 + have A: "(0::complex) + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by simp + have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)" + by (rule derivative_eq_intros DERIV_chain refl + has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1) + thus ?case by (simp add: rGamma_complex_plus1) +next + case (Suc n) + hence A: "(rGamma has_field_derivative (-1)^n * fact n) + (at (- of_nat (Suc n) + 1 :: complex))" by simp + have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative + (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))" + by (rule derivative_eq_intros refl A DERIV_chain)+ + (simp add: algebra_simps rGamma_complex_altdef) + thus ?case by (simp add: rGamma_complex_plus1) +qed + +instance proof + fix z :: complex show "(rGamma z = 0) \<longleftrightarrow> (\<exists>n. z = - of_nat n)" + by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases') +next + fix z :: complex assume "\<And>n. z \<noteq> - of_nat n" + hence "z \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases') + from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this] + show "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in (\<lambda>y. (rGamma y - rGamma z + + rGamma z * d * (y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0" + by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] + netlimit_at of_real_def[symmetric] suminf_def) +next + fix n :: nat + from has_field_derivative_rGamma_complex_nonpos_Int[of n] + show "let z = - of_nat n in (\<lambda>y. (rGamma y - rGamma z - (- 1) ^ n * setprod of_nat {1..n} * + (y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0" + by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def) +next + fix z :: complex + from rGamma_series_complex_converges[of z] have "rGamma_series z \<longlonglongrightarrow> rGamma z" + by (simp add: convergent_LIMSEQ_iff rGamma_complex_def) + thus "let fact' = \<lambda>n. setprod of_nat {1..n}; + exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e; + pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n + in (\<lambda>n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma z" + by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def + of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) +qed + +end +end + + +lemma Gamma_complex_altdef: + "Gamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))" + unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus) + +lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)" +proof - + have "rGamma_series (cnj z) = (\<lambda>n. cnj (rGamma_series z n))" + by (intro ext) (simp_all add: rGamma_series_def exp_cnj) + also have "... \<longlonglongrightarrow> cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros) + finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI]) +qed + +lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)" + unfolding Gamma_def by (simp add: cnj_rGamma) + +lemma Gamma_complex_real: + "z \<in> \<real> \<Longrightarrow> Gamma z \<in> (\<real> :: complex set)" and rGamma_complex_real: "z \<in> \<real> \<Longrightarrow> rGamma z \<in> \<real>" + by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma) + +lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)" + using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast + +lemma holomorphic_on_rGamma: "rGamma holomorphic_on A" + unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma) + +lemma analytic_on_rGamma: "rGamma analytic_on A" + unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_on_rGamma) + + +lemma field_differentiable_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma field_differentiable (at z within A)" + using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto + +lemma holomorphic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma holomorphic_on A" + unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma) + +lemma analytic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma analytic_on A" + by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open) + (auto intro!: holomorphic_on_Gamma) + +lemma has_field_derivative_rGamma_complex' [derivative_intros]: + "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-Re z\<rfloor>) * fact (nat \<lfloor>-Re z\<rfloor>) else + -rGamma z * Digamma z)) (at z within A)" + using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases') + +declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros] + + +lemma field_differentiable_Polygamma: + fixes z::complex + shows + "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Polygamma n field_differentiable (at z within A)" + using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto + +lemma holomorphic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n holomorphic_on A" + unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma) + +lemma analytic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n analytic_on A" + by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open) + (auto intro!: holomorphic_on_Polygamma) + + + +text \<open>The real Gamma function\<close> + +lemma rGamma_series_real: + "eventually (\<lambda>n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially" + using eventually_gt_at_top[of "0 :: nat"] +proof eventually_elim + fix n :: nat assume n: "n > 0" + have "Re (rGamma_series (of_real x) n) = + Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))" + using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real) + also from n have "\<dots> = Re (of_real ((pochhammer x (Suc n)) / + (fact n * (exp (x * ln (real_of_nat n))))))" + by (subst exp_of_real) simp + also from n have "\<dots> = rGamma_series x n" + by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def) + finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" .. +qed + +instantiation real :: Gamma +begin + +definition "rGamma_real x = Re (rGamma (of_real x :: complex))" + +instance proof + fix x :: real + have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def) + also have "of_real \<dots> = rGamma (of_real x :: complex)" + by (intro of_real_Re rGamma_complex_real) simp_all + also have "\<dots> = 0 \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff) + also have "\<dots> \<longleftrightarrow> (\<exists>n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases') + finally show "(rGamma x) = 0 \<longleftrightarrow> (\<exists>n. x = - real_of_nat n)" by simp +next + fix x :: real assume "\<And>n. x \<noteq> - of_nat n" + hence x: "complex_of_real x \<notin> \<int>\<^sub>\<le>\<^sub>0" + by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases') + then have "x \<noteq> 0" by auto + with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)" + by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex + simp: Polygamma_of_real rGamma_real_def [abs_def]) + thus "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (x + of_nat k)) + \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in (\<lambda>y. (rGamma y - rGamma x + + rGamma x * d * (y - x)) /\<^sub>R norm (y - x)) \<midarrow>x\<rightarrow> 0" + by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] + netlimit_at of_real_def[symmetric] suminf_def) +next + fix n :: nat + have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))" + by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex + simp: Polygamma_of_real rGamma_real_def [abs_def]) + thus "let x = - of_nat n in (\<lambda>y. (rGamma y - rGamma x - (- 1) ^ n * setprod of_nat {1..n} * + (y - x)) /\<^sub>R norm (y - x)) \<midarrow>x::real\<rightarrow> 0" + by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def) +next + fix x :: real + have "rGamma_series x \<longlonglongrightarrow> rGamma x" + proof (rule Lim_transform_eventually) + show "(\<lambda>n. Re (rGamma_series (of_real x) n)) \<longlonglongrightarrow> rGamma x" unfolding rGamma_real_def + by (intro tendsto_intros) + qed (insert rGamma_series_real, simp add: eq_commute) + thus "let fact' = \<lambda>n. setprod of_nat {1..n}; + exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e; + pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n + in (\<lambda>n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma x" + by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def + of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) +qed + +end + + +lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)" + unfolding rGamma_real_def using rGamma_complex_real by simp + +lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)" + unfolding Gamma_def by (simp add: rGamma_complex_of_real) + +lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))" + by (rule sym, rule limI, rule tendsto_intros) + +lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))" + by (rule sym, rule limI, rule tendsto_intros) + +lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))" + using rGamma_complex_real[OF Reals_of_real[of x]] + by (elim Reals_cases) + (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real) + +lemma ln_Gamma_series_complex_of_real: + "x > 0 \<Longrightarrow> n > 0 \<Longrightarrow> ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)" +proof - + assume xn: "x > 0" "n > 0" + have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \<ge> 1" for k + using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps) + with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real) +qed + +lemma ln_Gamma_real_converges: + assumes "(x::real) > 0" + shows "convergent (ln_Gamma_series x)" +proof - + have "(\<lambda>n. ln_Gamma_series (complex_of_real x) n) \<longlonglongrightarrow> ln_Gamma (of_real x)" using assms + by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff) + moreover from eventually_gt_at_top[of "0::nat"] + have "eventually (\<lambda>n. complex_of_real (ln_Gamma_series x n) = + ln_Gamma_series (complex_of_real x) n) sequentially" + by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms) + ultimately have "(\<lambda>n. complex_of_real (ln_Gamma_series x n)) \<longlonglongrightarrow> ln_Gamma (of_real x)" + by (subst tendsto_cong) assumption+ + from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def) +qed + +lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \<Longrightarrow> ln_Gamma_series x \<longlonglongrightarrow> ln_Gamma x" + using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff) + +lemma ln_Gamma_complex_of_real: "x > 0 \<Longrightarrow> ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)" +proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually) + assume x: "x > 0" + show "eventually (\<lambda>n. of_real (ln_Gamma_series x n) = + ln_Gamma_series (complex_of_real x) n) sequentially" + using eventually_gt_at_top[of "0::nat"] + by eventually_elim (simp add: ln_Gamma_series_complex_of_real x) +qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def) + +lemma Gamma_real_pos_exp: "x > (0 :: real) \<Longrightarrow> Gamma x = exp (ln_Gamma x)" + by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff + ln_Gamma_complex_of_real exp_of_real) + +lemma ln_Gamma_real_pos: "x > 0 \<Longrightarrow> ln_Gamma x = ln (Gamma x :: real)" + unfolding Gamma_real_pos_exp by simp + +lemma Gamma_real_pos: "x > (0::real) \<Longrightarrow> Gamma x > 0" + by (simp add: Gamma_real_pos_exp) + +lemma has_field_derivative_ln_Gamma_real [derivative_intros]: + assumes x: "x > (0::real)" + shows "(ln_Gamma has_field_derivative Digamma x) (at x)" +proof (subst DERIV_cong_ev[OF refl _ refl]) + from assms show "((Re \<circ> ln_Gamma \<circ> complex_of_real) has_field_derivative Digamma x) (at x)" + by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex + simp: Polygamma_of_real o_def) + from eventually_nhds_in_nhd[of x "{0<..}"] assms + show "eventually (\<lambda>y. ln_Gamma y = (Re \<circ> ln_Gamma \<circ> of_real) y) (nhds x)" + by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open) +qed + +declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros] + + +lemma has_field_derivative_rGamma_real' [derivative_intros]: + "(rGamma has_field_derivative (if x \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-x\<rfloor>) * fact (nat \<lfloor>-x\<rfloor>) else + -rGamma x * Digamma x)) (at x within A)" + using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases') + +declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros] + +lemma Polygamma_real_odd_pos: + assumes "(x::real) \<notin> \<int>\<^sub>\<le>\<^sub>0" "odd n" + shows "Polygamma n x > 0" +proof - + from assms have "x \<noteq> 0" by auto + with assms show ?thesis + unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] + by (auto simp: zero_less_power_eq simp del: power_Suc + dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos) +qed + +lemma Polygamma_real_even_neg: + assumes "(x::real) > 0" "n > 0" "even n" + shows "Polygamma n x < 0" + using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] + by (auto intro!: mult_pos_pos suminf_pos) + +lemma Polygamma_real_strict_mono: + assumes "x > 0" "x < (y::real)" "even n" + shows "Polygamma n x < Polygamma n y" +proof - + have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>" + using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) + then guess \<xi> by (elim exE conjE) note \<xi> = this + note \<xi>(3) + also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> > 0" + by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases) + finally show ?thesis by simp +qed + +lemma Polygamma_real_strict_antimono: + assumes "x > 0" "x < (y::real)" "odd n" + shows "Polygamma n x > Polygamma n y" +proof - + have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>" + using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) + then guess \<xi> by (elim exE conjE) note \<xi> = this + note \<xi>(3) + also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> < 0" + by (intro mult_pos_neg Polygamma_real_even_neg) simp_all + finally show ?thesis by simp +qed + +lemma Polygamma_real_mono: + assumes "x > 0" "x \<le> (y::real)" "even n" + shows "Polygamma n x \<le> Polygamma n y" + using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2) + by (cases "x = y") simp_all + +lemma Digamma_real_ge_three_halves_pos: + assumes "x \<ge> 3/2" + shows "Digamma (x :: real) > 0" +proof - + have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos) + also from assms have "\<dots> \<le> Digamma x" by (intro Polygamma_real_mono) simp_all + finally show ?thesis . +qed + +lemma ln_Gamma_real_strict_mono: + assumes "x \<ge> 3/2" "x < y" + shows "ln_Gamma (x :: real) < ln_Gamma y" +proof - + have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> ln_Gamma y - ln_Gamma x = (y - x) * Digamma \<xi>" + using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) + then guess \<xi> by (elim exE conjE) note \<xi> = this + note \<xi>(3) + also from \<xi>(1,2) assms have "(y - x) * Digamma \<xi> > 0" + by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all + finally show ?thesis by simp +qed + +lemma Gamma_real_strict_mono: + assumes "x \<ge> 3/2" "x < y" + shows "Gamma (x :: real) < Gamma y" +proof - + from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp + also have "\<dots> < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms) + also from Gamma_real_pos_exp[of y] assms have "\<dots> = Gamma y" by simp + finally show ?thesis . +qed + +lemma log_convex_Gamma_real: "convex_on {0<..} (ln \<circ> Gamma :: real \<Rightarrow> real)" + by (rule convex_on_realI[of _ _ Digamma]) + (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos + simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases') + + +subsection \<open>Beta function\<close> + +definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)" + +lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)" + by (simp add: inverse_eq_divide Beta_def Gamma_def) + +lemma Beta_commute: "Beta a b = Beta b a" + unfolding Beta_def by (simp add: ac_simps) + +lemma has_field_derivative_Beta1 [derivative_intros]: + assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "((\<lambda>x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y)))) + (at x within A)" unfolding Beta_altdef + by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps) + +lemma Beta_pole1: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0" + by (auto simp add: Beta_def elim!: nonpos_Ints_cases') + +lemma Beta_pole2: "y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0" + by (auto simp add: Beta_def elim!: nonpos_Ints_cases') + +lemma Beta_zero: "x + y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0" + by (auto simp add: Beta_def elim!: nonpos_Ints_cases') + +lemma has_field_derivative_Beta2 [derivative_intros]: + assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "((\<lambda>y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y)))) + (at y within A)" + using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac) + +lemma Beta_plus1_plus1: + assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "Beta (x + 1) y + Beta x (y + 1) = Beta x y" +proof - + have "Beta (x + 1) y + Beta x (y + 1) = + (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)" + by (simp add: Beta_altdef add_divide_distrib algebra_simps) + also have "\<dots> = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))" + by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps) + also from assms have "\<dots> = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp + finally show ?thesis . +qed + +lemma Beta_plus1_left: + assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(x + y) * Beta (x + 1) y = x * Beta x y" +proof - + have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))" + unfolding Beta_altdef by (simp only: ac_simps) + also have "\<dots> = x * Beta x y" unfolding Beta_altdef + by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps) + finally show ?thesis . +qed + +lemma Beta_plus1_right: + assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(x + y) * Beta x (y + 1) = y * Beta x y" + using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute) + +lemma Gamma_Gamma_Beta: + assumes "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "Gamma x * Gamma y = Beta x y * Gamma (x + y)" + unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"] + by (simp add: rGamma_inverse_Gamma) + + + +subsection \<open>Legendre duplication theorem\<close> + +context +begin + +private lemma Gamma_legendre_duplication_aux: + fixes z :: "'a :: Gamma" + assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)" +proof - + let ?powr = "\<lambda>b a. exp (a * of_real (ln (of_nat b)))" + let ?h = "\<lambda>n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) * + exp (1/2 * of_real (ln (real_of_nat n)))" + { + fix z :: 'a assume z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" + let ?g = "\<lambda>n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n / + Gamma_series' (2*z) (2*n)" + have "eventually (\<lambda>n. ?g n = ?h n) sequentially" using eventually_gt_at_top + proof eventually_elim + fix n :: nat assume n: "n > 0" + let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a" + have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp + have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) / + (pochhammer z n * pochhammer (z + 1/2) n)" + by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac) + have B: "Gamma_series' (2*z) (2*n) = + ?f' * ?powr 2 (2*z) * ?powr n (2*z) / + (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n + by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double) + from z have "pochhammer z n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) + moreover from z have "pochhammer (z + 1/2) n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) + ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) = + ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))" + using n unfolding A B by (simp add: divide_simps exp_minus) + also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)" + by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib) + finally show "?g n = ?h n" by (simp only: mult_ac) + qed + + moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto + hence "?g \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" + using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "op*2" "2*z"] + by (intro tendsto_intros Gamma_series'_LIMSEQ) + (simp_all add: o_def subseq_def Gamma_eq_zero_iff) + ultimately have "?h \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" + by (rule Lim_transform_eventually) + } note lim = this + + from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto + from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \<notin> \<int>\<^sub>\<le>\<^sub>0" + by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all + with lim[of "1/2 :: 'a"] have "?h \<longlonglongrightarrow> 2 * Gamma (1 / 2 :: 'a)" by (simp add: exp_of_real) + from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis + by (simp add: divide_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps) +qed + +(* TODO: perhaps this is unnecessary once we have the fact that a holomorphic function is + infinitely differentiable *) +private lemma Gamma_reflection_aux: + defines "h \<equiv> \<lambda>z::complex. if z \<in> \<int> then 0 else + (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" + defines "a \<equiv> complex_of_real pi" + obtains h' where "continuous_on UNIV h'" "\<And>z. (h has_field_derivative (h' z)) (at z)" +proof - + define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n + define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z + define g where "g n = complex_of_real (sin_coeff (n+1))" for n + define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z + have a_nz: "a \<noteq> 0" unfolding a_def by simp + + have "(\<lambda>n. f n * (a*z)^n) sums (F z) \<and> (\<lambda>n. g n * (a*z)^n) sums (G z)" + if "abs (Re z) < 1" for z + proof (cases "z = 0"; rule conjI) + assume "z \<noteq> 0" + note z = this that + + from z have sin_nz: "sin (a*z) \<noteq> 0" unfolding a_def by (auto simp: sin_eq_0) + have "(\<lambda>n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"] + by (simp add: scaleR_conv_of_real) + from sums_split_initial_segment[OF this, of 1] + have "(\<lambda>n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac) + from sums_mult[OF this, of "inverse (a*z)"] z a_nz + have A: "(\<lambda>n. g n * (a*z)^n) sums (sin (a*z)/(a*z))" + by (simp add: field_simps g_def) + with z show "(\<lambda>n. g n * (a*z)^n) sums (G z)" by (simp add: G_def) + from A z a_nz sin_nz have g_nz: "(\<Sum>n. g n * (a*z)^n) \<noteq> 0" by (simp add: sums_iff g_def) + + have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def) + from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1] + have "(\<lambda>n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))" + by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def) + from sums_mult[OF this, of "inverse z"] z assms + show "(\<lambda>n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def) + next + assume z: "z = 0" + have "(\<lambda>n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp + with z show "(\<lambda>n. f n * (a * z) ^ n) sums (F z)" + by (simp add: f_def F_def sin_coeff_def cos_coeff_def) + have "(\<lambda>n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp + with z show "(\<lambda>n. g n * (a * z) ^ n) sums (G z)" + by (simp add: g_def G_def sin_coeff_def cos_coeff_def) + qed + note sums = conjunct1[OF this] conjunct2[OF this] + + define h2 where [abs_def]: + "h2 z = (\<Sum>n. f n * (a*z)^n) / (\<Sum>n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z + define POWSER where [abs_def]: "POWSER f z = (\<Sum>n. f n * (z^n :: complex))" for f z + define POWSER' where [abs_def]: "POWSER' f z = (\<Sum>n. diffs f n * (z^n))" for f and z :: complex + define h2' where [abs_def]: + "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / + (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z + + have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t + proof - + from that have t: "t \<in> \<int> \<longleftrightarrow> t = 0" by (auto elim!: Ints_cases simp: dist_0_norm) + hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)" + unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def) + also have "a*cot (a*t) - 1/t = (F t) / (G t)" + using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def) + also have "\<dots> = (\<Sum>n. f n * (a*t)^n) / (\<Sum>n. g n * (a*t)^n)" + using sums[of t] that by (simp add: sums_iff dist_0_norm) + finally show "h t = h2 t" by (simp only: h2_def) + qed + + let ?A = "{z. abs (Re z) < 1}" + have "open ({z. Re z < 1} \<inter> {z. Re z > -1})" + using open_halfspace_Re_gt open_halfspace_Re_lt by auto + also have "({z. Re z < 1} \<inter> {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto + finally have open_A: "open ?A" . + hence [simp]: "interior ?A = ?A" by (simp add: interior_open) + + have summable_f: "summable (\<lambda>n. f n * z^n)" for z + by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"]) + (simp_all add: norm_mult a_def del: of_real_add) + have summable_g: "summable (\<lambda>n. g n * z^n)" for z + by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"]) + (simp_all add: norm_mult a_def del: of_real_add) + have summable_fg': "summable (\<lambda>n. diffs f n * z^n)" "summable (\<lambda>n. diffs g n * z^n)" for z + by (intro termdiff_converges_all summable_f summable_g)+ + have "(POWSER f has_field_derivative (POWSER' f z)) (at z)" + "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z + unfolding POWSER_def POWSER'_def + by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+ + note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def] + have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z" + for z unfolding POWSER_def POWSER'_def + by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+ + note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def] + + { + fix z :: complex assume z: "abs (Re z) < 1" + define d where "d = \<i> * of_real (norm z + 1)" + have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add) + have "eventually (\<lambda>z. h z = h2 z) (nhds z)" + using eventually_nhds_in_nhd[of z ?A] using h_eq z + by (auto elim!: eventually_mono simp: dist_0_norm) + + moreover from sums(2)[OF z] z have nz: "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0" + unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def) + have A: "z \<in> \<int> \<longleftrightarrow> z = 0" using z by (auto elim!: Ints_cases) + have no_int: "1 + z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of "1+z" 1] A + by (auto elim!: nonpos_Ints_cases) + have no_int': "1 - z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of 1 "1-z"] A + by (auto elim!: nonpos_Ints_cases) + from no_int no_int' have no_int: "1 - z \<notin> \<int>\<^sub>\<le>\<^sub>0" "1 + z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto + have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def + by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+) + (auto simp: h2'_def POWSER_def field_simps power2_eq_square) + ultimately have deriv: "(h has_field_derivative h2' z) (at z)" + by (subst DERIV_cong_ev[OF refl _ refl]) + + from sums(2)[OF z] z have "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0" + unfolding G_def by (auto simp: sums_iff a_def sin_eq_0) + hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def + by (intro continuous_intros cont + continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto + note deriv and this + } note A = this + + interpret h: periodic_fun_simple' h + proof + fix z :: complex + show "h (z + 1) = h z" + proof (cases "z \<in> \<int>") + assume z: "z \<notin> \<int>" + hence A: "z + 1 \<notin> \<int>" "z \<noteq> 0" using Ints_diff[of "z+1" 1] by auto + hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)" + by (subst (1 2) Digamma_plus1) simp_all + with A z show "h (z + 1) = h z" + by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def) + qed (simp add: h_def) + qed + + have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z + proof - + have "((\<lambda>z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)" + by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) + (insert z, auto intro!: derivative_eq_intros) + hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1) + moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all + ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique) + qed + + define h2'' where "h2'' z = h2' (z - of_int \<lfloor>Re z\<rfloor>)" for z + have deriv: "(h has_field_derivative h2'' z) (at z)" for z + proof - + fix z :: complex + have B: "\<bar>Re z - real_of_int \<lfloor>Re z\<rfloor>\<bar> < 1" by linarith + have "((\<lambda>t. h (t - of_int \<lfloor>Re z\<rfloor>)) has_field_derivative h2'' z) (at z)" + unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) + (insert B, auto intro!: derivative_intros) + thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int) + qed + + have cont: "continuous_on UNIV h2''" + proof (intro continuous_at_imp_continuous_on ballI) + fix z :: complex + define r where "r = \<lfloor>Re z\<rfloor>" + define A where "A = {t. of_int r - 1 < Re t \<and> Re t < of_int r + 1}" + have "continuous_on A (\<lambda>t. h2' (t - of_int r))" unfolding A_def + by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros) + (simp_all add: abs_real_def) + moreover have "h2'' t = h2' (t - of_int r)" if t: "t \<in> A" for t + proof (cases "Re t \<ge> of_int r") + case True + from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) + with True have "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor>" unfolding r_def by linarith + thus ?thesis by (auto simp: r_def h2''_def) + next + case False + from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) + with False have t': "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor> - 1" unfolding r_def by linarith + moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)" + by (intro h2'_eq) simp_all + ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t') + qed + ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl]) + moreover { + have "open ({t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t})" + by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt) + also have "{t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t} = A" + unfolding A_def by blast + finally have "open A" . + } + ultimately have C: "isCont h2'' t" if "t \<in> A" for t using that + by (subst (asm) continuous_on_eq_continuous_at) auto + have "of_int r - 1 < Re z" "Re z < of_int r + 1" unfolding r_def by linarith+ + thus "isCont h2'' z" by (intro C) (simp_all add: A_def) + qed + + from that[OF cont deriv] show ?thesis . +qed + +lemma Gamma_reflection_complex: + fixes z :: complex + shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)" +proof - + let ?g = "\<lambda>z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)" + define g where [abs_def]: "g z = (if z \<in> \<int> then of_real pi else ?g z)" for z :: complex + let ?h = "\<lambda>z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" + define h where [abs_def]: "h z = (if z \<in> \<int> then 0 else ?h z)" for z :: complex + + \<comment> \<open>@{term g} is periodic with period 1.\<close> + interpret g: periodic_fun_simple' g + proof + fix z :: complex + show "g (z + 1) = g z" + proof (cases "z \<in> \<int>") + case False + hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def) + also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)" + using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints + by (subst Beta_plus1_left [symmetric]) auto + also have "\<dots> * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))" + using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints + by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi) + also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)" + using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def) + finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto + qed (simp add: g_def) + qed + + \<comment> \<open>@{term g} is entire.\<close> + have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex + proof (cases "z \<in> \<int>") + let ?h' = "\<lambda>z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) + + of_real pi * cos (z * of_real pi))" + case False + from False have "eventually (\<lambda>t. t \<in> UNIV - \<int>) (nhds z)" + by (intro eventually_nhds_in_open) (auto simp: open_Diff) + hence "eventually (\<lambda>t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def) + moreover { + from False Ints_diff[of 1 "1-z"] have "1 - z \<notin> \<int>" by auto + hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints + by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def) + also from False have "sin (of_real pi * z) \<noteq> 0" by (subst sin_eq_0) auto + hence "?h' z = h z * g z" + using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def) + finally have "(?g has_field_derivative (h z * g z)) (at z)" . + } + ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) + next + case True + then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases) + let ?t = "(\<lambda>z::complex. if z = 0 then 1 else sin z / z) \<circ> (\<lambda>z. of_real pi * z)" + have deriv_0: "(g has_field_derivative 0) (at 0)" + proof (subst DERIV_cong_ev[OF refl _ refl]) + show "eventually (\<lambda>z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)" + using eventually_nhds_ball[OF zero_less_one, of "0::complex"] + proof eventually_elim + fix z :: complex assume z: "z \<in> ball 0 1" + show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z" + proof (cases "z = 0") + assume z': "z \<noteq> 0" + with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases simp: dist_0_norm) + from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp + with z'' z' show ?thesis by (simp add: g_def ac_simps) + qed (simp add: g_def) + qed + have "(?t has_field_derivative (0 * of_real pi)) (at 0)" + using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"] + by (intro DERIV_chain) simp_all + thus "((\<lambda>z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)" + by (auto intro!: derivative_eq_intros simp: o_def) + qed + + have "((g \<circ> (\<lambda>x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))" + using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros) + also have "g \<circ> (\<lambda>x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int) + finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def) + qed + + have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z + proof (cases "z \<in> \<int>") + case True + with that have "z = 0 \<or> z = 1" by (force elim!: Ints_cases) + moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0" + using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square) + moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1" + using fraction_not_in_ints[where 'a = complex, of 2 1] + by (simp add: g_def power2_eq_square Beta_def algebra_simps) + ultimately show ?thesis by force + next + case False + hence z: "z/2 \<notin> \<int>" "(z+1)/2 \<notin> \<int>" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases) + hence z': "z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "(z+1)/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases) + from z have "1-z/2 \<notin> \<int>" "1-((z+1)/2) \<notin> \<int>" + using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto + hence z'': "1-z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "1-((z+1)/2) \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases) + from z have "g (z/2) * g ((z+1)/2) = + (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) * + (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))" + by (simp add: g_def) + also from z' Gamma_legendre_duplication_aux[of "z/2"] + have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z" + by (simp add: add_divide_distrib) + also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"] + have "Gamma (1-z/2) * Gamma (1-(z+1)/2) = + Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))" + by (simp add: add_divide_distrib ac_simps) + finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) * + (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))" + by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real) + also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)" + using cos_sin_eq[of "- of_real pi * z/2", symmetric] + by (simp add: ring_distribs add_divide_distrib ac_simps) + also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)" + by (subst sin_times_cos) (simp add: field_simps) + also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z" + using \<open>z \<notin> \<int>\<close> by (simp add: g_def) + finally show ?thesis . + qed + have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z + proof - + define r where "r = \<lfloor>Re z / 2\<rfloor>" + have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int) + also have "of_int (2*r) = 2 * of_int r" by simp + also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+ + hence "Gamma (1/2)^2 * g (z - 2 * of_int r) = + g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)" + unfolding r_def by (intro g_eq[symmetric]) simp_all + also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp + also have "g \<dots> = g (z/2)" by (rule g.minus_of_int) + also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp + also have "g \<dots> = g ((z+1)/2)" by (rule g.minus_of_int) + finally show ?thesis .. + qed + + have g_nz [simp]: "g z \<noteq> 0" for z :: complex + unfolding g_def using Ints_diff[of 1 "1 - z"] + by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int) + + have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z + proof - + have "((\<lambda>t. g (t/2) * g ((t+1)/2)) has_field_derivative + (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" + by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps) + hence "((\<lambda>t. Gamma (1/2)^2 * g t) has_field_derivative + Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" + by (subst (1 2) g_eq[symmetric]) simp + from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"] + have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)" + using fraction_not_in_ints[where 'a = complex, of 2 1] + by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints) + moreover have "(g has_field_derivative (g z * h z)) (at z)" + using g_g'[of z] by (simp add: ac_simps) + ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)" + by (intro DERIV_unique) + thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp + qed + + obtain h' where h'_cont: "continuous_on UNIV h'" and + h_h': "\<And>z. (h has_field_derivative h' z) (at z)" + unfolding h_def by (erule Gamma_reflection_aux) + + have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z + proof - + have "((\<lambda>t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative + ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)" + by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2]) + hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)" + by (subst (asm) h_eq[symmetric]) + from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique) + qed + + have h'_zero: "h' z = 0" for z + proof - + define m where "m = max 1 \<bar>Re z\<bar>" + define B where "B = {t. abs (Re t) \<le> m \<and> abs (Im t) \<le> abs (Im z)}" + have "closed ({t. Re t \<ge> -m} \<inter> {t. Re t \<le> m} \<inter> + {t. Im t \<ge> -\<bar>Im z\<bar>} \<inter> {t. Im t \<le> \<bar>Im z\<bar>})" + (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le + closed_halfspace_Im_ge closed_halfspace_Im_le) + also have "?B = B" unfolding B_def by fastforce + finally have "closed B" . + moreover have "bounded B" unfolding bounded_iff + proof (intro ballI exI) + fix t assume t: "t \<in> B" + have "norm t \<le> \<bar>Re t\<bar> + \<bar>Im t\<bar>" by (rule cmod_le) + also from t have "\<bar>Re t\<bar> \<le> m" unfolding B_def by blast + also from t have "\<bar>Im t\<bar> \<le> \<bar>Im z\<bar>" unfolding B_def by blast + finally show "norm t \<le> m + \<bar>Im z\<bar>" by - simp + qed + ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast + + define M where "M = (SUP z:B. norm (h' z))" + have "compact (h' ` B)" + by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+ + hence bdd: "bdd_above ((\<lambda>z. norm (h' z)) ` B)" + using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded) + have "norm (h' z) \<le> M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def) + also have "M \<le> M/2" + proof (subst M_def, subst cSUP_le_iff) + have "z \<in> B" unfolding B_def m_def by simp + thus "B \<noteq> {}" by auto + next + show "\<forall>z\<in>B. norm (h' z) \<le> M/2" + proof + fix t :: complex assume t: "t \<in> B" + from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm) + also have "norm \<dots> = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp + also have "norm (h' (t/2) + h' ((t+1)/2)) \<le> norm (h' (t/2)) + norm (h' ((t+1)/2))" + by (rule norm_triangle_ineq) + also from t have "abs (Re ((t + 1)/2)) \<le> m" unfolding m_def B_def by auto + with t have "t/2 \<in> B" "(t+1)/2 \<in> B" unfolding B_def by auto + hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \<le> M + M" unfolding M_def + by (intro add_mono cSUP_upper bdd) (auto simp: B_def) + also have "(M + M) / 4 = M / 2" by simp + finally show "norm (h' t) \<le> M/2" by - simp_all + qed + qed (insert bdd, auto simp: cball_eq_empty) + hence "M \<le> 0" by simp + finally show "h' z = 0" by simp + qed + have h_h'_2: "(h has_field_derivative 0) (at z)" for z + using h_h'[of z] h'_zero[of z] by simp + + have g_real: "g z \<in> \<real>" if "z \<in> \<real>" for z + unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real) + have h_real: "h z \<in> \<real>" if "z \<in> \<real>" for z + unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real) + have g_nz: "g z \<noteq> 0" for z unfolding g_def using Ints_diff[of 1 "1-z"] + by (auto simp: Gamma_eq_zero_iff sin_eq_0) + + from h'_zero h_h'_2 have "\<exists>c. \<forall>z\<in>UNIV. h z = c" + by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm) + then obtain c where c: "\<And>z. h z = c" by auto + have "\<exists>u. u \<in> closed_segment 0 1 \<and> Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))" + by (intro complex_mvt_line g_g') + find_theorems name:deriv Reals + then guess u by (elim exE conjE) note u = this + from u(1) have u': "u \<in> \<real>" unfolding closed_segment_def + by (auto simp: scaleR_conv_of_real) + from u' g_real[of u] g_nz[of u] have "Re (g u) \<noteq> 0" by (auto elim!: Reals_cases) + with u(2) c[of u] g_real[of u] g_nz[of u] u' + have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1) + with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases) + with c have A: "h z * g z = 0" for z by simp + hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp + hence "\<exists>c'. \<forall>z\<in>UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all + then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm) + from this[of 0] have "c' = pi" unfolding g_def by simp + with c have "g z = pi" by simp + + show ?thesis + proof (cases "z \<in> \<int>") + case False + with \<open>g z = pi\<close> show ?thesis by (auto simp: g_def divide_simps) + next + case True + then obtain n where n: "z = of_int n" by (elim Ints_cases) + with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force + moreover have "of_int (1 - n) \<in> \<int>\<^sub>\<le>\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp + ultimately show ?thesis using n + by (cases "n \<le> 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int) + qed +qed + +lemma rGamma_reflection_complex: + "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi" + using Gamma_reflection_complex[of z] + by (simp add: Gamma_def divide_simps split: if_split_asm) + +lemma rGamma_reflection_complex': + "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi" +proof - + have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))" + using rGamma_plus1[of "-z", symmetric] by simp + also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi" + by (rule rGamma_reflection_complex) + finally show ?thesis by simp +qed + +lemma Gamma_reflection_complex': + "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))" + using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def divide_simps mult_ac) + + + +lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi" +proof - + from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1] + have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square) + hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp + also have "\<dots> = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all + finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all + moreover have "Gamma (1/2 :: real) \<ge> 0" using Gamma_real_pos[of "1/2"] by simp + ultimately show ?thesis by (rule real_sqrt_unique [symmetric]) +qed + +lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)" +proof - + have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp + also have "\<dots> = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real) + finally show ?thesis . +qed + +lemma Gamma_legendre_duplication: + fixes z :: complex + assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "Gamma z * Gamma (z + 1/2) = + exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)" + using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex) + +end + + +subsection \<open>Limits and residues\<close> + +text \<open> + The inverse of the Gamma function has simple zeros: +\<close> + +lemma rGamma_zeros: + "(\<lambda>z. rGamma z / (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n * fact n :: 'a :: Gamma)" +proof (subst tendsto_cong) + let ?f = "\<lambda>z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a" + from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] + show "eventually (\<lambda>z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))" + by (subst pochhammer_rGamma[of _ "Suc n"]) + (auto elim!: eventually_mono simp: divide_simps pochhammer_rec' eq_neg_iff_add_eq_0) + have "isCont ?f (- of_nat n)" by (intro continuous_intros) + thus "?f \<midarrow> (- of_nat n) \<rightarrow> (- 1) ^ n * fact n" unfolding isCont_def + by (simp add: pochhammer_same) +qed + + +text \<open> + The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function, + and their residues can easily be computed from the limit we have just proven: +\<close> + +lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))" +proof - + from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] + have "eventually (\<lambda>z. rGamma z \<noteq> (0 :: 'a)) (at (- of_nat n))" + by (auto elim!: eventually_mono nonpos_Ints_cases' + simp: rGamma_eq_zero_iff dist_of_nat dist_minus) + with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident] + have "filterlim (\<lambda>z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))" + unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity]) + (simp_all add: filterlim_at) + moreover have "(\<lambda>z. inverse (rGamma z) :: 'a) = Gamma" + by (intro ext) (simp add: rGamma_inverse_Gamma) + ultimately show ?thesis by (simp only: ) +qed + +lemma Gamma_residues: + "(\<lambda>z. Gamma z * (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n / fact n :: 'a :: Gamma)" +proof (subst tendsto_cong) + let ?c = "(- 1) ^ n / fact n :: 'a" + from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] + show "eventually (\<lambda>z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n))) + (at (- of_nat n))" + by (auto elim!: eventually_mono simp: divide_simps rGamma_inverse_Gamma) + have "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> + inverse ((- 1) ^ n * fact n :: 'a)" + by (intro tendsto_intros rGamma_zeros) simp_all + also have "inverse ((- 1) ^ n * fact n) = ?c" + by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib) + finally show "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> ?c" . +qed + + + +subsection \<open>Alternative definitions\<close> + + +subsubsection \<open>Variant of the Euler form\<close> + + +definition Gamma_series_euler' where + "Gamma_series_euler' z n = + inverse z * (\<Prod>k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))" + +context +begin +private lemma Gamma_euler'_aux1: + fixes z :: "'a :: {real_normed_field,banach}" + assumes n: "n > 0" + shows "exp (z * of_real (ln (of_nat n + 1))) = (\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))" +proof - + have "(\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) = + exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))" + by (subst exp_setsum [symmetric]) (simp_all add: setsum_right_distrib) + also have "(\<Sum>k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\<Prod>k=1..n. 1 + 1 / real_of_nat k)" + by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg) + also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)" + by (intro setprod.cong) (simp_all add: divide_simps) + also have "(\<Prod>k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1" + by (induction n) (simp_all add: setprod_nat_ivl_Suc' divide_simps) + finally show ?thesis .. +qed + +lemma Gamma_series_euler': + assumes z: "(z :: 'a :: Gamma) \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(\<lambda>n. Gamma_series_euler' z n) \<longlonglongrightarrow> Gamma z" +proof (rule Gamma_seriesI, rule Lim_transform_eventually) + let ?f = "\<lambda>n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)" + let ?r = "\<lambda>n. ?f n / Gamma_series z n" + let ?r' = "\<lambda>n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))" + from z have z': "z \<noteq> 0" by auto + + have "eventually (\<lambda>n. ?r' n = ?r n) sequentially" using eventually_gt_at_top[of "0::nat"] + using z by (auto simp: divide_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac + elim!: eventually_mono dest: pochhammer_eq_0_imp_nonpos_Int) + moreover have "?r' \<longlonglongrightarrow> exp (z * of_real (ln 1))" + by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all + ultimately show "?r \<longlonglongrightarrow> 1" by (force dest!: Lim_transform_eventually) + + from eventually_gt_at_top[of "0::nat"] + show "eventually (\<lambda>n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially" + proof eventually_elim + fix n :: nat assume n: "n > 0" + from n z' have "Gamma_series_euler' z n = + exp (z * of_real (ln (of_nat n + 1))) / (z * (\<Prod>k=1..n. (1 + z / of_nat k)))" + by (subst Gamma_euler'_aux1) + (simp_all add: Gamma_series_euler'_def setprod.distrib + setprod_inversef[symmetric] divide_inverse) + also have "(\<Prod>k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n" + by (cases n) (simp_all add: pochhammer_setprod fact_setprod atLeastLessThanSuc_atLeastAtMost + setprod_dividef [symmetric] field_simps setprod.atLeast_Suc_atMost_Suc_shift) + also have "z * \<dots> = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec) + finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp + qed +qed + +end + + + +subsubsection \<open>Weierstrass form\<close> + +definition Gamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where + "Gamma_series_weierstrass z n = + exp (-euler_mascheroni * z) / z * (\<Prod>k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))" + +definition rGamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where + "rGamma_series_weierstrass z n = + exp (euler_mascheroni * z) * z * (\<Prod>k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))" + +lemma Gamma_series_weierstrass_nonpos_Ints: + "eventually (\<lambda>k. Gamma_series_weierstrass (- of_nat n) k = 0) sequentially" + using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_weierstrass_def) + +lemma rGamma_series_weierstrass_nonpos_Ints: + "eventually (\<lambda>k. rGamma_series_weierstrass (- of_nat n) k = 0) sequentially" + using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_weierstrass_def) + +lemma Gamma_weierstrass_complex: "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma (z :: complex)" +proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") + case True + then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') + also from True have "Gamma_series_weierstrass \<dots> \<longlonglongrightarrow> Gamma z" + by (simp add: tendsto_cong[OF Gamma_series_weierstrass_nonpos_Ints] Gamma_nonpos_Int) + finally show ?thesis . +next + case False + hence z: "z \<noteq> 0" by auto + let ?f = "(\<lambda>x. \<Prod>x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))" + have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \<ge> 1" for n :: nat + using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp) + have "(\<lambda>n. \<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \<longlonglongrightarrow> ln_Gamma z + euler_mascheroni * z + ln z" + using ln_Gamma_series'_aux[OF False] + by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def + setsum_shift_bounds_Suc_ivl sums_def atLeast0LessThan) + from tendsto_exp[OF this] False z have "?f \<longlonglongrightarrow> z * exp (euler_mascheroni * z) * Gamma z" + by (simp add: exp_add exp_setsum exp_diff mult_ac Gamma_complex_altdef A) + from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z + show "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma z" + by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def]) +qed + +lemma tendsto_complex_of_real_iff: "((\<lambda>x. complex_of_real (f x)) \<longlongrightarrow> of_real c) F = (f \<longlongrightarrow> c) F" + by (rule tendsto_of_real_iff) + +lemma Gamma_weierstrass_real: "Gamma_series_weierstrass x \<longlonglongrightarrow> Gamma (x :: real)" + using Gamma_weierstrass_complex[of "of_real x"] unfolding Gamma_series_weierstrass_def[abs_def] + by (subst tendsto_complex_of_real_iff [symmetric]) + (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real) + +lemma rGamma_weierstrass_complex: "rGamma_series_weierstrass z \<longlonglongrightarrow> rGamma (z :: complex)" +proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") + case True + then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') + also from True have "rGamma_series_weierstrass \<dots> \<longlonglongrightarrow> rGamma z" + by (simp add: tendsto_cong[OF rGamma_series_weierstrass_nonpos_Ints] rGamma_nonpos_Int) + finally show ?thesis . +next + case False + have "rGamma_series_weierstrass z = (\<lambda>n. inverse (Gamma_series_weierstrass z n))" + by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def + exp_minus divide_inverse setprod_inversef[symmetric] mult_ac) + also from False have "\<dots> \<longlonglongrightarrow> inverse (Gamma z)" + by (intro tendsto_intros Gamma_weierstrass_complex) (simp add: Gamma_eq_zero_iff) + finally show ?thesis by (simp add: Gamma_def) +qed + +subsubsection \<open>Binomial coefficient form\<close> + +lemma Gamma_gbinomial: + "(\<lambda>n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \<longlonglongrightarrow> rGamma (z+1)" +proof (cases "z = 0") + case False + show ?thesis + proof (rule Lim_transform_eventually) + let ?powr = "\<lambda>a b. exp (b * of_real (ln (of_nat a)))" + show "eventually (\<lambda>n. rGamma_series z n / z = + ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially" + proof (intro always_eventually allI) + fix n :: nat + from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n" + by (simp add: gbinomial_pochhammer' pochhammer_rec) + also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z" + by (simp add: rGamma_series_def divide_simps exp_minus) + finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" .. + qed + + from False have "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma z / z" by (intro tendsto_intros) + also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z] + by (simp add: field_simps) + finally show "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma (z+1)" . + qed +qed (simp_all add: binomial_gbinomial [symmetric]) + +lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)" + by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric]) + +lemma gbinomial_asymptotic: + fixes z :: "'a :: Gamma" + shows "(\<lambda>n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \<longlonglongrightarrow> + inverse (Gamma (- z))" + unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"] + by (subst (asm) gbinomial_minus') + (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric]) + +lemma fact_binomial_limit: + "(\<lambda>n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \<longlonglongrightarrow> 1 / fact k" +proof (rule Lim_transform_eventually) + have "(\<lambda>n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n)))) + \<longlonglongrightarrow> 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \<longlonglongrightarrow> _") + using Gamma_gbinomial[of "of_nat k :: 'a"] + by (simp add: binomial_gbinomial add_ac Gamma_def divide_simps exp_of_real [symmetric] exp_minus) + also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact) + finally show "?f \<longlonglongrightarrow> 1 / fact k" . + + show "eventually (\<lambda>n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)" + by (simp add: exp_of_nat_mult) + thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp + qed +qed + +lemma binomial_asymptotic': + "(\<lambda>n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \<longlonglongrightarrow> 1" + using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp + +lemma gbinomial_Beta: + assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))" +using assms +proof (induction n arbitrary: z) + case 0 + hence "z + 2 \<notin> \<int>\<^sub>\<le>\<^sub>0" + using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute) + with 0 show ?case + by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute) +next + case (Suc n z) + show ?case + proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") + case True + with Suc.prems have "z = 0" + by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff) + show ?thesis + proof (cases "n = 0") + case True + with \<open>z = 0\<close> show ?thesis + by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric]) + next + case False + with \<open>z = 0\<close> show ?thesis + by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1) + qed + next + case False + have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp + also have "\<dots> = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)" + by (subst gbinomial_factors) (simp add: field_simps) + also from False have "\<dots> = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))" + (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1) + also have "of_nat (Suc n) \<notin> (\<int>\<^sub>\<le>\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all + hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)" + by (subst Beta_plus1_right [symmetric]) simp_all + finally show ?thesis . + qed +qed + +lemma gbinomial_Gamma: + assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" + shows "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))" +proof - + have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))" + by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac) + also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)" + using Gamma_plus1[of "z+1"] by (auto simp add: divide_simps mult_ac add_ac) + finally show ?thesis . +qed + + +subsubsection \<open>Integral form\<close> + +lemma integrable_Gamma_integral_bound: + fixes a c :: real + assumes a: "a > -1" and c: "c \<ge> 0" + defines "f \<equiv> \<lambda>x. if x \<in> {0..c} then x powr a else exp (-x/2)" + shows "f integrable_on {0..}" +proof - + have "f integrable_on {0..c}" + by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]]) + (insert a c, simp_all add: f_def) + moreover have A: "(\<lambda>x. exp (-x/2)) integrable_on {c..}" + using integrable_on_exp_minus_to_infinity[of "1/2"] by simp + have "f integrable_on {c..}" + by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def) + ultimately show "f integrable_on {0..}" + by (rule integrable_union') (insert c, auto simp: max_def) +qed + +lemma Gamma_integral_complex: + assumes z: "Re z > 0" + shows "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" +proof - + have A: "((\<lambda>t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n)) + has_integral (fact n / pochhammer z (n+1))) {0..1}" + if "Re z > 0" for n z using that + proof (induction n arbitrary: z) + case 0 + have "((\<lambda>t. complex_of_real t powr (z - 1)) has_integral + (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0 + by (intro fundamental_theorem_of_calculus_interior) + (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_complex) + thus ?case by simp + next + case (Suc n) + let ?f = "\<lambda>t. complex_of_real t powr z / z" + let ?f' = "\<lambda>t. complex_of_real t powr (z - 1)" + let ?g = "\<lambda>t. (1 - complex_of_real t) ^ Suc n" + let ?g' = "\<lambda>t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)" + have "((\<lambda>t. ?f' t * ?g t) has_integral + (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}" + (is "(_ has_integral ?I) _") + proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g]) + from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g" + by (auto intro!: continuous_intros) + next + fix t :: real assume t: "t \<in> {0<..<1}" + show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems + by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex) + show "(?g has_vector_derivative ?g' t) (at t)" + by (rule has_vector_derivative_real_complex derivative_eq_intros refl)+ simp_all + next + from Suc.prems have [simp]: "z \<noteq> 0" by auto + from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp + have [simp]: "z + of_nat n \<noteq> 0" "z + 1 + of_nat n \<noteq> 0" for n + using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel) + have "((\<lambda>x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral + fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}" + (is "(?A has_integral ?B) _") + using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec) + also have "?A = (\<lambda>t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps) + also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))" + by (simp add: divide_simps pochhammer_rec + setprod_shift_bounds_cl_Suc_ivl del: of_nat_Suc) + finally show "((\<lambda>t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}" + by simp + qed (simp_all add: bounded_bilinear_mult) + thus ?case by simp + qed + + have B: "((\<lambda>t. if t \<in> {0..of_nat n} then + of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0) + has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n + proof (cases "n > 0") + case [simp]: True + hence [simp]: "n \<noteq> 0" by auto + with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0] + have "((\<lambda>x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n) + has_integral fact n * of_nat n / pochhammer z (n+1)) ((\<lambda>x. real n * x)`{0..1})" + (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real) + also from True have "((\<lambda>x. real n*x)`{0..1}) = {0..real n}" + by (subst image_mult_atLeastAtMost) simp_all + also have "?f = (\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)" + using True by (intro ext) (simp add: field_simps) + finally have "((\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n) + has_integral ?I) {0..real n}" (is ?P) . + also have "?P \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n) + has_integral ?I) {0..real n}" + by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric]) + also have "\<dots> \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n) + has_integral ?I) {0..real n}" + by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div) + finally have \<dots> . + note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"] + have "((\<lambda>x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n) + has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P) + by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric]) + (simp add: Ln_of_nat algebra_simps) + also have "?P \<longleftrightarrow> ((\<lambda>x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n) + has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" + by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real) + also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) = + (of_nat n powr z * fact n / pochhammer z (n+1))" + by (auto simp add: powr_def algebra_simps exp_diff) + finally show ?thesis by (subst has_integral_restrict) simp_all + next + case False + thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl) + qed + + have "eventually (\<lambda>n. Gamma_series z n = + of_nat n powr z * fact n / pochhammer z (n+1)) sequentially" + using eventually_gt_at_top[of "0::nat"] + by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def) + from this and Gamma_series_LIMSEQ[of z] + have C: "(\<lambda>k. of_nat k powr z * fact k / pochhammer z (k+1)) \<longlonglongrightarrow> Gamma z" + by (rule Lim_transform_eventually) + + { + fix x :: real assume x: "x \<ge> 0" + have lim_exp: "(\<lambda>k. (1 - x / real k) ^ k) \<longlonglongrightarrow> exp (-x)" + using tendsto_exp_limit_sequentially[of "-x"] by simp + have "(\<lambda>k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k)) + \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))" (is ?P) + by (intro tendsto_intros lim_exp) + also from eventually_gt_at_top[of "nat \<lceil>x\<rceil>"] + have "eventually (\<lambda>k. of_nat k > x) sequentially" by eventually_elim linarith + hence "?P \<longleftrightarrow> (\<lambda>k. if x \<le> of_nat k then + of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0) + \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))" + by (intro tendsto_cong) (auto elim!: eventually_mono) + finally have \<dots> . + } + hence D: "\<forall>x\<in>{0..}. (\<lambda>k. if x \<in> {0..real k} then + of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0) + \<longlonglongrightarrow> of_real x powr (z - 1) / of_real (exp x)" + by (simp add: exp_minus field_simps cong: if_cong) + + have "((\<lambda>x. (Re z - 1) * (ln x / x)) \<longlongrightarrow> (Re z - 1) * 0) at_top" + by (intro tendsto_intros ln_x_over_x_tendsto_0) + hence "((\<lambda>x. ((Re z - 1) * ln x) / x) \<longlongrightarrow> 0) at_top" by simp + from order_tendstoD(2)[OF this, of "1/2"] + have "eventually (\<lambda>x. (Re z - 1) * ln x / x < 1/2) at_top" by simp + from eventually_conj[OF this eventually_gt_at_top[of 0]] + obtain x0 where "\<forall>x\<ge>x0. (Re z - 1) * ln x / x < 1/2 \<and> x > 0" + by (auto simp: eventually_at_top_linorder) + hence x0: "x0 > 0" "\<And>x. x \<ge> x0 \<Longrightarrow> (Re z - 1) * ln x < x / 2" by auto + + define h where "h = (\<lambda>x. if x \<in> {0..x0} then x powr (Re z - 1) else exp (-x/2))" + have le_h: "x powr (Re z - 1) * exp (-x) \<le> h x" if x: "x \<ge> 0" for x + proof (cases "x > x0") + case True + from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)" + by (simp add: powr_def exp_diff exp_minus field_simps exp_add) + also from x0(2)[of x] True have "\<dots> < exp (-x/2)" + by (simp add: field_simps) + finally show ?thesis using True by (auto simp add: h_def) + next + case False + from x have "x powr (Re z - 1) * exp (- x) \<le> x powr (Re z - 1) * 1" + by (intro mult_left_mono) simp_all + with False show ?thesis by (auto simp add: h_def) + qed + + have E: "\<forall>x\<in>{0..}. cmod (if x \<in> {0..real k} then of_real x powr (z - 1) * + (1 - complex_of_real x / of_nat k) ^ k else 0) \<le> h x" + (is "\<forall>x\<in>_. ?f x \<le> _") for k + proof safe + fix x :: real assume x: "x \<ge> 0" + { + fix x :: real and n :: nat assume x: "x \<le> of_nat n" + have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp + also have "norm \<dots> = \<bar>(1 - x / real n)\<bar>" by (subst norm_of_real) (rule refl) + also from x have "\<dots> = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: divide_simps) + finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" . + } note D = this + from D[of x k] x + have "?f x \<le> (if of_nat k \<ge> x \<and> k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)" + by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg) + also have "\<dots> \<le> x powr (Re z - 1) * exp (-x)" + by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n) + also from x have "\<dots> \<le> h x" by (rule le_h) + finally show "?f x \<le> h x" . + qed + + have F: "h integrable_on {0..}" unfolding h_def + by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all) + show ?thesis + by (rule has_integral_dominated_convergence[OF B F E D C]) +qed + +lemma Gamma_integral_real: + assumes x: "x > (0 :: real)" + shows "((\<lambda>t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}" +proof - + have A: "((\<lambda>t. complex_of_real t powr (complex_of_real x - 1) / + complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}" + using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real) + have "((\<lambda>t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}" + by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric]) + from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def) +qed + + + +subsection \<open>The Weierstraß product formula for the sine\<close> + +lemma sin_product_formula_complex: + fixes z :: complex + shows "(\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k^2)) \<longlonglongrightarrow> sin (of_real pi * z)" +proof - + let ?f = "rGamma_series_weierstrass" + have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n)) + \<longlonglongrightarrow> (- of_real pi * inverse z) * (rGamma z * rGamma (- z))" + by (intro tendsto_intros rGamma_weierstrass_complex) + also have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) = + (\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2))" + proof + fix n :: nat + have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = + of_real pi * z * (\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)" + by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus + divide_simps setprod.distrib[symmetric] power2_eq_square) + also have "(\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) = + (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2)" + by (intro setprod.cong) (simp_all add: power2_eq_square field_simps) + finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \<dots>" + by (simp add: divide_simps) + qed + also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)" + by (subst rGamma_reflection_complex') (simp add: divide_simps) + finally show ?thesis . +qed + +lemma sin_product_formula_real: + "(\<lambda>n. pi * (x::real) * (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x)" +proof - + from sin_product_formula_complex[of "of_real x"] + have "(\<lambda>n. of_real pi * of_real x * (\<Prod>k=1..n. 1 - (of_real x)^2 / (of_nat k)^2)) + \<longlonglongrightarrow> sin (of_real pi * of_real x :: complex)" (is "?f \<longlonglongrightarrow> ?y") . + also have "?f = (\<lambda>n. of_real (pi * x * (\<Prod>k=1..n. 1 - x^2 / (of_nat k^2))))" by simp + also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult) + finally show ?thesis by (subst (asm) tendsto_of_real_iff) +qed + +lemma sin_product_formula_real': + assumes "x \<noteq> (0::real)" + shows "(\<lambda>n. (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x) / (pi * x)" + using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms + by simp + + +subsection \<open>The Solution to the Basel problem\<close> + +theorem inverse_squares_sums: "(\<lambda>n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)" +proof - + define P where "P x n = (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n + define K where "K = (\<Sum>n. inverse (real_of_nat (Suc n))^2)" + define f where [abs_def]: "f x = (\<Sum>n. P x n / of_nat (Suc n)^2)" for x + define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x + + have sums: "(\<lambda>n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x + proof (cases "x = 0") + assume x: "x = 0" + have "summable (\<lambda>n. inverse ((real_of_nat (Suc n))\<^sup>2))" + using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp + thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums) + next + assume x: "x \<noteq> 0" + have "(\<lambda>n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))" + unfolding P_def using x by (intro telescope_sums' sin_product_formula_real') + also have "(\<lambda>n. P x n - P x (Suc n)) = (\<lambda>n. (x^2 / of_nat (Suc n)^2) * P x n)" + unfolding P_def by (simp add: setprod_nat_ivl_Suc' algebra_simps) + also have "P x 0 = 1" by (simp add: P_def) + finally have "(\<lambda>n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" . + from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp + qed + + have "continuous_on (ball 0 1) f" + proof (rule uniform_limit_theorem; (intro always_eventually allI)?) + show "uniform_limit (ball 0 1) (\<lambda>n x. \<Sum>k<n. P x k / of_nat (Suc k)^2) f sequentially" + proof (unfold f_def, rule weierstrass_m_test) + fix n :: nat and x :: real assume x: "x \<in> ball 0 1" + { + fix k :: nat assume k: "k \<ge> 1" + from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1) + also from k have "\<dots> \<le> of_nat k^2" by simp + finally have "(1 - x^2 / of_nat k^2) \<in> {0..1}" using k + by (simp_all add: field_simps del: of_nat_Suc) + } + hence "(\<Prod>k=1..n. abs (1 - x^2 / of_nat k^2)) \<le> (\<Prod>k=1..n. 1)" by (intro setprod_mono) simp + thus "norm (P x n / (of_nat (Suc n)^2)) \<le> 1 / of_nat (Suc n)^2" + unfolding P_def by (simp add: field_simps abs_setprod del: of_nat_Suc) + qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide) + qed (auto simp: P_def intro!: continuous_intros) + hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all + hence "(f \<midarrow> 0 \<rightarrow> f 0)" by (simp add: isCont_def) + also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide) + finally have "f \<midarrow> 0 \<rightarrow> K" . + + moreover have "f \<midarrow> 0 \<rightarrow> pi^2 / 6" + proof (rule Lim_transform_eventually) + define f' where [abs_def]: "f' x = (\<Sum>n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x + have "eventually (\<lambda>x. x \<noteq> (0::real)) (at 0)" + by (auto simp add: eventually_at intro!: exI[of _ 1]) + thus "eventually (\<lambda>x. f' x = f x) (at 0)" + proof eventually_elim + fix x :: real assume x: "x \<noteq> 0" + have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def) + with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"] + have "(\<lambda>n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))" + by (simp add: eval_nat_numeral) + from sums_divide[OF this, of "x^3 * pi"] x + have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)" + by (simp add: divide_simps eval_nat_numeral power_mult_distrib mult_ac) + with x have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)" + by (simp add: g_def) + hence "f' x = g x / x^2" by (simp add: sums_iff f'_def) + also have "\<dots> = f x" using sums[of x] x by (simp add: sums_iff g_def f_def) + finally show "f' x = f x" . + qed + + have "isCont f' 0" unfolding f'_def + proof (intro isCont_powser_converges_everywhere) + fix x :: real show "summable (\<lambda>n. -sin_coeff (n+3) * pi^(n+2) * x^n)" + proof (cases "x = 0") + assume x: "x \<noteq> 0" + from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF + sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x + show ?thesis by (simp add: mult_ac power_mult_distrib divide_simps eval_nat_numeral) + qed (simp only: summable_0_powser) + qed + hence "f' \<midarrow> 0 \<rightarrow> f' 0" by (simp add: isCont_def) + also have "f' 0 = pi * pi / fact 3" unfolding f'_def + by (subst powser_zero) (simp add: sin_coeff_def) + finally show "f' \<midarrow> 0 \<rightarrow> pi^2 / 6" by (simp add: eval_nat_numeral) + qed + + ultimately have "K = pi^2 / 6" by (rule LIM_unique) + moreover from inverse_power_summable[of 2] + have "summable (\<lambda>n. (inverse (real_of_nat (Suc n)))\<^sup>2)" + by (subst summable_Suc_iff) (simp add: power_inverse) + ultimately show ?thesis unfolding K_def + by (auto simp add: sums_iff power_divide inverse_eq_divide) +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Generalised_Binomial_Theorem.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,256 @@ +(* Title: HOL/Analysis/Generalised_Binomial_Theorem.thy + Author: Manuel Eberl, TU München +*) + +section \<open>Generalised Binomial Theorem\<close> + +text \<open> + The proof of the Generalised Binomial Theorem and related results. + We prove the generalised binomial theorem for complex numbers, following the proof at: + \url{https://proofwiki.org/wiki/Binomial_Theorem/General_Binomial_Theorem} +\<close> + +theory Generalised_Binomial_Theorem +imports + Complex_Main + Complex_Transcendental + Summation_Tests +begin + +lemma gbinomial_ratio_limit: + fixes a :: "'a :: real_normed_field" + assumes "a \<notin> \<nat>" + shows "(\<lambda>n. (a gchoose n) / (a gchoose Suc n)) \<longlonglongrightarrow> -1" +proof (rule Lim_transform_eventually) + let ?f = "\<lambda>n. inverse (a / of_nat (Suc n) - of_nat n / of_nat (Suc n))" + from eventually_gt_at_top[of "0::nat"] + show "eventually (\<lambda>n. ?f n = (a gchoose n) /(a gchoose Suc n)) sequentially" + proof eventually_elim + fix n :: nat assume n: "n > 0" + then obtain q where q: "n = Suc q" by (cases n) blast + let ?P = "\<Prod>i=0..<n. a - of_nat i" + from n have "(a gchoose n) / (a gchoose Suc n) = (of_nat (Suc n) :: 'a) * + (?P / (\<Prod>i=0..n. a - of_nat i))" + by (simp add: gbinomial_setprod_rev atLeastLessThanSuc_atLeastAtMost) + also from q have "(\<Prod>i=0..n. a - of_nat i) = ?P * (a - of_nat n)" + by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) + also have "?P / \<dots> = (?P / ?P) / (a - of_nat n)" by (rule divide_divide_eq_left[symmetric]) + also from assms have "?P / ?P = 1" by auto + also have "of_nat (Suc n) * (1 / (a - of_nat n)) = + inverse (inverse (of_nat (Suc n)) * (a - of_nat n))" by (simp add: field_simps) + also have "inverse (of_nat (Suc n)) * (a - of_nat n) = a / of_nat (Suc n) - of_nat n / of_nat (Suc n)" + by (simp add: field_simps del: of_nat_Suc) + finally show "?f n = (a gchoose n) / (a gchoose Suc n)" by simp + qed + + have "(\<lambda>n. norm a / (of_nat (Suc n))) \<longlonglongrightarrow> 0" + unfolding divide_inverse + by (intro tendsto_mult_right_zero LIMSEQ_inverse_real_of_nat) + hence "(\<lambda>n. a / of_nat (Suc n)) \<longlonglongrightarrow> 0" + by (subst tendsto_norm_zero_iff[symmetric]) (simp add: norm_divide del: of_nat_Suc) + hence "?f \<longlonglongrightarrow> inverse (0 - 1)" + by (intro tendsto_inverse tendsto_diff LIMSEQ_n_over_Suc_n) simp_all + thus "?f \<longlonglongrightarrow> -1" by simp +qed + +lemma conv_radius_gchoose: + fixes a :: "'a :: {real_normed_field,banach}" + shows "conv_radius (\<lambda>n. a gchoose n) = (if a \<in> \<nat> then \<infinity> else 1)" +proof (cases "a \<in> \<nat>") + assume a: "a \<in> \<nat>" + have "eventually (\<lambda>n. (a gchoose n) = 0) sequentially" + using eventually_gt_at_top[of "nat \<lfloor>norm a\<rfloor>"] + by eventually_elim (insert a, auto elim!: Nats_cases simp: binomial_gbinomial[symmetric]) + from conv_radius_cong[OF this] a show ?thesis by simp +next + assume a: "a \<notin> \<nat>" + from tendsto_norm[OF gbinomial_ratio_limit[OF this]] + have "conv_radius (\<lambda>n. a gchoose n) = 1" + by (intro conv_radius_ratio_limit_nonzero[of _ 1]) (simp_all add: norm_divide) + with a show ?thesis by simp +qed + +lemma gen_binomial_complex: + fixes z :: complex + assumes "norm z < 1" + shows "(\<lambda>n. (a gchoose n) * z^n) sums (1 + z) powr a" +proof - + define K where "K = 1 - (1 - norm z) / 2" + from assms have K: "K > 0" "K < 1" "norm z < K" + unfolding K_def by (auto simp: field_simps intro!: add_pos_nonneg) + let ?f = "\<lambda>n. a gchoose n" and ?f' = "diffs (\<lambda>n. a gchoose n)" + have summable_strong: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < 1" for z using that + by (intro summable_in_conv_radius) (simp_all add: conv_radius_gchoose) + with K have summable: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < K" for z using that by auto + hence summable': "summable (\<lambda>n. ?f' n * z ^ n)" if "norm z < K" for z using that + by (intro termdiff_converges[of _ K]) simp_all + + define f f' where [abs_def]: "f z = (\<Sum>n. ?f n * z ^ n)" "f' z = (\<Sum>n. ?f' n * z ^ n)" for z + { + fix z :: complex assume z: "norm z < K" + from summable_mult2[OF summable'[OF z], of z] + have summable1: "summable (\<lambda>n. ?f' n * z ^ Suc n)" by (simp add: mult_ac) + hence summable2: "summable (\<lambda>n. of_nat n * ?f n * z^n)" + unfolding diffs_def by (subst (asm) summable_Suc_iff) + + have "(1 + z) * f' z = (\<Sum>n. ?f' n * z^n) + (\<Sum>n. ?f' n * z^Suc n)" + unfolding f_f'_def using summable' z by (simp add: algebra_simps suminf_mult) + also have "(\<Sum>n. ?f' n * z^n) = (\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n)" + by (intro suminf_cong) (simp add: diffs_def) + also have "(\<Sum>n. ?f' n * z^Suc n) = (\<Sum>n. of_nat n * ?f n * z ^ n)" + using summable1 suminf_split_initial_segment[OF summable1] unfolding diffs_def + by (subst suminf_split_head, subst (asm) summable_Suc_iff) simp_all + also have "(\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n) + (\<Sum>n. of_nat n * ?f n * z^n) = + (\<Sum>n. a * ?f n * z^n)" + by (subst gbinomial_mult_1, subst suminf_add) + (insert summable'[OF z] summable2, + simp_all add: summable_powser_split_head algebra_simps diffs_def) + also have "\<dots> = a * f z" unfolding f_f'_def + by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac) + finally have "a * f z = (1 + z) * f' z" by simp + } note deriv = this + + have [derivative_intros]: "(f has_field_derivative f' z) (at z)" if "norm z < of_real K" for z + unfolding f_f'_def using K that + by (intro termdiffs_strong[of "?f" K z] summable_strong) simp_all + have "f 0 = (\<Sum>n. if n = 0 then 1 else 0)" unfolding f_f'_def by (intro suminf_cong) simp + also have "\<dots> = 1" using sums_single[of 0 "\<lambda>_. 1::complex"] unfolding sums_iff by simp + finally have [simp]: "f 0 = 1" . + + have "\<exists>c. \<forall>z\<in>ball 0 K. f z * (1 + z) powr (-a) = c" + proof (rule has_field_derivative_zero_constant) + fix z :: complex assume z': "z \<in> ball 0 K" + hence z: "norm z < K" by simp + with K have nz: "1 + z \<noteq> 0" by (auto dest!: minus_unique) + from z K have "norm z < 1" by simp + hence "(1 + z) \<notin> \<real>\<^sub>\<le>\<^sub>0" by (cases z) (auto simp: complex_nonpos_Reals_iff) + hence "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative + f' z * (1 + z) powr (-a) - a * f z * (1 + z) powr (-a-1)) (at z)" using z + by (auto intro!: derivative_eq_intros) + also from z have "a * f z = (1 + z) * f' z" by (rule deriv) + finally show "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z within ball 0 K)" + using nz by (simp add: field_simps powr_diff_complex at_within_open[OF z']) + qed simp_all + then obtain c where c: "\<And>z. z \<in> ball 0 K \<Longrightarrow> f z * (1 + z) powr (-a) = c" by blast + from c[of 0] and K have "c = 1" by simp + with c[of z] have "f z = (1 + z) powr a" using K + by (simp add: powr_minus_complex field_simps dist_complex_def) + with summable K show ?thesis unfolding f_f'_def by (simp add: sums_iff) +qed + +lemma gen_binomial_complex': + fixes x y :: real and a :: complex + assumes "\<bar>x\<bar> < \<bar>y\<bar>" + shows "(\<lambda>n. (a gchoose n) * of_real x^n * of_real y powr (a - of_nat n)) sums + of_real (x + y) powr a" (is "?P x y") +proof - + { + fix x y :: real assume xy: "\<bar>x\<bar> < \<bar>y\<bar>" "y \<ge> 0" + hence "y > 0" by simp + note xy = xy this + from xy have "(\<lambda>n. (a gchoose n) * of_real (x / y) ^ n) sums (1 + of_real (x / y)) powr a" + by (intro gen_binomial_complex) (simp add: norm_divide) + hence "(\<lambda>n. (a gchoose n) * of_real (x / y) ^ n * y powr a) sums + ((1 + of_real (x / y)) powr a * y powr a)" + by (rule sums_mult2) + also have "(1 + complex_of_real (x / y)) = complex_of_real (1 + x/y)" by simp + also from xy have "\<dots> powr a * of_real y powr a = (\<dots> * y) powr a" + by (subst powr_times_real[symmetric]) (simp_all add: field_simps) + also from xy have "complex_of_real (1 + x / y) * complex_of_real y = of_real (x + y)" + by (simp add: field_simps) + finally have "?P x y" using xy by (simp add: field_simps powr_diff_complex powr_nat) + } note A = this + + show ?thesis + proof (cases "y < 0") + assume y: "y < 0" + with assms have xy: "x + y < 0" by simp + with assms have "\<bar>-x\<bar> < \<bar>-y\<bar>" "-y \<ge> 0" by simp_all + note A[OF this] + also have "complex_of_real (-x + -y) = - complex_of_real (x + y)" by simp + also from xy assms have "... powr a = (-1) powr -a * of_real (x + y) powr a" + by (subst powr_neg_real_complex) (simp add: abs_real_def split: if_split_asm) + also { + fix n :: nat + from y have "(a gchoose n) * of_real (-x) ^ n * of_real (-y) powr (a - of_nat n) = + (a gchoose n) * (-of_real x / -of_real y) ^ n * (- of_real y) powr a" + by (subst power_divide) (simp add: powr_diff_complex powr_nat) + also from y have "(- of_real y) powr a = (-1) powr -a * of_real y powr a" + by (subst powr_neg_real_complex) simp + also have "-complex_of_real x / -complex_of_real y = complex_of_real x / complex_of_real y" + by simp + also have "... ^ n = of_real x ^ n / of_real y ^ n" by (simp add: power_divide) + also have "(a gchoose n) * ... * ((-1) powr -a * of_real y powr a) = + (-1) powr -a * ((a gchoose n) * of_real x ^ n * of_real y powr (a - n))" + by (simp add: algebra_simps powr_diff_complex powr_nat) + finally have "(a gchoose n) * of_real (- x) ^ n * of_real (- y) powr (a - of_nat n) = + (-1) powr -a * ((a gchoose n) * of_real x ^ n * of_real y powr (a - of_nat n))" . + } + note sums_cong[OF this] + finally show ?thesis by (simp add: sums_mult_iff) + qed (insert A[of x y] assms, simp_all add: not_less) +qed + +lemma gen_binomial_complex'': + fixes x y :: real and a :: complex + assumes "\<bar>y\<bar> < \<bar>x\<bar>" + shows "(\<lambda>n. (a gchoose n) * of_real x powr (a - of_nat n) * of_real y ^ n) sums + of_real (x + y) powr a" + using gen_binomial_complex'[OF assms] by (simp add: mult_ac add.commute) + +lemma gen_binomial_real: + fixes z :: real + assumes "\<bar>z\<bar> < 1" + shows "(\<lambda>n. (a gchoose n) * z^n) sums (1 + z) powr a" +proof - + from assms have "norm (of_real z :: complex) < 1" by simp + from gen_binomial_complex[OF this] + have "(\<lambda>n. (of_real a gchoose n :: complex) * of_real z ^ n) sums + (of_real (1 + z)) powr (of_real a)" by simp + also have "(of_real (1 + z) :: complex) powr (of_real a) = of_real ((1 + z) powr a)" + using assms by (subst powr_of_real) simp_all + also have "(of_real a gchoose n :: complex) = of_real (a gchoose n)" for n + by (simp add: gbinomial_setprod_rev) + hence "(\<lambda>n. (of_real a gchoose n :: complex) * of_real z ^ n) = + (\<lambda>n. of_real ((a gchoose n) * z ^ n))" by (intro ext) simp + finally show ?thesis by (simp only: sums_of_real_iff) +qed + +lemma gen_binomial_real': + fixes x y a :: real + assumes "\<bar>x\<bar> < y" + shows "(\<lambda>n. (a gchoose n) * x^n * y powr (a - of_nat n)) sums (x + y) powr a" +proof - + from assms have "y > 0" by simp + note xy = this assms + from assms have "\<bar>x / y\<bar> < 1" by simp + hence "(\<lambda>n. (a gchoose n) * (x / y) ^ n) sums (1 + x / y) powr a" + by (rule gen_binomial_real) + hence "(\<lambda>n. (a gchoose n) * (x / y) ^ n * y powr a) sums ((1 + x / y) powr a * y powr a)" + by (rule sums_mult2) + with xy show ?thesis + by (simp add: field_simps powr_divide powr_divide2[symmetric] powr_realpow) +qed + +lemma one_plus_neg_powr_powser: + fixes z s :: complex + assumes "norm (z :: complex) < 1" + shows "(\<lambda>n. (-1)^n * ((s + n - 1) gchoose n) * z^n) sums (1 + z) powr (-s)" + using gen_binomial_complex[OF assms, of "-s"] by (simp add: gbinomial_minus) + +lemma gen_binomial_real'': + fixes x y a :: real + assumes "\<bar>y\<bar> < x" + shows "(\<lambda>n. (a gchoose n) * x powr (a - of_nat n) * y^n) sums (x + y) powr a" + using gen_binomial_real'[OF assms] by (simp add: mult_ac add.commute) + +lemma sqrt_series': + "\<bar>z\<bar> < a \<Longrightarrow> (\<lambda>n. ((1/2) gchoose n) * a powr (1/2 - real_of_nat n) * z ^ n) sums + sqrt (a + z :: real)" + using gen_binomial_real''[of z a "1/2"] by (simp add: powr_half_sqrt) + +lemma sqrt_series: + "\<bar>z\<bar> < 1 \<Longrightarrow> (\<lambda>n. ((1/2) gchoose n) * z ^ n) sums sqrt (1 + z)" + using gen_binomial_real[of z "1/2"] by (simp add: powr_half_sqrt) + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Harmonic_Numbers.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,533 @@ +(* Title: HOL/Analysis/Harmonic_Numbers.thy + Author: Manuel Eberl, TU München +*) + +section \<open>Harmonic Numbers\<close> + +theory Harmonic_Numbers +imports + Complex_Transcendental + Summation_Tests + Integral_Test +begin + +text \<open> + The definition of the Harmonic Numbers and the Euler-Mascheroni constant. + Also provides a reasonably accurate approximation of @{term "ln 2 :: real"} + and the Euler-Mascheroni constant. +\<close> + +lemma ln_2_less_1: "ln 2 < (1::real)" +proof - + have "2 < 5/(2::real)" by simp + also have "5/2 \<le> exp (1::real)" using exp_lower_taylor_quadratic[of 1, simplified] by simp + finally have "exp (ln 2) < exp (1::real)" by simp + thus "ln 2 < (1::real)" by (subst (asm) exp_less_cancel_iff) simp +qed + +lemma setsum_Suc_diff': + fixes f :: "nat \<Rightarrow> 'a::ab_group_add" + assumes "m \<le> n" + shows "(\<Sum>i = m..<n. f (Suc i) - f i) = f n - f m" +using assms by (induct n) (auto simp: le_Suc_eq) + + +subsection \<open>The Harmonic numbers\<close> + +definition harm :: "nat \<Rightarrow> 'a :: real_normed_field" where + "harm n = (\<Sum>k=1..n. inverse (of_nat k))" + +lemma harm_altdef: "harm n = (\<Sum>k<n. inverse (of_nat (Suc k)))" + unfolding harm_def by (induction n) simp_all + +lemma harm_Suc: "harm (Suc n) = harm n + inverse (of_nat (Suc n))" + by (simp add: harm_def) + +lemma harm_nonneg: "harm n \<ge> (0 :: 'a :: {real_normed_field,linordered_field})" + unfolding harm_def by (intro setsum_nonneg) simp_all + +lemma harm_pos: "n > 0 \<Longrightarrow> harm n > (0 :: 'a :: {real_normed_field,linordered_field})" + unfolding harm_def by (intro setsum_pos) simp_all + +lemma of_real_harm: "of_real (harm n) = harm n" + unfolding harm_def by simp + +lemma norm_harm: "norm (harm n) = harm n" + by (subst of_real_harm [symmetric]) (simp add: harm_nonneg) + +lemma harm_expand: + "harm 0 = 0" + "harm (Suc 0) = 1" + "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)" +proof - + have "numeral n = Suc (pred_numeral n)" by simp + also have "harm \<dots> = harm (pred_numeral n) + inverse (numeral n)" + by (subst harm_Suc, subst numeral_eq_Suc[symmetric]) simp + finally show "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)" . +qed (simp_all add: harm_def) + +lemma not_convergent_harm: "\<not>convergent (harm :: nat \<Rightarrow> 'a :: real_normed_field)" +proof - + have "convergent (\<lambda>n. norm (harm n :: 'a)) \<longleftrightarrow> + convergent (harm :: nat \<Rightarrow> real)" by (simp add: norm_harm) + also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k=Suc 0..Suc n. inverse (of_nat k) :: real)" + unfolding harm_def[abs_def] by (subst convergent_Suc_iff) simp_all + also have "... \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k\<le>n. inverse (of_nat (Suc k)) :: real)" + by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: atLeast0AtMost) + also have "... \<longleftrightarrow> summable (\<lambda>n. inverse (of_nat n) :: real)" + by (subst summable_Suc_iff [symmetric]) (simp add: summable_iff_convergent') + also have "\<not>..." by (rule not_summable_harmonic) + finally show ?thesis by (blast dest: convergent_norm) +qed + +lemma harm_pos_iff [simp]: "harm n > (0 :: 'a :: {real_normed_field,linordered_field}) \<longleftrightarrow> n > 0" + by (rule iffI, cases n, simp add: harm_expand, simp, rule harm_pos) + +lemma ln_diff_le_inverse: + assumes "x \<ge> (1::real)" + shows "ln (x + 1) - ln x < 1 / x" +proof - + from assms have "\<exists>z>x. z < x + 1 \<and> ln (x + 1) - ln x = (x + 1 - x) * inverse z" + by (intro MVT2) (auto intro!: derivative_eq_intros simp: field_simps) + then obtain z where z: "z > x" "z < x + 1" "ln (x + 1) - ln x = inverse z" by auto + have "ln (x + 1) - ln x = inverse z" by fact + also from z(1,2) assms have "\<dots> < 1 / x" by (simp add: field_simps) + finally show ?thesis . +qed + +lemma ln_le_harm: "ln (real n + 1) \<le> (harm n :: real)" +proof (induction n) + fix n assume IH: "ln (real n + 1) \<le> harm n" + have "ln (real (Suc n) + 1) = ln (real n + 1) + (ln (real n + 2) - ln (real n + 1))" by simp + also have "(ln (real n + 2) - ln (real n + 1)) \<le> 1 / real (Suc n)" + using ln_diff_le_inverse[of "real n + 1"] by (simp add: add_ac) + also note IH + also have "harm n + 1 / real (Suc n) = harm (Suc n)" by (simp add: harm_Suc field_simps) + finally show "ln (real (Suc n) + 1) \<le> harm (Suc n)" by - simp +qed (simp_all add: harm_def) + + +subsection \<open>The Euler--Mascheroni constant\<close> + +text \<open> + The limit of the difference between the partial harmonic sum and the natural logarithm + (approximately 0.577216). This value occurs e.g. in the definition of the Gamma function. + \<close> +definition euler_mascheroni :: "'a :: real_normed_algebra_1" where + "euler_mascheroni = of_real (lim (\<lambda>n. harm n - ln (of_nat n)))" + +lemma of_real_euler_mascheroni [simp]: "of_real euler_mascheroni = euler_mascheroni" + by (simp add: euler_mascheroni_def) + +interpretation euler_mascheroni: antimono_fun_sum_integral_diff "\<lambda>x. inverse (x + 1)" + by unfold_locales (auto intro!: continuous_intros) + +lemma euler_mascheroni_sum_integral_diff_series: + "euler_mascheroni.sum_integral_diff_series n = harm (Suc n) - ln (of_nat (Suc n))" +proof - + have "harm (Suc n) = (\<Sum>k=0..n. inverse (of_nat k + 1) :: real)" unfolding harm_def + unfolding One_nat_def by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: add_ac) + moreover have "((\<lambda>x. inverse (x + 1) :: real) has_integral ln (of_nat n + 1) - ln (0 + 1)) + {0..of_nat n}" + by (intro fundamental_theorem_of_calculus) + (auto intro!: derivative_eq_intros simp: divide_inverse + has_field_derivative_iff_has_vector_derivative[symmetric]) + hence "integral {0..of_nat n} (\<lambda>x. inverse (x + 1) :: real) = ln (of_nat (Suc n))" + by (auto dest!: integral_unique) + ultimately show ?thesis + by (simp add: euler_mascheroni.sum_integral_diff_series_def atLeast0AtMost) +qed + +lemma euler_mascheroni_sequence_decreasing: + "m > 0 \<Longrightarrow> m \<le> n \<Longrightarrow> harm n - ln (of_nat n) \<le> harm m - ln (of_nat m :: real)" + by (cases m, simp, cases n, simp, hypsubst, + subst (1 2) euler_mascheroni_sum_integral_diff_series [symmetric], + rule euler_mascheroni.sum_integral_diff_series_antimono, simp) + +lemma euler_mascheroni_sequence_nonneg: + "n > 0 \<Longrightarrow> harm n - ln (of_nat n) \<ge> (0::real)" + by (cases n, simp, hypsubst, subst euler_mascheroni_sum_integral_diff_series [symmetric], + rule euler_mascheroni.sum_integral_diff_series_nonneg) + +lemma euler_mascheroni_convergent: "convergent (\<lambda>n. harm n - ln (of_nat n) :: real)" +proof - + have A: "(\<lambda>n. harm (Suc n) - ln (of_nat (Suc n))) = + euler_mascheroni.sum_integral_diff_series" + by (subst euler_mascheroni_sum_integral_diff_series [symmetric]) (rule refl) + have "convergent (\<lambda>n. harm (Suc n) - ln (of_nat (Suc n) :: real))" + by (subst A) (fact euler_mascheroni.sum_integral_diff_series_convergent) + thus ?thesis by (subst (asm) convergent_Suc_iff) +qed + +lemma euler_mascheroni_LIMSEQ: + "(\<lambda>n. harm n - ln (of_nat n) :: real) \<longlonglongrightarrow> euler_mascheroni" + unfolding euler_mascheroni_def + by (simp add: convergent_LIMSEQ_iff [symmetric] euler_mascheroni_convergent) + +lemma euler_mascheroni_LIMSEQ_of_real: + "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow> + (euler_mascheroni :: 'a :: {real_normed_algebra_1, topological_space})" +proof - + have "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow> (of_real (euler_mascheroni) :: 'a)" + by (intro tendsto_of_real euler_mascheroni_LIMSEQ) + thus ?thesis by simp +qed + +lemma euler_mascheroni_sum: + "(\<lambda>n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real) + sums euler_mascheroni" + using sums_add[OF telescope_sums[OF LIMSEQ_Suc[OF euler_mascheroni_LIMSEQ]] + telescope_sums'[OF LIMSEQ_inverse_real_of_nat]] + by (simp_all add: harm_def algebra_simps) + +lemma alternating_harmonic_series_sums: "(\<lambda>k. (-1)^k / real_of_nat (Suc k)) sums ln 2" +proof - + let ?f = "\<lambda>n. harm n - ln (real_of_nat n)" + let ?g = "\<lambda>n. if even n then 0 else (2::real)" + let ?em = "\<lambda>n. harm n - ln (real_of_nat n)" + have "eventually (\<lambda>n. ?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))) at_top" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + have "(\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k)) = + (\<Sum>k<2*n. ((-1)^k + ?g k) / of_nat (Suc k)) - (\<Sum>k<2*n. ?g k / of_nat (Suc k))" + by (simp add: setsum.distrib algebra_simps divide_inverse) + also have "(\<Sum>k<2*n. ((-1)^k + ?g k) / real_of_nat (Suc k)) = harm (2*n)" + unfolding harm_altdef by (intro setsum.cong) (auto simp: field_simps) + also have "(\<Sum>k<2*n. ?g k / real_of_nat (Suc k)) = (\<Sum>k|k<2*n \<and> odd k. ?g k / of_nat (Suc k))" + by (intro setsum.mono_neutral_right) auto + also have "\<dots> = (\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k)))" + by (intro setsum.cong) auto + also have "(\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k))) = harm n" + unfolding harm_altdef + by (intro setsum.reindex_cong[of "\<lambda>n. 2*n+1"]) (auto simp: inj_on_def field_simps elim!: oddE) + also have "harm (2*n) - harm n = ?em (2*n) - ?em n + ln 2" using n + by (simp_all add: algebra_simps ln_mult) + finally show "?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))" .. + qed + moreover have "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real)) + \<longlonglongrightarrow> euler_mascheroni - euler_mascheroni + ln 2" + by (intro tendsto_intros euler_mascheroni_LIMSEQ filterlim_compose[OF euler_mascheroni_LIMSEQ] + filterlim_subseq) (auto simp: subseq_def) + hence "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real)) \<longlonglongrightarrow> ln 2" by simp + ultimately have "(\<lambda>n. (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2" + by (rule Lim_transform_eventually) + + moreover have "summable (\<lambda>k. (-1)^k * inverse (real_of_nat (Suc k)))" + using LIMSEQ_inverse_real_of_nat + by (intro summable_Leibniz(1) decseq_imp_monoseq decseq_SucI) simp_all + hence A: "(\<lambda>n. \<Sum>k<n. (-1)^k / real_of_nat (Suc k)) \<longlonglongrightarrow> (\<Sum>k. (-1)^k / real_of_nat (Suc k))" + by (simp add: summable_sums_iff divide_inverse sums_def) + from filterlim_compose[OF this filterlim_subseq[of "op * (2::nat)"]] + have "(\<lambda>n. \<Sum>k<2*n. (-1)^k / real_of_nat (Suc k)) \<longlonglongrightarrow> (\<Sum>k. (-1)^k / real_of_nat (Suc k))" + by (simp add: subseq_def) + ultimately have "(\<Sum>k. (- 1) ^ k / real_of_nat (Suc k)) = ln 2" by (intro LIMSEQ_unique) + with A show ?thesis by (simp add: sums_def) +qed + +lemma alternating_harmonic_series_sums': + "(\<lambda>k. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2))) sums ln 2" +unfolding sums_def +proof (rule Lim_transform_eventually) + show "(\<lambda>n. \<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2" + using alternating_harmonic_series_sums unfolding sums_def + by (rule filterlim_compose) (rule mult_nat_left_at_top, simp) + show "eventually (\<lambda>n. (\<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) = + (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))) sequentially" + proof (intro always_eventually allI) + fix n :: nat + show "(\<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) = + (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))" + by (induction n) (simp_all add: inverse_eq_divide) + qed +qed + + +subsection \<open>Bounds on the Euler--Mascheroni constant\<close> + +(* TODO: Move? *) +lemma ln_inverse_approx_le: + assumes "(x::real) > 0" "a > 0" + shows "ln (x + a) - ln x \<le> a * (inverse x + inverse (x + a))/2" (is "_ \<le> ?A") +proof - + define f' where "f' = (inverse (x + a) - inverse x)/a" + have f'_nonpos: "f' \<le> 0" using assms by (simp add: f'_def divide_simps) + let ?f = "\<lambda>t. (t - x) * f' + inverse x" + let ?F = "\<lambda>t. (t - x)^2 * f' / 2 + t * inverse x" + have diff: "\<forall>t\<in>{x..x+a}. (?F has_vector_derivative ?f t) + (at t within {x..x+a})" using assms + by (auto intro!: derivative_eq_intros + simp: has_field_derivative_iff_has_vector_derivative[symmetric]) + from assms have "(?f has_integral (?F (x+a) - ?F x)) {x..x+a}" + by (intro fundamental_theorem_of_calculus[OF _ diff]) + (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] field_simps + intro!: derivative_eq_intros) + also have "?F (x+a) - ?F x = (a*2 + f'*a\<^sup>2*x) / (2*x)" using assms by (simp add: field_simps) + also have "f'*a^2 = - (a^2) / (x*(x + a))" using assms + by (simp add: divide_simps f'_def power2_eq_square) + also have "(a*2 + - a\<^sup>2/(x*(x+a))*x) / (2*x) = ?A" using assms + by (simp add: divide_simps power2_eq_square) (simp add: algebra_simps) + finally have int1: "((\<lambda>t. (t - x) * f' + inverse x) has_integral ?A) {x..x + a}" . + + from assms have int2: "(inverse has_integral (ln (x + a) - ln x)) {x..x+a}" + by (intro fundamental_theorem_of_calculus) + (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps + intro!: derivative_eq_intros) + hence "ln (x + a) - ln x = integral {x..x+a} inverse" by (simp add: integral_unique) + also have ineq: "\<forall>xa\<in>{x..x + a}. inverse xa \<le> (xa - x) * f' + inverse x" + proof + fix t assume t': "t \<in> {x..x+a}" + with assms have t: "0 \<le> (t - x) / a" "(t - x) / a \<le> 1" by simp_all + have "inverse t = inverse ((1 - (t - x) / a) *\<^sub>R x + ((t - x) / a) *\<^sub>R (x + a))" (is "_ = ?A") + using assms t' by (simp add: field_simps) + also from assms have "convex_on {x..x+a} inverse" by (intro convex_on_inverse) auto + from convex_onD_Icc[OF this _ t] assms + have "?A \<le> (1 - (t - x) / a) * inverse x + (t - x) / a * inverse (x + a)" by simp + also have "\<dots> = (t - x) * f' + inverse x" using assms + by (simp add: f'_def divide_simps) (simp add: f'_def field_simps) + finally show "inverse t \<le> (t - x) * f' + inverse x" . + qed + hence "integral {x..x+a} inverse \<le> integral {x..x+a} ?f" using f'_nonpos assms + by (intro integral_le has_integral_integrable[OF int1] has_integral_integrable[OF int2] ineq) + also have "\<dots> = ?A" using int1 by (rule integral_unique) + finally show ?thesis . +qed + +lemma ln_inverse_approx_ge: + assumes "(x::real) > 0" "x < y" + shows "ln y - ln x \<ge> 2 * (y - x) / (x + y)" (is "_ \<ge> ?A") +proof - + define m where "m = (x+y)/2" + define f' where "f' = -inverse (m^2)" + from assms have m: "m > 0" by (simp add: m_def) + let ?F = "\<lambda>t. (t - m)^2 * f' / 2 + t / m" + from assms have "((\<lambda>t. (t - m) * f' + inverse m) has_integral (?F y - ?F x)) {x..y}" + by (intro fundamental_theorem_of_calculus) + (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps + intro!: derivative_eq_intros) + also from m have "?F y - ?F x = ((y - m)^2 - (x - m)^2) * f' / 2 + (y - x) / m" + by (simp add: field_simps) + also have "((y - m)^2 - (x - m)^2) = 0" by (simp add: m_def power2_eq_square field_simps) + also have "0 * f' / 2 + (y - x) / m = ?A" by (simp add: m_def) + finally have int1: "((\<lambda>t. (t - m) * f' + inverse m) has_integral ?A) {x..y}" . + + from assms have int2: "(inverse has_integral (ln y - ln x)) {x..y}" + by (intro fundamental_theorem_of_calculus) + (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps + intro!: derivative_eq_intros) + hence "ln y - ln x = integral {x..y} inverse" by (simp add: integral_unique) + also have ineq: "\<forall>xa\<in>{x..y}. inverse xa \<ge> (xa - m) * f' + inverse m" + proof + fix t assume t: "t \<in> {x..y}" + from t assms have "inverse t - inverse m \<ge> f' * (t - m)" + by (intro convex_on_imp_above_tangent[of "{0<..}"] convex_on_inverse) + (auto simp: m_def interior_open f'_def power2_eq_square intro!: derivative_eq_intros) + thus "(t - m) * f' + inverse m \<le> inverse t" by (simp add: algebra_simps) + qed + hence "integral {x..y} inverse \<ge> integral {x..y} (\<lambda>t. (t - m) * f' + inverse m)" + using int1 int2 by (intro integral_le has_integral_integrable) + also have "integral {x..y} (\<lambda>t. (t - m) * f' + inverse m) = ?A" + using integral_unique[OF int1] by simp + finally show ?thesis . +qed + + +lemma euler_mascheroni_lower: + "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))" + and euler_mascheroni_upper: + "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))" +proof - + define D :: "_ \<Rightarrow> real" + where "D n = inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2))" for n + let ?g = "\<lambda>n. ln (of_nat (n+2)) - ln (of_nat (n+1)) - inverse (of_nat (n+1)) :: real" + define inv where [abs_def]: "inv n = inverse (real_of_nat n)" for n + fix n :: nat + note summable = sums_summable[OF euler_mascheroni_sum, folded D_def] + have sums: "(\<lambda>k. (inv (Suc (k + (n+1))) - inv (Suc (Suc k + (n+1))))/2) sums ((inv (Suc (0 + (n+1))) - 0)/2)" + unfolding inv_def + by (intro sums_divide telescope_sums' LIMSEQ_ignore_initial_segment LIMSEQ_inverse_real_of_nat) + have sums': "(\<lambda>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) sums ((inv (Suc (0 + n)) - 0)/2)" + unfolding inv_def + by (intro sums_divide telescope_sums' LIMSEQ_ignore_initial_segment LIMSEQ_inverse_real_of_nat) + from euler_mascheroni_sum have "euler_mascheroni = (\<Sum>k. D k)" + by (simp add: sums_iff D_def) + also have "\<dots> = (\<Sum>k. D (k + Suc n)) + (\<Sum>k\<le>n. D k)" + by (subst suminf_split_initial_segment[OF summable, of "Suc n"], subst lessThan_Suc_atMost) simp + finally have sum: "(\<Sum>k\<le>n. D k) - euler_mascheroni = -(\<Sum>k. D (k + Suc n))" by simp + + note sum + also have "\<dots> \<le> -(\<Sum>k. (inv (k + Suc n + 1) - inv (k + Suc n + 2)) / 2)" + proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable]) + fix k' :: nat + define k where "k = k' + Suc n" + hence k: "k > 0" by (simp add: k_def) + have "real_of_nat (k+1) > 0" by (simp add: k_def) + with ln_inverse_approx_le[OF this zero_less_one] + have "ln (of_nat k + 2) - ln (of_nat k + 1) \<le> (inv (k+1) + inv (k+2))/2" + by (simp add: inv_def add_ac) + hence "(inv (k+1) - inv (k+2))/2 \<le> inv (k+1) + ln (of_nat (k+1)) - ln (of_nat (k+2))" + by (simp add: field_simps) + also have "\<dots> = D k" unfolding D_def inv_def .. + finally show "D (k' + Suc n) \<ge> (inv (k' + Suc n + 1) - inv (k' + Suc n + 2)) / 2" + by (simp add: k_def) + from sums_summable[OF sums] + show "summable (\<lambda>k. (inv (k + Suc n + 1) - inv (k + Suc n + 2))/2)" by simp + qed + also from sums have "\<dots> = -inv (n+2) / 2" by (simp add: sums_iff) + finally have "euler_mascheroni \<ge> (\<Sum>k\<le>n. D k) + 1 / (of_nat (2 * (n+2)))" + by (simp add: inv_def field_simps) + also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))" + unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: setsum.distrib setsum_subtractf) + also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))" + by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all + finally show "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))" + by simp + + note sum + also have "-(\<Sum>k. D (k + Suc n)) \<ge> -(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" + proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable]) + fix k' :: nat + define k where "k = k' + Suc n" + hence k: "k > 0" by (simp add: k_def) + have "real_of_nat (k+1) > 0" by (simp add: k_def) + from ln_inverse_approx_ge[of "of_nat k + 1" "of_nat k + 2"] + have "2 / (2 * real_of_nat k + 3) \<le> ln (of_nat (k+2)) - ln (real_of_nat (k+1))" + by (simp add: add_ac) + hence "D k \<le> 1 / real_of_nat (k+1) - 2 / (2 * real_of_nat k + 3)" + by (simp add: D_def inverse_eq_divide inv_def) + also have "\<dots> = inv ((k+1)*(2*k+3))" unfolding inv_def by (simp add: field_simps) + also have "\<dots> \<le> inv (2*k*(k+1))" unfolding inv_def using k + by (intro le_imp_inverse_le) + (simp add: algebra_simps, simp del: of_nat_add) + also have "\<dots> = (inv k - inv (k+1))/2" unfolding inv_def using k + by (simp add: divide_simps del: of_nat_mult) (simp add: algebra_simps) + finally show "D k \<le> (inv (Suc (k' + n)) - inv (Suc (Suc k' + n)))/2" unfolding k_def by simp + next + from sums_summable[OF sums'] + show "summable (\<lambda>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" by simp + qed + also from sums' have "(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) = inv (n+1)/2" + by (simp add: sums_iff) + finally have "euler_mascheroni \<le> (\<Sum>k\<le>n. D k) + 1 / of_nat (2 * (n+1))" + by (simp add: inv_def field_simps) + also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))" + unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: setsum.distrib setsum_subtractf) + also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))" + by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all + finally show "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))" + by simp +qed + +lemma euler_mascheroni_pos: "euler_mascheroni > (0::real)" + using euler_mascheroni_lower[of 0] ln_2_less_1 by (simp add: harm_def) + +context +begin + +private lemma ln_approx_aux: + fixes n :: nat and x :: real + defines "y \<equiv> (x-1)/(x+1)" + assumes x: "x > 0" "x \<noteq> 1" + shows "inverse (2*y^(2*n+1)) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in> + {0..(1 / (1 - y^2) / of_nat (2*n+1))}" +proof - + from x have norm_y: "norm y < 1" unfolding y_def by simp + from power_strict_mono[OF this, of 2] have norm_y': "norm y^2 < 1" by simp + + let ?f = "\<lambda>k. 2 * y ^ (2*k+1) / of_nat (2*k+1)" + note sums = ln_series_quadratic[OF x(1)] + define c where "c = inverse (2*y^(2*n+1))" + let ?d = "c * (ln x - (\<Sum>k<n. ?f k))" + have "\<forall>k. y\<^sup>2^k / of_nat (2*(k+n)+1) \<le> y\<^sup>2 ^ k / of_nat (2*n+1)" + by (intro allI divide_left_mono mult_right_mono mult_pos_pos zero_le_power[of "y^2"]) simp_all + moreover { + have "(\<lambda>k. ?f (k + n)) sums (ln x - (\<Sum>k<n. ?f k))" + using sums_split_initial_segment[OF sums] by (simp add: y_def) + hence "(\<lambda>k. c * ?f (k + n)) sums ?d" by (rule sums_mult) + also have "(\<lambda>k. c * (2*y^(2*(k+n)+1) / of_nat (2*(k+n)+1))) = + (\<lambda>k. (c * (2*y^(2*n+1))) * ((y^2)^k / of_nat (2*(k+n)+1)))" + by (simp only: ring_distribs power_add power_mult) (simp add: mult_ac) + also from x have "c * (2*y^(2*n+1)) = 1" by (simp add: c_def y_def) + finally have "(\<lambda>k. (y^2)^k / of_nat (2*(k+n)+1)) sums ?d" by simp + } note sums' = this + moreover from norm_y' have "(\<lambda>k. (y^2)^k / of_nat (2*n+1)) sums (1 / (1 - y^2) / of_nat (2*n+1))" + by (intro sums_divide geometric_sums) (simp_all add: norm_power) + ultimately have "?d \<le> (1 / (1 - y^2) / of_nat (2*n+1))" by (rule sums_le) + moreover have "c * (ln x - (\<Sum>k<n. 2 * y ^ (2 * k + 1) / real_of_nat (2 * k + 1))) \<ge> 0" + by (intro sums_le[OF _ sums_zero sums']) simp_all + ultimately show ?thesis unfolding c_def by simp +qed + +lemma + fixes n :: nat and x :: real + defines "y \<equiv> (x-1)/(x+1)" + defines "approx \<equiv> (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))" + defines "d \<equiv> y^(2*n+1) / (1 - y^2) / of_nat (2*n+1)" + assumes x: "x > 1" + shows ln_approx_bounds: "ln x \<in> {approx..approx + 2*d}" + and ln_approx_abs: "abs (ln x - (approx + d)) \<le> d" +proof - + define c where "c = 2*y^(2*n+1)" + from x have c_pos: "c > 0" unfolding c_def y_def + by (intro mult_pos_pos zero_less_power) simp_all + have A: "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in> + {0.. (1 / (1 - y^2) / of_nat (2*n+1))}" using assms unfolding y_def c_def + by (intro ln_approx_aux) simp_all + hence "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1)/of_nat (2*k+1))) \<le> (1 / (1-y^2) / of_nat (2*n+1))" + by simp + hence "(ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) / c \<le> (1 / (1 - y^2) / of_nat (2*n+1))" + by (auto simp add: divide_simps) + with c_pos have "ln x \<le> c / (1 - y^2) / of_nat (2*n+1) + approx" + by (subst (asm) pos_divide_le_eq) (simp_all add: mult_ac approx_def) + moreover { + from A c_pos have "0 \<le> c * (inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))))" + by (intro mult_nonneg_nonneg[of c]) simp_all + also have "\<dots> = (c * inverse c) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1)))" + by (simp add: mult_ac) + also from c_pos have "c * inverse c = 1" by simp + finally have "ln x \<ge> approx" by (simp add: approx_def) + } + ultimately show "ln x \<in> {approx..approx + 2*d}" by (simp add: c_def d_def) + thus "abs (ln x - (approx + d)) \<le> d" by auto +qed + +end + +lemma euler_mascheroni_bounds: + fixes n :: nat assumes "n \<ge> 1" defines "t \<equiv> harm n - ln (of_nat (Suc n)) :: real" + shows "euler_mascheroni \<in> {t + inverse (of_nat (2*(n+1)))..t + inverse (of_nat (2*n))}" + using assms euler_mascheroni_upper[of "n-1"] euler_mascheroni_lower[of "n-1"] + unfolding t_def by (cases n) (simp_all add: harm_Suc t_def inverse_eq_divide) + +lemma euler_mascheroni_bounds': + fixes n :: nat assumes "n \<ge> 1" "ln (real_of_nat (Suc n)) \<in> {l<..<u}" + shows "euler_mascheroni \<in> + {harm n - u + inverse (of_nat (2*(n+1)))<..<harm n - l + inverse (of_nat (2*n))}" + using euler_mascheroni_bounds[OF assms(1)] assms(2) by auto + + +text \<open> + Approximation of @{term "ln 2"}. The lower bound is accurate to about 0.03; the upper + bound is accurate to about 0.0015. +\<close> +lemma ln2_ge_two_thirds: "2/3 \<le> ln (2::real)" + and ln2_le_25_over_36: "ln (2::real) \<le> 25/36" + using ln_approx_bounds[of 2 1, simplified, simplified eval_nat_numeral, simplified] by simp_all + + +text \<open> + Approximation of the Euler--Mascheroni constant. The lower bound is accurate to about 0.0015; + the upper bound is accurate to about 0.015. +\<close> +lemma euler_mascheroni_gt_19_over_33: "(euler_mascheroni :: real) > 19/33" (is ?th1) + and euler_mascheroni_less_13_over_22: "(euler_mascheroni :: real) < 13/22" (is ?th2) +proof - + have "ln (real (Suc 7)) = 3 * ln 2" by (simp add: ln_powr [symmetric] powr_numeral) + also from ln_approx_bounds[of 2 3] have "\<dots> \<in> {3*307/443<..<3*4615/6658}" + by (simp add: eval_nat_numeral) + finally have "ln (real (Suc 7)) \<in> \<dots>" . + from euler_mascheroni_bounds'[OF _ this] have "?th1 \<and> ?th2" by (simp_all add: harm_expand) + thus ?th1 ?th2 by blast+ +qed + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Aug 08 14:13:14 2016 +0200 @@ -0,0 +1,12512 @@ +(* Author: John Harrison + Author: Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP +*) + +section \<open>Henstock-Kurzweil gauge integration in many dimensions.\<close> + +theory Henstock_Kurzweil_Integration +imports + Derivative + Uniform_Limit + "~~/src/HOL/Library/Indicator_Function" +begin + +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 \<open>Sundries\<close> + +lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto +lemma conjunctD3: assumes "a \<and> b \<and> c" shows a b c using assms by auto +lemma conjunctD4: assumes "a \<and> b \<and> c \<and> d" shows a b c d using assms by auto + +declare norm_triangle_ineq4[intro] + +lemma simple_image: "{f x |x . x \<in> s} = f ` s" + by blast + +lemma linear_simps: + assumes "bounded_linear f" + shows + "f (a + b) = f a + f b" + "f (a - b) = f a - f b" + "f 0 = 0" + "f (- a) = - f a" + "f (s *\<^sub>R v) = s *\<^sub>R (f v)" +proof - + interpret f: bounded_linear f by fact + show "f (a + b) = f a + f b" by (rule f.add) + show "f (a - b) = f a - f b" by (rule f.diff) + show "f 0 = 0" by (rule f.zero) + show "f (- a) = - f a" by (rule f.minus) + show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scaleR) +qed + +lemma bounded_linearI: + assumes "\<And>x y. f (x + y) = f x + f y" + and "\<And>r x. f (r *\<^sub>R x) = r *\<^sub>R f x" + and "\<And>x. norm (f x) \<le> norm x * K" + shows "bounded_linear f" + using assms by (rule bounded_linear_intro) (* FIXME: duplicate *) + +lemma bounded_linear_component [intro]: "bounded_linear (\<lambda>x::'a::euclidean_space. x \<bullet> k)" + by (rule bounded_linear_inner_left) + +lemma transitive_stepwise_lt_eq: + assumes "(\<And>x y z::nat. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z)" + shows "((\<forall>m. \<forall>n>m. R m n) \<longleftrightarrow> (\<forall>n. R n (Suc n)))" + (is "?l = ?r") +proof safe + assume ?r + fix n m :: nat + assume "m < n" + then show "R m n" + proof (induct n arbitrary: m) + case 0 + then show ?case by auto + next + case (Suc n) + show ?case + proof (cases "m < n") + case True + show ?thesis + apply (rule assms[OF Suc(1)[OF True]]) + using \<open>?r\<close> + apply auto + done + next + case False + then have "m = n" + using Suc(2) by auto + then show ?thesis + using \<open>?r\<close> by auto + qed + qed +qed auto + +lemma transitive_stepwise_gt: + assumes "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z" "\<And>n. R n (Suc n)" + shows "\<forall>n>m. R m n" +proof - + have "\<forall>m. \<forall>n>m. R m n" + apply (subst transitive_stepwise_lt_eq) + apply (blast intro: assms)+ + done + then show ?thesis by auto +qed + +lemma transitive_stepwise_le_eq: + assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z" + shows "(\<forall>m. \<forall>n\<ge>m. R m n) \<longleftrightarrow> (\<forall>n. R n (Suc n))" + (is "?l = ?r") +proof safe + assume ?r + fix m n :: nat + assume "m \<le> n" + then show "R m n" + proof (induct n arbitrary: m) + case 0 + with assms show ?case by auto + next + case (Suc n) + show ?case + proof (cases "m \<le> n") + case True + with Suc.hyps \<open>\<forall>n. R n (Suc n)\<close> assms show ?thesis + by blast + next + case False + then have "m = Suc n" + using Suc(2) by auto + then show ?thesis + using assms(1) by auto + qed + qed +qed auto + +lemma transitive_stepwise_le: + assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z" + and "\<And>n. R n (Suc n)" + shows "\<forall>n\<ge>m. R m n" +proof - + have "\<forall>m. \<forall>n\<ge>m. R m n" + apply (subst transitive_stepwise_le_eq) + apply (blast intro: assms)+ + done + then show ?thesis by auto +qed + + +subsection \<open>Some useful lemmas about intervals.\<close> + +lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" + using nonempty_Basis + by (fastforce simp add: set_eq_iff mem_box) + +lemma interior_subset_union_intervals: + assumes "i = cbox a b" + and "j = cbox c d" + and "interior j \<noteq> {}" + and "i \<subseteq> j \<union> s" + and "interior i \<inter> interior j = {}" + shows "interior i \<subseteq> interior s" +proof - + have "box a b \<inter> cbox c d = {}" + using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5) + unfolding assms(1,2) interior_cbox by auto + moreover + have "box a b \<subseteq> cbox c d \<union> s" + apply (rule order_trans,rule box_subset_cbox) + using assms(4) unfolding assms(1,2) + apply auto + done + 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: "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a b" "\<And>s. s \<in> f \<Longrightarrow> interior s \<subseteq> t" + and t: "closed t" + shows "interior (\<Union>f) \<subseteq> t" +proof - + have [simp]: "s \<in> f \<Longrightarrow> closed s" for s + using f by auto + define E where "E = {s\<in>f. interior s = {}}" + then have "finite E" "E \<subseteq> {s\<in>f. interior s = {}}" + using \<open>finite f\<close> by auto + then have "interior (\<Union>f) = interior (\<Union>(f - E))" + proof (induction E rule: finite_subset_induct') + case (insert s f') + have "interior (\<Union>(f - insert s f') \<union> s) = interior (\<Union>(f - insert s f'))" + using insert.hyps \<open>finite f\<close> by (intro interior_closed_Un_empty_interior) auto + also have "\<Union>(f - insert s f') \<union> s = \<Union>(f - f')" + using insert.hyps by auto + finally show ?case + by (simp add: insert.IH) + qed simp + also have "\<dots> \<subseteq> \<Union>(f - E)" + by (rule interior_subset) + also have "\<dots> \<subseteq> t" + proof (rule Union_least) + fix s assume "s \<in> f - E" + with f[of s] obtain a b where s: "s \<in> f" "s = cbox a b" "box a b \<noteq> {}" + by (fastforce simp: E_def) + have "closure (interior s) \<subseteq> closure t" + by (intro closure_mono f \<open>s \<in> f\<close>) + with s \<open>closed t\<close> show "s \<subseteq> t" + by (simp add: closure_box) + qed + finally show ?thesis . +qed + +lemma inter_interior_unions_intervals: + "finite f \<Longrightarrow> open s \<Longrightarrow> \<forall>t\<in>f. \<exists>a b. t = cbox a b \<Longrightarrow> \<forall>t\<in>f. s \<inter> (interior t) = {} \<Longrightarrow> s \<inter> interior (\<Union>f) = {}" + using interior_Union_subset_cbox[of f "UNIV - s"] by auto + +subsection \<open>Bounds on intervals where they exist.\<close> + +definition interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a" + where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)" + +definition interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a" + where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)" + +lemma interval_upperbound[simp]: + "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> + interval_upperbound (cbox a b) = (b::'a::euclidean_space)" + unfolding interval_upperbound_def euclidean_representation_setsum cbox_def + by (safe intro!: cSup_eq) auto + +lemma interval_lowerbound[simp]: + "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> + interval_lowerbound (cbox a b) = (a::'a::euclidean_space)" + unfolding interval_lowerbound_def euclidean_representation_setsum 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 \<noteq> {}" + 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 \<noteq> {}" and "B \<noteq> {}" + shows "interval_upperbound (A \<times> B) = (interval_upperbound A, interval_upperbound B)" +proof- + from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp + have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:A. x \<bullet> i) *\<^sub>R i)" + by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0) + moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp + have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:B. x \<bullet> i) *\<^sub>R i)" + by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0) + ultimately show ?thesis unfolding interval_upperbound_def + by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod) +qed + +lemma interval_lowerbound_Times: + assumes "A \<noteq> {}" and "B \<noteq> {}" + shows "interval_lowerbound (A \<times> B) = (interval_lowerbound A, interval_lowerbound B)" +proof- + from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp + have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:A. x \<bullet> i) *\<^sub>R i)" + by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0) + moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp + have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:B. x \<bullet> i) *\<^sub>R i)" + by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0) + ultimately show ?thesis unfolding interval_lowerbound_def + by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod) +qed + +subsection \<open>Content (length, area, volume...) of an interval.\<close> + +definition "content (s::('a::euclidean_space) set) = + (if s = {} then 0 else (\<Prod>i\<in>Basis. (interval_upperbound s)\<bullet>i - (interval_lowerbound s)\<bullet>i))" + +lemma interval_not_empty: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> cbox a b \<noteq> {}" + unfolding box_eq_empty unfolding not_ex not_less by auto + +lemma content_cbox: + fixes a :: "'a::euclidean_space" + assumes "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i" + shows "content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)" + using interval_not_empty[OF assms] + unfolding content_def + by auto + +lemma content_cbox': + fixes a :: "'a::euclidean_space" + assumes "cbox a b \<noteq> {}" + shows "content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)" + using assms box_ne_empty(1) content_cbox by blast + +lemma content_real: "a \<le> b \<Longrightarrow> content {a..b} = b - a" + by (auto simp: interval_upperbound_def interval_lowerbound_def content_def) + +lemma abs_eq_content: "\<bar>y - x\<bar> = (if x\<le>y then content {x .. y} else content {y..x})" + by (auto simp: content_real) + +lemma content_singleton[simp]: "content {a} = 0" +proof - + have "content (cbox a a) = 0" + by (subst content_cbox) (auto simp: ex_in_conv) + then show ?thesis by (simp add: cbox_sing) +qed + +lemma content_unit[iff]: "content(cbox 0 (One::'a::euclidean_space)) = 1" + proof - + have *: "\<forall>i\<in>Basis. (0::'a)\<bullet>i \<le> (One::'a)\<bullet>i" + by auto + have "0 \<in> cbox 0 (One::'a)" + unfolding mem_box by auto + then show ?thesis + unfolding content_def interval_bounds[OF *] using setprod.neutral_const by auto + qed + +lemma content_pos_le[intro]: + fixes a::"'a::euclidean_space" + shows "0 \<le> content (cbox a b)" +proof (cases "cbox a b = {}") + case False + then have *: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" + unfolding box_ne_empty . + have "0 \<le> (\<Prod>i\<in>Basis. interval_upperbound (cbox a b) \<bullet> i - interval_lowerbound (cbox a b) \<bullet> i)" + apply (rule setprod_nonneg) + unfolding interval_bounds[OF *] + using * + apply auto + done + also have "\<dots> = content (cbox a b)" using False by (simp add: content_def) + finally show ?thesis . +qed (simp add: content_def) + +corollary content_nonneg [simp]: + fixes a::"'a::euclidean_space" + shows "~ content (cbox a b) < 0" +using not_le by blast + +lemma content_pos_lt: + fixes a :: "'a::euclidean_space" + assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" + shows "0 < content (cbox a b)" + using assms + by (auto simp: content_def box_eq_empty intro!: setprod_pos) + +lemma content_eq_0: + "content (cbox a b) = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i)" + by (auto simp: content_def box_eq_empty intro!: setprod_pos bexI) + +lemma cond_cases: "(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)" + by auto + +lemma content_cbox_cases: + "content (cbox a (b::'a::euclidean_space)) = + (if \<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i then setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis else 0)" + by (auto simp: not_le content_eq_0 intro: less_imp_le content_cbox) + +lemma content_eq_0_interior: "content (cbox a b) = 0 \<longleftrightarrow> 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)) \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)" +proof (rule iffI) + assume "0 < content (cbox a b)" + then have "content (cbox a b) \<noteq> 0" by auto + then show "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" + unfolding content_eq_0 not_ex not_le by fastforce +next + assume "\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i" + then show "0 < content (cbox a b)" + by (metis content_pos_lt) +qed + +lemma content_empty [simp]: "content {} = 0" + unfolding content_def by auto + +lemma content_real_if [simp]: "content {a..b} = (if a \<le> b then b - a else 0)" + by (simp add: content_real) + +lemma content_subset: + assumes "cbox a b \<subseteq> cbox c d" + shows "content (cbox a b) \<le> content (cbox c d)" +proof (cases "cbox a b = {}") + case True + then show ?thesis + using content_pos_le[of c d] by auto +next + case False + then have ab_ne: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" + unfolding box_ne_empty by auto + then have ab_ab: "a\<in>cbox a b" "b\<in>cbox a b" + unfolding mem_box by auto + have "cbox c d \<noteq> {}" using assms False by auto + then have cd_ne: "\<forall>i\<in>Basis. c \<bullet> i \<le> d \<bullet> i" + using assms unfolding box_ne_empty by auto + have "\<And>i. i \<in> Basis \<Longrightarrow> 0 \<le> b \<bullet> i - a \<bullet> i" + using ab_ne by auto + moreover + have "\<And>i. i \<in> Basis \<Longrightarrow> b \<bullet> i - a \<bullet> i \<le> d \<bullet> i - c \<bullet> i" + using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2)] + assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1)] + by (metis diff_mono) + ultimately show ?thesis + unfolding content_def interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] + by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF \<open>cbox c d \<noteq> {}\<close>]) +qed + +lemma content_lt_nz: "0 < content (cbox a b) \<longleftrightarrow> content (cbox a b) \<noteq> 0" + unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce + +lemma content_times[simp]: "content (A \<times> B) = content A * content B" +proof (cases "A \<times> B = {}") + let ?ub1 = "interval_upperbound" and ?lb1 = "interval_lowerbound" + let ?ub2 = "interval_upperbound" and ?lb2 = "interval_lowerbound" + assume nonempty: "A \<times> B \<noteq> {}" + hence "content (A \<times> B) = (\<Prod>i\<in>Basis. (?ub1 A, ?ub2 B) \<bullet> i - (?lb1 A, ?lb2 B) \<bullet> i)" + unfolding content_def by (simp add: interval_upperbound_Times interval_lowerbound_Times) + also have "... = content A * content B" unfolding content_def using nonempty + apply (subst Basis_prod_def, subst setprod.union_disjoint, force, force, force, simp) + apply (subst (1 2) setprod.reindex, auto intro: inj_onI) + done + finally show ?thesis . +qed (auto simp: content_def) + +lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" + by (simp add: cbox_Pair_eq) + +lemma content_cbox_pair_eq0_D: + "content (cbox (a,c) (b,d)) = 0 \<Longrightarrow> content (cbox a b) = 0 \<or> content (cbox c d) = 0" + by (simp add: content_Pair) + +lemma content_eq_0_gen: + fixes s :: "'a::euclidean_space set" + assumes "bounded s" + shows "content s = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. \<exists>v. \<forall>x \<in> s. x \<bullet> i = v)" (is "_ = ?rhs") +proof safe + assume "content s = 0" then show ?rhs + apply (clarsimp simp: ex_in_conv content_def split: if_split_asm) + apply (rule_tac x=a in bexI) + apply (rule_tac x="interval_lowerbound s \<bullet> a" in exI) + apply (clarsimp simp: interval_upperbound_def interval_lowerbound_def) + apply (drule cSUP_eq_cINF_D) + apply (auto simp: bounded_inner_imp_bdd_above [OF assms] bounded_inner_imp_bdd_below [OF assms]) + done +next + fix i a + assume "i \<in> Basis" "\<forall>x\<in>s. x \<bullet> i = a" + then show "content s = 0" + apply (clarsimp simp: content_def) + apply (rule_tac x=i in bexI) + apply (auto simp: interval_upperbound_def interval_lowerbound_def) + done +qed + +lemma content_0_subset_gen: + fixes a :: "'a::euclidean_space" + assumes "content t = 0" "s \<subseteq> t" "bounded t" shows "content s = 0" +proof - + have "bounded s" + using assms by (metis bounded_subset) + then show ?thesis + using assms + by (auto simp: content_eq_0_gen) +qed + +lemma content_0_subset: "\<lbrakk>content(cbox a b) = 0; s \<subseteq> cbox a b\<rbrakk> \<Longrightarrow> content s = 0" + by (simp add: content_0_subset_gen bounded_cbox) + + +lemma interval_split: + fixes a :: "'a::euclidean_space" + assumes "k \<in> Basis" + shows + "cbox a b \<inter> {x. x\<bullet>k \<le> c} = cbox a (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i) *\<^sub>R i)" + "cbox a b \<inter> {x. x\<bullet>k \<ge> c} = cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i) *\<^sub>R i) b" + apply (rule_tac[!] set_eqI) + unfolding Int_iff mem_box mem_Collect_eq + using assms + apply auto + done + +lemma content_split: + fixes a :: "'a::euclidean_space" + assumes "k \<in> Basis" + shows "content (cbox a b) = content(cbox a b \<inter> {x. x\<bullet>k \<le> c}) + content(cbox a b \<inter> {x. x\<bullet>k \<ge> c})" +proof cases + note simps = interval_split[OF assms] content_cbox_cases + have *: "Basis = insert k (Basis - {k})" "\<And>x. finite (Basis-{x})" "\<And>x. x\<notin>Basis-{x}" + using assms by auto + have *: "\<And>X Y Z. (\<Prod>i\<in>Basis. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>Basis-{k}. Z i (Y i))" + "(\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i) = (\<Prod>i\<in>Basis-{k}. b\<bullet>i - a\<bullet>i) * (b\<bullet>k - a\<bullet>k)" + apply (subst *(1)) + defer + apply (subst *(1)) + unfolding setprod.insert[OF *(2-)] + apply auto + done + assume as: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" + moreover + have "\<And>x. min (b \<bullet> k) c = max (a \<bullet> k) c \<Longrightarrow> + x * (b\<bullet>k - a\<bullet>k) = x * (max (a \<bullet> k) c - a \<bullet> k) + x * (b \<bullet> k - max (a \<bullet> k) c)" + by (auto simp add: field_simps) + moreover + have **: "(\<Prod>i\<in>Basis. ((\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) *\<^sub>R i) \<bullet> i - a \<bullet> i)) = + (\<Prod>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) - a \<bullet> i)" + "(\<Prod>i\<in>Basis. b \<bullet> i - ((\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i)) = + (\<Prod>i\<in>Basis. b \<bullet> i - (if i = k then max (a \<bullet> k) c else a \<bullet> i))" + by (auto intro!: setprod.cong) + have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False" + unfolding not_le + using as[unfolded ,rule_format,of k] assms + by auto + ultimately show ?thesis + using assms + unfolding simps ** + unfolding *(1)[of "\<lambda>i x. b\<bullet>i - x"] *(1)[of "\<lambda>i x. x - a\<bullet>i"] + unfolding *(2) + by auto +next + assume "\<not> (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" + then have "cbox a b = {}" + unfolding box_eq_empty by (auto simp: not_le) + then show ?thesis + by (auto simp: not_le) +qed + +subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close> + +definition "gauge d \<longleftrightarrow> (\<forall>x. x \<in> d x \<and> open (d x))" + +lemma gaugeI: + assumes "\<And>x. x \<in> g x" + and "\<And>x. open (g x)" + shows "gauge g" + using assms unfolding gauge_def by auto + +lemma gaugeD[dest]: + assumes "gauge d" + shows "x \<in> d x" + and "open (d x)" + using assms unfolding gauge_def by auto + +lemma gauge_ball_dependent: "\<forall>x. 0 < e x \<Longrightarrow> gauge (\<lambda>x. ball x (e x))" + unfolding gauge_def by auto + +lemma gauge_ball[intro]: "0 < e \<Longrightarrow> gauge (\<lambda>x. ball x e)" + unfolding gauge_def by auto + +lemma gauge_trivial[intro!]: "gauge (\<lambda>x. ball x 1)" + by (rule gauge_ball) auto + +lemma gauge_inter[intro]: "gauge d1 \<Longrightarrow> gauge d2 \<Longrightarrow> gauge (\<lambda>x. d1 x \<inter> d2 x)" + unfolding gauge_def by auto + +lemma gauge_inters: + assumes "finite s" + and "\<forall>d\<in>s. gauge (f d)" + shows "gauge (\<lambda>x. \<Inter>{f d x | d. d \<in> s})" +proof - + have *: "\<And>x. {f d x |d. d \<in> s} = (\<lambda>d. f d 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: + "(\<forall>x. \<exists>d :: real. p x \<longrightarrow> 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. p x \<longrightarrow> q d x)" + by (metis zero_less_one) + + +subsection \<open>Divisions.\<close> + +definition division_of (infixl "division'_of" 40) +where + "s division_of i \<longleftrightarrow> + finite s \<and> + (\<forall>k\<in>s. k \<subseteq> i \<and> k \<noteq> {} \<and> (\<exists>a b. k = cbox a b)) \<and> + (\<forall>k1\<in>s. \<forall>k2\<in>s. k1 \<noteq> k2 \<longrightarrow> interior(k1) \<inter> interior(k2) = {}) \<and> + (\<Union>s = i)" + +lemma division_ofD[dest]: + assumes "s division_of i" + shows "finite s" + and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i" + and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}" + and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b" + and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}" + and "\<Union>s = i" + using assms unfolding division_of_def by auto + +lemma division_ofI: + assumes "finite s" + and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i" + and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}" + and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b" + and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}" + and "\<Union>s = i" + shows "s division_of i" + using assms unfolding division_of_def by auto + +lemma division_of_finite: "s division_of i \<Longrightarrow> finite s" + unfolding division_of_def by auto + +lemma division_of_self[intro]: "cbox a b \<noteq> {} \<Longrightarrow> {cbox a b} division_of (cbox a b)" + unfolding division_of_def by auto + +lemma division_of_trivial[simp]: "s division_of {} \<longleftrightarrow> s = {}" + unfolding division_of_def by auto + +lemma division_of_sing[simp]: + "s division_of cbox a (a::'a::euclidean_space) \<longleftrightarrow> s = {cbox a a}" + (is "?l = ?r") +proof + assume ?r + moreover + { fix k + assume "s = {{a}}" "k\<in>s" + then have "\<exists>x y. k = cbox x y" + apply (rule_tac x=a in exI)+ + apply (force simp: cbox_sing) + done + } + ultimately show ?l + unfolding division_of_def cbox_sing by auto +next + assume ?l + note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]] + { + fix x + assume x: "x \<in> s" have "x = {a}" + using *(2)[rule_format,OF x] by auto + } + moreover have "s \<noteq> {}" + using *(4) 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 \<Longrightarrow> \<forall>x\<in>i. \<exists>k\<in>s. x \<in> k" + unfolding division_of_def by auto + +lemma forall_in_division: + "d division_of i \<Longrightarrow> (\<forall>x\<in>d. P x) \<longleftrightarrow> (\<forall>a b. cbox a b \<in> d \<longrightarrow> P (cbox a b))" + unfolding division_of_def by fastforce + +lemma division_of_subset: + assumes "p division_of (\<Union>p)" + and "q \<subseteq> p" + shows "q division_of (\<Union>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 \<in> q" + then have kp: "k \<in> p" + using assms(2) by auto + show "k \<subseteq> \<Union>q" + using \<open>k \<in> q\<close> by auto + show "\<exists>a b. k = cbox a b" + using *(4)[OF kp] by auto + show "k \<noteq> {}" + using *(3)[OF kp] by auto + } + fix k1 k2 + assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2" + then have **: "k1 \<in> p" "k2 \<in> p" "k1 \<noteq> k2" + using assms(2) by auto + show "interior k1 \<inter> interior k2 = {}" + using *(5)[OF **] by auto +qed auto + +lemma division_of_union_self[intro]: "p division_of s \<Longrightarrow> p division_of (\<Union>p)" + unfolding division_of_def by auto + +lemma division_of_content_0: + assumes "content (cbox a b) = 0" "d division_of (cbox a b)" + shows "\<forall>k\<in>d. content k = 0" + unfolding forall_in_division[OF assms(2)] + by (metis antisym_conv assms content_pos_le content_subset division_ofD(2)) + +lemma division_inter: + fixes s1 s2 :: "'a::euclidean_space set" + assumes "p1 division_of s1" + and "p2 division_of s2" + shows "{k1 \<inter> k2 | k1 k2. k1 \<in> p1 \<and> k2 \<in> p2 \<and> k1 \<inter> k2 \<noteq> {}} division_of (s1 \<inter> s2)" + (is "?A' division_of _") +proof - + let ?A = "{s. s \<in> (\<lambda>(k1,k2). k1 \<inter> k2) ` (p1 \<times> p2) \<and> s \<noteq> {}}" + have *: "?A' = ?A" by auto + show ?thesis + unfolding * + proof (rule division_ofI) + have "?A \<subseteq> (\<lambda>(x, y). x \<inter> y) ` (p1 \<times> p2)" + by auto + moreover have "finite (p1 \<times> p2)" + using assms unfolding division_of_def by auto + ultimately show "finite ?A" by auto + have *: "\<And>s. \<Union>{x\<in>s. x \<noteq> {}} = \<Union>s" + by auto + show "\<Union>?A = s1 \<inter> s2" + apply (rule set_eqI) + unfolding * and UN_iff + using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] + apply auto + done + { + fix k + assume "k \<in> ?A" + then obtain k1 k2 where k: "k = k1 \<inter> k2" "k1 \<in> p1" "k2 \<in> p2" "k \<noteq> {}" + by auto + then show "k \<noteq> {}" + by auto + show "k \<subseteq> s1 \<inter> 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 "\<exists>a b. k = cbox a b" + unfolding k k1 k2 unfolding inter_interval by auto + } + fix k1 k2 + assume "k1 \<in> ?A" + then obtain x1 y1 where k1: "k1 = x1 \<inter> y1" "x1 \<in> p1" "y1 \<in> p2" "k1 \<noteq> {}" + by auto + assume "k2 \<in> ?A" + then obtain x2 y2 where k2: "k2 = x2 \<inter> y2" "x2 \<in> p1" "y2 \<in> p2" "k2 \<noteq> {}" + by auto + assume "k1 \<noteq> k2" + then have th: "x1 \<noteq> x2 \<or> y1 \<noteq> y2" + unfolding k1 k2 by auto + have *: "interior x1 \<inter> interior x2 = {} \<or> interior y1 \<inter> interior y2 = {} \<Longrightarrow> + interior (x1 \<inter> y1) \<subseteq> interior x1 \<Longrightarrow> interior (x1 \<inter> y1) \<subseteq> interior y1 \<Longrightarrow> + interior (x2 \<inter> y2) \<subseteq> interior x2 \<Longrightarrow> interior (x2 \<inter> y2) \<subseteq> interior y2 \<Longrightarrow> + interior (x1 \<inter> y1) \<inter> interior (x2 \<inter> y2) = {}" by auto + show "interior k1 \<inter> 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) \<subseteq> i" + shows "{cbox a b \<inter> k | k. k \<in> d \<and> cbox a b \<inter> k \<noteq> {}} 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 \<inter> 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_inter: + fixes s t :: "'a::euclidean_space set" + assumes "p1 division_of s" + and "p2 division_of t" + shows "\<exists>p. p division_of (s \<inter> t)" +using assms division_inter by blast + +lemma elementary_inters: + assumes "finite f" + and "f \<noteq> {}" + and "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::euclidean_space) set)" + shows "\<exists>p. p division_of (\<Inter>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 \<Inter>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_inter Inter_insert) + qed +qed auto + +lemma division_disjoint_union: + assumes "p1 division_of s1" + and "p2 division_of s2" + and "interior s1 \<inter> interior s2 = {}" + shows "(p1 \<union> p2) division_of (s1 \<union> s2)" +proof (rule division_ofI) + note d1 = division_ofD[OF assms(1)] + note d2 = division_ofD[OF assms(2)] + show "finite (p1 \<union> p2)" + using d1(1) d2(1) by auto + show "\<Union>(p1 \<union> p2) = s1 \<union> s2" + using d1(6) d2(6) by auto + { + fix k1 k2 + assume as: "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2" + moreover + let ?g="interior k1 \<inter> interior k2 = {}" + { + assume as: "k1\<in>p1" "k2\<in>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\<in>p2" "k2\<in>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 \<in> p1 \<union> p2" + show "k \<subseteq> s1 \<union> s2" + using k d1(2) d2(2) by auto + show "k \<noteq> {}" + using k d1(3) d2(3) by auto + show "\<exists>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 \<subseteq> cbox a b" + and nonempty: "cbox c d \<noteq> {}" + obtains p where "p division_of (cbox a b)" "cbox c d \<in> p" +proof + let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a. + cbox (\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)" + define p where "p = ?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})" + + show "cbox c d \<in> p" + unfolding p_def + by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"]) + { + fix i :: 'a + assume "i \<in> Basis" + with incl nonempty have "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i" + unfolding box_eq_empty subset_box by (auto simp: not_le) + } + note ord = this + + 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 \<in> p" + then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" + by (auto simp: p_def) + then show "\<exists>a b. k = cbox a b" + by auto + have "k \<subseteq> cbox a b \<and> k \<noteq> {}" + proof (simp add: k box_eq_empty subset_box not_less, safe) + fix i :: 'a + assume i: "i \<in> Basis" + with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)" + by (auto simp: PiE_iff) + with i ord[of i] + show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i" + by auto + qed + then show "k \<noteq> {}" "k \<subseteq> cbox a b" + by auto + { + fix l + assume "l \<in> p" + then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g" + by (auto simp: p_def) + assume "l \<noteq> k" + have "\<exists>i\<in>Basis. f i \<noteq> g i" + proof (rule ccontr) + assume "\<not> ?thesis" + with f g have "f = g" + by (auto simp: PiE_iff extensional_def intro!: ext) + with \<open>l \<noteq> k\<close> show False + by (simp add: l k) + qed + then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" .. + then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)" + "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)" + using f g by (auto simp: PiE_iff) + with * ord[of i] show "interior l \<inter> interior k = {}" + by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i]) + } + note \<open>k \<subseteq> cbox a b\<close> + } + moreover + { + fix x assume x: "x \<in> cbox a b" + have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}" + proof + fix i :: 'a + assume "i \<in> Basis" + with x ord[of i] + have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or> + (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)" + by (auto simp: cbox_def) + then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}" + by auto + qed + then obtain f where + f: "\<forall>i\<in>Basis. x \<bullet> i \<in> {fst (f i) \<bullet> i..snd (f i) \<bullet> i} \<and> f i \<in> {(a, c), (c, d), (d, b)}" + unfolding bchoice_iff .. + moreover from f have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" + by auto + moreover from f have "x \<in> ?B (restrict f Basis)" + by (auto simp: mem_box) + ultimately have "\<exists>k\<in>p. x \<in> k" + unfolding p_def by blast + } + ultimately show "\<Union>p = cbox a b" + by auto + qed +qed + +lemma partial_division_extend_interval: + assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> cbox a b" + obtains q where "p \<subseteq> 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: "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q" + proof + fix k + assume kp: "k \<in> p" + obtain c d where k: "k = cbox c d" + using p(4)[OF kp] by blast + have *: "cbox c d \<subseteq> cbox a b" "cbox c d \<noteq> {}" + 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 \<in> q" + by (rule partial_division_extend_1[OF *]) + then show "\<exists>q. q division_of cbox a b \<and> k \<in> q" + unfolding k by auto + qed + obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x" + using bchoice[OF div_cbox] by blast + { fix x + assume x: "x \<in> p" + have "q x division_of \<Union>q x" + apply (rule division_ofI) + using division_ofD[OF q(1)[OF x]] + apply auto + done } + then have "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})" + by (meson Diff_subset division_of_subset) + then have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" + apply - + apply (rule elementary_inters [OF finite_imageI[OF p(1)]]) + apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]]) + done + then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" .. + have "d \<union> p division_of cbox a b" + proof - + have te: "\<And>s f t. s \<noteq> {} \<Longrightarrow> \<forall>i\<in>s. f i \<union> i = t \<Longrightarrow> t = \<Inter>(f ` s) \<union> \<Union>s" by auto + have cbox_eq: "cbox a b = \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p" + proof (rule te[OF False], clarify) + fix i + assume i: "i \<in> p" + show "\<Union>(q i - {i}) \<union> 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 \<in> p" + have *: "\<And>u t s. t \<inter> s = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<inter> t = {}" + by auto + have "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior k = {}" + proof (rule *[OF inter_interior_unions_intervals]) + note qk=division_ofD[OF q(1)[OF k]] + show "finite (q k - {k})" "open (interior k)" "\<forall>t\<in>q k - {k}. \<exists>a b. t = cbox a b" + using qk by auto + show "\<forall>t\<in>q k - {k}. interior k \<inter> interior t = {}" + using qk(5) using q(2)[OF k] by auto + show "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<subseteq> interior (\<Union>(q k - {k}))" + apply (rule interior_mono)+ + using k + apply auto + done + qed } note [simp] = this + show "d \<union> p division_of (cbox a b)" + unfolding cbox_eq + apply (rule division_disjoint_union[OF d assms(1)]) + apply (rule inter_interior_unions_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 \<Longrightarrow> bounded s" + unfolding division_of_def by (metis bounded_Union bounded_cbox) + +lemma elementary_subset_cbox: + "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> cbox a (b::'a::euclidean_space)" + by (meson elementary_bounded bounded_subset_cbox) + +lemma division_union_intervals_exists: + fixes a b :: "'a::euclidean_space" + assumes "cbox a b \<noteq> {}" + obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)" +proof (cases "cbox c d = {}") + case True + show ?thesis + apply (rule that[of "{}"]) + unfolding True + using assms + apply auto + done +next + case False + show ?thesis + proof (cases "cbox a b \<inter> 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 \<inter> cbox c d = cbox u v" + unfolding inter_interval by auto + have uv_sub: "cbox u v \<subseteq> cbox c d" using uv by auto + obtain p where "p division_of cbox c d" "cbox u v \<in> p" + by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]]) + note p = this division_ofD[OF this(1)] + have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v \<inter> \<Union>(p - {cbox u v}))" + apply (rule arg_cong[of _ _ interior]) + using p(8) uv by auto + also have "\<dots> = {}" + unfolding interior_Int + apply (rule inter_interior_unions_intervals) + using p(6) p(7)[OF p(2)] p(3) + apply auto + done + finally have [simp]: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" by simp + have cbe: "cbox a b \<union> cbox c d = cbox a b \<union> \<Union>(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 \<union> \<Union>(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 \<union> \<Union>(p - {cbox u v})" + by (metis (no_types) Diff_subset \<open>interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}\<close> 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_unions: + assumes "finite f" + and "\<And>p. p \<in> f \<Longrightarrow> p division_of (\<Union>p)" + and "\<And>k1 k2. k1 \<in> \<Union>f \<Longrightarrow> k2 \<in> \<Union>f \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}" + shows "\<Union>f division_of \<Union>\<Union>f" + using assms + by (auto intro!: division_ofI) + +lemma elementary_union_interval: + fixes a b :: "'a::euclidean_space" + assumes "p division_of \<Union>p" + obtains q where "q division_of (cbox a b \<union> \<Union>p)" +proof - + note assm = division_ofD[OF assms] + have lem1: "\<And>f s. \<Union>\<Union>(f ` s) = \<Union>((\<lambda>x. \<Union>(f x)) ` s)" + by auto + have lem2: "\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f" + by auto + { + presume "p = {} \<Longrightarrow> thesis" + "cbox a b = {} \<Longrightarrow> thesis" + "cbox a b \<noteq> {} \<Longrightarrow> interior (cbox a b) = {} \<Longrightarrow> thesis" + "p \<noteq> {} \<Longrightarrow> interior (cbox a b)\<noteq>{} \<Longrightarrow> cbox a b \<noteq> {} \<Longrightarrow> thesis" + then show thesis by auto + next + assume as: "p = {}" + obtain p where "p division_of (cbox a b)" + by (rule elementary_interval) + then show thesis + using as that by auto + next + assume as: "cbox a b = {}" + show thesis + using as assms that by auto + next + assume as: "interior (cbox a b) = {}" "cbox a b \<noteq> {}" + show thesis + apply (rule that[of "insert (cbox a b) p"],rule division_ofI) + unfolding finite_insert + apply (rule assm(1)) unfolding Union_insert + using assm(2-4) as + apply - + apply (fast dest: assm(5))+ + done + next + assume as: "p \<noteq> {}" "interior (cbox a b) \<noteq> {}" "cbox a b \<noteq> {}" + have "\<forall>k\<in>p. \<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)" + proof + fix k + assume kp: "k \<in> p" + from assm(4)[OF kp] obtain c d where "k = cbox c d" by blast + then show "\<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)" + by (meson as(3) division_union_intervals_exists) + qed + from bchoice[OF this] obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" .. + note q = division_ofD[OF this[rule_format]] + let ?D = "\<Union>{insert (cbox a b) (q k) | k. k \<in> p}" + show thesis + proof (rule that[OF division_ofI]) + have *: "{insert (cbox a b) (q k) |k. k \<in> p} = (\<lambda>k. insert (cbox a b) (q k)) ` p" + by auto + show "finite ?D" + using "*" assm(1) q(1) by auto + show "\<Union>?D = cbox a b \<union> \<Union>p" + unfolding * lem1 + unfolding lem2[OF as(1), of "cbox a b", symmetric] + using q(6) + by auto + fix k + assume k: "k \<in> ?D" + then show "k \<subseteq> cbox a b \<union> \<Union>p" + using q(2) by auto + show "k \<noteq> {}" + using q(3) k by auto + show "\<exists>a b. k = cbox a b" + using q(4) k by auto + fix k' + assume k': "k' \<in> ?D" "k \<noteq> k'" + obtain x where x: "k \<in> insert (cbox a b) (q x)" "x\<in>p" + using k by auto + obtain x' where x': "k'\<in>insert (cbox a b) (q x')" "x'\<in>p" + using k' by auto + show "interior k \<inter> interior k' = {}" + proof (cases "x = x'") + case True + show ?thesis + using True k' q(5) x' x by auto + next + case False + { + presume "k = cbox a b \<Longrightarrow> ?thesis" + and "k' = cbox a b \<Longrightarrow> ?thesis" + and "k \<noteq> cbox a b \<Longrightarrow> k' \<noteq> cbox a b \<Longrightarrow> ?thesis" + then show ?thesis by linarith + next + assume as': "k = cbox a b" + show ?thesis + using as' k' q(5) x' by blast + next + assume as': "k' = cbox a b" + show ?thesis + using as' k'(2) q(5) x by blast + } + assume as': "k \<noteq> cbox a b" "k' \<noteq> cbox a b" + obtain c d where k: "k = cbox c d" + using q(4)[OF x(2,1)] by blast + have "interior k \<inter> interior (cbox a b) = {}" + using as' k'(2) q(5) x by blast + then have "interior k \<subseteq> interior x" + using interior_subset_union_intervals + by (metis as(2) 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' \<inter> interior (cbox a b) = {}" + using as'(2) q(5) x' by blast + then have "interior k' \<subseteq> interior x'" + by (metis as(2) c_d interior_subset_union_intervals q(2) x'(1) x'(2)) + ultimately show ?thesis + using assm(5)[OF x(2) x'(2) False] by auto + qed + qed + } +qed + +lemma elementary_unions_intervals: + assumes fin: "finite f" + and "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)" + obtains p where "p division_of (\<Union>f)" +proof - + have "\<exists>p. p division_of (\<Union>f)" + proof (induct_tac f rule:finite_subset_induct) + show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto + next + fix x F + assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f" + from this(3) obtain p where p: "p division_of \<Union>F" .. + from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast + have *: "\<Union>F = \<Union>p" + using division_ofD[OF p] by auto + show "\<exists>p. p division_of \<Union>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 \<union> t)" +proof - + have *: "s \<union> t = \<Union>ps \<union> \<Union>pt" + using assms unfolding division_of_def by auto + show ?thesis + apply (rule elementary_unions_intervals[of "ps \<union> 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 \<subseteq> t" + obtains r where "p \<subseteq> 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 \<subseteq> cbox a b" + using elementary_subset_cbox[OF assms(2)] by auto + obtain r1 where "p \<subseteq> 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 \<Union>(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 (\<Union>(r1 - p)) \<inter> (\<Union>q)" + by (metis assms(2) divq(6) elementary_inter) + { + fix x + assume x: "x \<in> t" "x \<notin> s" + then have "x\<in>\<Union>r1" + unfolding r1 using ab by auto + then obtain r where r: "r \<in> r1" "x \<in> r" + unfolding Union_iff .. + moreover + have "r \<notin> p" + proof + assume "r \<in> p" + then have "x \<in> s" using divp(2) r by auto + then show False using x by auto + qed + ultimately have "x\<in>\<Union>(r1 - p)" by auto + } + then have *: "t = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)" + unfolding divp divq using assms(3) by auto + show ?thesis + apply (rule that[of "p \<union> r2"]) + unfolding * + defer + apply (rule division_disjoint_union) + unfolding divp(6) + apply(rule assms r2)+ + proof - + have "interior s \<inter> interior (\<Union>(r1-p)) = {}" + proof (rule inter_interior_unions_intervals) + show "finite (r1 - p)" and "open (interior s)" and "\<forall>t\<in>r1-p. \<exists>a b. t = cbox a b" + using r1 by auto + have *: "\<And>s. (\<And>x. x \<in> s \<Longrightarrow> False) \<Longrightarrow> s = {}" + by auto + show "\<forall>t\<in>r1-p. interior s \<inter> interior t = {}" + proof + fix m x + assume as: "m \<in> r1 - p" + have "interior m \<inter> interior (\<Union>p) = {}" + proof (rule inter_interior_unions_intervals) + show "finite p" and "open (interior m)" and "\<forall>t\<in>p. \<exists>a b. t = cbox a b" + using divp by auto + show "\<forall>t\<in>p. interior m \<inter> interior t = {}" + by (metis DiffD1 DiffD2 as r1(1) r1(7) set_rev_mp) + qed + then show "interior s \<inter> interior m = {}" + unfolding divp by auto + qed + qed + then show "interior s \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}" + using interior_subset by auto + qed auto +qed + +lemma division_split_left_inj: + fixes type :: "'a::euclidean_space" + assumes "d division_of i" + and "k1 \<in> d" + and "k2 \<in> d" + and "k1 \<noteq> k2" + and "k1 \<inter> {x::'a. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}" + and k: "k\<in>Basis" + shows "content(k1 \<inter> {x. x\<bullet>k \<le> c}) = 0" +proof - + note d=division_ofD[OF assms(1)] + have *: "\<And>(a::'a) b c. content (cbox a b \<inter> {x. x\<bullet>k \<le> c}) = 0 \<longleftrightarrow> + interior(cbox a b \<inter> {x. x\<bullet>k \<le> c}) = {}" + unfolding interval_split[OF k] content_eq_0_interior by auto + guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this + guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this + have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}" + by auto + show ?thesis + unfolding uv1 uv2 * + apply (rule **[OF d(5)[OF assms(2-4)]]) + apply (simp add: uv1) + using assms(5) uv1 by auto +qed + +lemma division_split_right_inj: + fixes type :: "'a::euclidean_space" + assumes "d division_of i" + and "k1 \<in> d" + and "k2 \<in> d" + and "k1 \<noteq> k2" + and "k1 \<inter> {x::'a. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}" + and k: "k \<in> Basis" + shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0" +proof - + note d=division_ofD[OF assms(1)] + have *: "\<And>a b::'a. \<And>c. content(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = 0 \<longleftrightarrow> + interior(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = {}" + unfolding interval_split[OF k] content_eq_0_interior by auto + guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this + guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this + have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}" + by auto + show ?thesis + unfolding uv1 uv2 * + apply (rule **[OF d(5)[OF assms(2-4)]]) + apply (simp add: uv1) + using assms(5) uv1 by auto +qed + + +lemma division_split: + fixes a :: "'a::euclidean_space" + assumes "p division_of (cbox a b)" + and k: "k\<in>Basis" + shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} division_of(cbox a b \<inter> {x. x\<bullet>k \<le> c})" + (is "?p1 division_of ?I1") + and "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> 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 "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2" + unfolding p(6)[symmetric] by auto + { + fix k + assume "k \<in> ?p1" + then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this + guess u v using p(4)[OF l(2)] by (elim exE) note uv=this + show "k \<subseteq> ?I1" + using l p(2) uv by force + show "k \<noteq> {}" + by (simp add: l) + show "\<exists>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' \<in> ?p1" + then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this + assume "k \<noteq> k'" + then show "interior k \<inter> interior k' = {}" + unfolding l l' using p(5)[OF l(2) l'(2)] by auto + } + { + fix k + assume "k \<in> ?p2" + then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this + guess u v using p(4)[OF l(2)] by (elim exE) note uv=this + show "k \<subseteq> ?I2" + using l p(2) uv by force + show "k \<noteq> {}" + by (simp add: l) + show "\<exists>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' \<in> ?p2" + then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this + assume "k \<noteq> k'" + then show "interior k \<inter> interior k' = {}" + unfolding l l' using p(5)[OF l(2) l'(2)] by auto + } +qed + +subsection \<open>Tagged (partial) divisions.\<close> + +definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) + where "s tagged_partial_division_of i \<longleftrightarrow> + finite s \<and> + (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and> + (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow> + interior k1 \<inter> interior k2 = {})" + +lemma tagged_partial_division_ofD[dest]: + assumes "s tagged_partial_division_of i" + shows "finite s" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b" + and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> + (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow> interior k1 \<inter> interior k2 = {}" + using assms unfolding tagged_partial_division_of_def by blast+ + +definition tagged_division_of (infixr "tagged'_division'_of" 40) + where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)" + +lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s" + unfolding tagged_division_of_def tagged_partial_division_of_def by auto + +lemma tagged_division_of: + "s tagged_division_of i \<longleftrightarrow> + finite s \<and> + (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and> + (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow> + interior k1 \<inter> interior k2 = {}) \<and> + (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)" + unfolding tagged_division_of_def tagged_partial_division_of_def by auto + +lemma tagged_division_ofI: + assumes "finite s" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b" + and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow> + interior k1 \<inter> interior k2 = {}" + and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)" + shows "s tagged_division_of i" + unfolding tagged_division_of + using assms + apply auto + apply fastforce+ + done + +lemma tagged_division_ofD[dest]: (*FIXME USE A LOCALE*) + assumes "s tagged_division_of i" + shows "finite s" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i" + and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b" + and "\<And>x1 k1 x2 k2. (x1, k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow> + interior k1 \<inter> interior k2 = {}" + and "(\<Union>{k. \<exists>x. (x,k) \<in> 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 "\<Union>(snd ` s) = i" "finite (snd ` s)" + using assm by auto + fix k + assume k: "k \<in> snd ` s" + then obtain xk where xk: "(xk, k) \<in> s" + by auto + then show "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = cbox a b" + using assm by fastforce+ + fix k' + assume k': "k' \<in> snd ` s" "k \<noteq> k'" + from this(1) obtain xk' where xk': "(xk', k') \<in> s" + by auto + then show "interior k \<inter> 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 \<Union>(snd ` s)" +proof (rule division_ofI) + note assm = tagged_partial_division_ofD[OF assms] + show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)" + using assm by auto + fix k + assume k: "k \<in> snd ` s" + then obtain xk where xk: "(xk, k) \<in> s" + by auto + then show "k \<noteq> {}" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>(snd ` s)" + using assm by auto + fix k' + assume k': "k' \<in> snd ` s" "k \<noteq> k'" + from this(1) obtain xk' where xk': "(xk', k') \<in> s" + by auto + then show "interior k \<inter> 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 \<subseteq> s" + shows "t tagged_partial_division_of i" + using assms + unfolding tagged_partial_division_of_def + using finite_subset[OF assms(2)] + by blast + +lemma (in comm_monoid_set) over_tagged_division_lemma: + assumes "p tagged_division_of i" + and "\<And>u v. cbox u v \<noteq> {} \<Longrightarrow> content (cbox u v) = 0 \<Longrightarrow> d (cbox u v) = \<^bold>1" + shows "F (\<lambda>(x,k). d k) p = F d (snd ` p)" +proof - + have *: "(\<lambda>(x,k). d k) = d \<circ> snd" + unfolding o_def by (rule ext) auto + 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\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y" + obtain a b where ab: "snd x = cbox a b" + using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto + have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y" + by (metis prod.collapse \<open>x\<in>p\<close> \<open>snd x = snd y\<close> \<open>x \<noteq> y\<close>)+ + with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}" + by (intro assm(5)[of "fst x" _ "fst y"]) auto + then have "content (cbox a b) = 0" + unfolding \<open>snd x = snd y\<close>[symmetric] ab content_eq_0_interior by auto + then have "d (cbox a b) = \<^bold>1" + using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto + then show "d (snd x) = \<^bold>1" + unfolding ab by auto + qed +qed + +lemma tag_in_interval: "p tagged_division_of i \<Longrightarrow> (x, k) \<in> p \<Longrightarrow> x \<in> 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 {} \<longleftrightarrow> p = {}" + unfolding tagged_partial_division_of_def by auto + +lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \<longleftrightarrow> p = {}" + unfolding tagged_division_of by auto + +lemma tagged_division_of_self: "x \<in> cbox a b \<Longrightarrow> {(x,cbox a b)} tagged_division_of (cbox a b)" + by (rule tagged_division_ofI) auto + +lemma tagged_division_of_self_real: "x \<in> {a .. b::real} \<Longrightarrow> {(x,{a .. b})} tagged_division_of {a .. b}" + unfolding box_real[symmetric] + by (rule tagged_division_of_self) + +lemma tagged_division_union: + assumes "p1 tagged_division_of s1" + and "p2 tagged_division_of s2" + and "interior s1 \<inter> interior s2 = {}" + shows "(p1 \<union> p2) tagged_division_of (s1 \<union> 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 \<union> p2)" + using p1(1) p2(1) by auto + show "\<Union>{k. \<exists>x. (x, k) \<in> p1 \<union> p2} = s1 \<union> s2" + using p1(6) p2(6) by blast + fix x k + assume xk: "(x, k) \<in> p1 \<union> p2" + show "x \<in> k" "\<exists>a b. k = cbox a b" + using xk p1(2,4) p2(2,4) by auto + show "k \<subseteq> s1 \<union> s2" + using xk p1(3) p2(3) by blast + fix x' k' + assume xk': "(x', k') \<in> p1 \<union> p2" "(x, k) \<noteq> (x', k')" + have *: "\<And>a b. a \<subseteq> s1 \<Longrightarrow> b \<subseteq> s2 \<Longrightarrow> interior a \<inter> interior b = {}" + using assms(3) interior_mono by blast + show "interior k \<inter> interior k' = {}" + apply (cases "(x, k) \<in> 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_unions: + assumes "finite iset" + and "\<forall>i\<in>iset. pfn i tagged_division_of i" + and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior(i1) \<inter> interior(i2) = {}" + shows "\<Union>(pfn ` iset) tagged_division_of (\<Union>iset)" +proof (rule tagged_division_ofI) + note assm = tagged_division_ofD[OF assms(2)[rule_format]] + show "finite (\<Union>(pfn ` iset))" + apply (rule finite_Union) + using assms + apply auto + done + have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)" + by blast + also have "\<dots> = \<Union>iset" + using assm(6) by auto + finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>iset" . + fix x k + assume xk: "(x, k) \<in> \<Union>(pfn ` iset)" + then obtain i where i: "i \<in> iset" "(x, k) \<in> pfn i" + by auto + show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>iset" + using assm(2-4)[OF i] using i(1) by auto + fix x' k' + assume xk': "(x', k') \<in> \<Union>(pfn ` iset)" "(x, k) \<noteq> (x', k')" + then obtain i' where i': "i' \<in> iset" "(x', k') \<in> pfn i'" + by auto + have *: "\<And>a b. i \<noteq> i' \<Longrightarrow> a \<subseteq> i \<Longrightarrow> b \<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}" + using i(1) i'(1) + using assms(3)[rule_format] interior_mono + by blast + show "interior k \<inter> interior k' = {}" + apply (cases "i = i'") + using assm(5) i' i(2) xk'(2) apply blast + using "*" assm(3) i' i by auto +qed + +lemma tagged_partial_division_of_union_self: + assumes "p tagged_partial_division_of s" + shows "p tagged_division_of (\<Union>(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 (\<Union>(snd ` p))" + apply (rule tagged_division_ofI) + using tagged_division_ofD[OF assms] + apply auto + done + +subsection \<open>Functions closed on boxes: morphisms from boxes to monoids\<close> + +text \<open>This auxiliary structure is used to sum up over the elements of a division. Main theorem is + @{text operative_division}. Instances for the monoid are @{typ "'a option"}, @{typ real}, and + @{typ bool}.\<close> + +lemma property_empty_interval: "\<forall>a b. content (cbox a b) = 0 \<longrightarrow> P (cbox a b) \<Longrightarrow> P {}" + using content_empty unfolding empty_as_interval by auto + +paragraph \<open>Using additivity of lifted function to encode definedness.\<close> + +definition lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option" +where + "lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>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 (in comm_monoid) comm_monoid_lift_option: "comm_monoid (lift_option op \<^bold>*) (Some \<^bold>1)" + proof qed (auto simp: lift_option_def ac_simps split: bind_split) + +lemma (in comm_monoid) comm_monoid_set_lift_option: "comm_monoid_set (lift_option op \<^bold>*) (Some \<^bold>1)" + proof qed (auto simp: lift_option_def ac_simps split: bind_split) + +lemma comm_monoid_and: "comm_monoid op \<and> True" + proof qed auto + +lemma comm_monoid_set_and: "comm_monoid_set op \<and> True" + proof qed auto + +paragraph \<open>Operative\<close> + +definition (in comm_monoid) operative :: "('b::euclidean_space set \<Rightarrow> 'a) \<Rightarrow> bool" + where "operative g \<longleftrightarrow> + (\<forall>a b. content (cbox a b) = 0 \<longrightarrow> g (cbox a b) = \<^bold>1) \<and> + (\<forall>a b c. \<forall>k\<in>Basis. g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c}))" + +lemma (in comm_monoid) operativeD[dest]: + assumes "operative g" + shows "\<And>a b. content (cbox a b) = 0 \<Longrightarrow> g (cbox a b) = \<^bold>1" + and "\<And>a b c k. k \<in> Basis \<Longrightarrow> g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" + using assms unfolding operative_def by auto + +lemma (in comm_monoid) operative_empty: "operative g \<Longrightarrow> g {} = \<^bold>1" + unfolding operative_def by (rule property_empty_interval) auto + +lemma operative_content[intro]: "add.operative content" + by (force simp add: add.operative_def content_split[symmetric]) + +definition "division_points (k::('a::euclidean_space) set) d = + {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and> + (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>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 = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and> + (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}" + have *: "division_points i d = \<Union>(?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 "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" "a\<bullet>k < c" "c < b\<bullet>k" + and k: "k \<in> Basis" + shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subseteq> + division_points (cbox a b) d" (is ?t1) + and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq> + division_points (cbox a b) d" (is ?t2) +proof - + note assm = division_ofD[OF assms(1)] + have *: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i" + "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) *\<^sub>R i) \<bullet> i" + "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i \<le> b\<bullet>i" + "min (b \<bullet> k) c = c" "max (a \<bullet> k) c = c" + using assms using less_imp_le by auto + show ?t1 (*FIXME a horrible mess*) + 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 * + apply (rule subsetI) + unfolding mem_Collect_eq split_beta + apply (erule bexE conjE)+ + apply (simp add: ) + apply (erule exE conjE)+ + proof + fix i l x + assume as: + "a \<bullet> fst x < snd x" "snd x < (if fst x = k then c else b \<bullet> fst x)" + "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x" + "i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}" + and fstx: "fst x \<in> Basis" + from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this + have *: "\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i" + using as(6) unfolding l interval_split[OF k] box_ne_empty as . + have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" + using l using as(6) unfolding box_ne_empty[symmetric] by auto + show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x" + apply (rule bexI[OF _ \<open>l \<in> d\<close>]) + using as(1-3,5) fstx + unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as + apply (auto split: if_split_asm) + done + show "snd x < b \<bullet> fst x" + using as(2) \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm) + qed + 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 * + unfolding subset_eq + apply rule + unfolding mem_Collect_eq split_beta + apply (erule bexE conjE)+ + apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) + apply (erule exE conjE)+ + proof + fix i l x + assume as: + "(if fst x = k then c else a \<bullet> fst x) < snd x" "snd x < b \<bullet> fst x" + "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x" + "i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}" + and fstx: "fst x \<in> Basis" + from assm(4)[OF this(5)] guess u v by (elim exE) note l=this + have *: "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i" + using as(6) unfolding l interval_split[OF k] box_ne_empty as . + have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" + using l using as(6) unfolding box_ne_empty[symmetric] by auto + show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x" + apply (rule bexI[OF _ \<open>l \<in> d\<close>]) + using as(1-3,5) fstx + unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as + apply (auto split: if_split_asm) + done + show "a \<bullet> fst x < snd x" + using as(1) \<open>a\<bullet>k < c\<close> by (auto split: if_split_asm) + qed +qed + +lemma division_points_psubset: + fixes a :: "'a::euclidean_space" + assumes "d division_of (cbox a b)" + and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" "a\<bullet>k < c" "c < b\<bullet>k" + and "l \<in> d" + and "interval_lowerbound l\<bullet>k = c \<or> interval_upperbound l\<bullet>k = c" + and k: "k \<in> Basis" + shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subset> + division_points (cbox a b) d" (is "?D1 \<subset> ?D") + and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} \<subset> + division_points (cbox a b) d" (is "?D2 \<subset> ?D") +proof - + have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i" + using assms(2) by (auto intro!:less_imp_le) + guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this + have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "\<forall>i\<in>Basis. a\<bullet>i \<le> u\<bullet>i \<and> v\<bullet>i \<le> b\<bullet>i" + using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty + using subset_box(1) + apply auto + apply blast+ + done + have *: "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k" + "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k" + unfolding l interval_split[OF k] interval_bounds[OF uv(1)] + using uv[rule_format, of k] ab k + by auto + have "\<exists>x. x \<in> ?D - ?D1" + using assms(3-) + unfolding division_points_def interval_bounds[OF ab] + apply - + apply (erule disjE) + apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *) + apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *) + done + moreover have "?D1 \<subseteq> ?D" + by (auto simp add: assms division_points_subset) + ultimately show "?D1 \<subset> ?D" + by blast + have *: "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k" + "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k" + unfolding l interval_split[OF k] interval_bounds[OF uv(1)] + using uv[rule_format, of k] ab k + by auto + have "\<exists>x. x \<in> ?D - ?D2" + using assms(3-) + unfolding division_points_def interval_bounds[OF ab] + apply - + apply (erule disjE) + apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *) + apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *) + done + moreover have "?D2 \<subseteq> ?D" + by (auto simp add: assms division_points_subset) + ultimately show "?D2 \<subset> ?D" + by blast +qed + +lemma (in comm_monoid_set) operative_division: + fixes g :: "'b::euclidean_space set \<Rightarrow> 'a" + assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)" +proof - + define C where [abs_def]: "C = card (division_points (cbox a b) d)" + then show ?thesis + using d + proof (induction C arbitrary: a b d rule: less_induct) + case (less a b d) + show ?case + proof cases + show "content (cbox a b) = 0 \<Longrightarrow> F g d = g (cbox a b)" + using division_of_content_0[OF _ less.prems] operativeD(1)[OF g] division_ofD(4)[OF less.prems] + by (fastforce intro!: neutral) + next + assume "content (cbox a b) \<noteq> 0" + note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq] + then have ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i" + by (auto intro!: less_imp_le) + 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 \<in> Basis" and as: "cbox u v \<in> d" + then have "cbox u v \<noteq> {}" + using less.prems by blast + then have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "u\<bullet>j \<le> v\<bullet>j" + using j unfolding box_ne_empty by auto + have *: "\<And>p r Q. \<not> j\<in>Basis \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> Q (cbox u v)" + using as j by auto + have "(j, u\<bullet>j) \<notin> division_points (cbox a b) d" + "(j, v\<bullet>j) \<notin> 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\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j" + using division_ofD(2,2,3)[OF \<open>d division_of cbox a b\<close> as] + apply (metis j subset_box(1) uv(1)) + by (metis \<open>cbox u v \<subseteq> cbox a b\<close> j subset_box(1) uv(1)) + ultimately have "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j" + unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force } + then have d': "\<forall>i\<in>d. \<exists>u v. i = cbox u v \<and> + (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)" + unfolding forall_in_division[OF less.prems] by blast + have "(1/2) *\<^sub>R (a+b) \<in> cbox a b" + unfolding mem_box using ab by(auto intro!: less_imp_le simp: inner_simps) + note this[unfolded division_ofD(6)[OF \<open>d division_of cbox a b\<close>,symmetric] Union_iff] + then guess i .. note i=this + guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this + have "cbox a b \<in> d" + proof - + have "u = a" "v = b" + unfolding euclidean_eq_iff[where 'a='b] + proof safe + fix j :: 'b + assume j: "j \<in> Basis" + note i(2)[unfolded uv mem_box,rule_format,of j] + then show "u \<bullet> j = a \<bullet> j" and "v \<bullet> j = b \<bullet> 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 \<in> d - {cbox a b}" + then have "x\<in>d" + by auto note d'[rule_format,OF this] + then guess u v by (elim exE conjE) note uv=this + have "u \<noteq> a \<or> v \<noteq> b" + using x[unfolded uv] by auto + then obtain j where "u\<bullet>j \<noteq> a\<bullet>j \<or> v\<bullet>j \<noteq> b\<bullet>j" and j: "j \<in> Basis" + unfolding euclidean_eq_iff[where 'a='b] by auto + then have "u\<bullet>j = v\<bullet>j" + using uv(2)[rule_format,OF j] by auto + then have "content (cbox u v) = 0" + unfolding content_eq_0 using j + by force + then show "g x = \<^bold>1" + unfolding uv(1) by (rule operativeD(1)[OF g]) + 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 "\<exists>x. x \<in> division_points (cbox a b) d" + by auto + then guess k c + unfolding split_paired_Ex division_points_def mem_Collect_eq split_conv + apply (elim exE conjE) + done + note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']] + from this(3) guess j .. note j=this + define d1 where "d1 = {l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}" + define d2 where "d2 = {l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}" + define cb where "cb = (\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)" + define ca where "ca = (\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)" + note division_points_psubset[OF \<open>d division_of cbox a b\<close> ab kc(1-2) j] + note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)] + then have *: "F g d1 = g (cbox a b \<inter> {x. x\<bullet>k \<le> c})" "F g d2 = g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" + unfolding interval_split[OF kc(4)] + apply (rule_tac[!] "less.hyps"[rule_format]) + using division_split[OF \<open>d division_of cbox a b\<close>, where k=k and c=c] + apply (simp_all add: interval_split kc d1_def d2_def division_points_finite[OF \<open>d division_of cbox a b\<close>]) + done + { fix l y + assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} = y \<inter> {x. x \<bullet> k \<le> c}" "l \<noteq> y" + from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this + have "g (l \<inter> {x. x \<bullet> k \<le> c}) = \<^bold>1" + unfolding leq interval_split[OF kc(4)] + apply (rule operativeD[OF g]) + unfolding interval_split[symmetric, OF kc(4)] + using division_split_left_inj less as kc leq by blast + } note fxk_le = this + { fix l y + assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} = y \<inter> {x. c \<le> x \<bullet> k}" "l \<noteq> y" + from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this + have "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1" + unfolding leq interval_split[OF kc(4)] + apply (rule operativeD(1)[OF g]) + unfolding interval_split[symmetric,OF kc(4)] + using division_split_right_inj less leq as kc by blast + } note fxk_ge = this + have d1_alt: "d1 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<le> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}" + using d1_def by auto + have d2_alt: "d2 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<ge> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}" + using d2_def by auto + have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev") + unfolding * using g kc(4) by blast + also have "F g d1 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d" + unfolding d1_alt using division_of_finite[OF less.prems] fxk_le + by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g]) + also have "F g d2 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d" + unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge + by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g]) + also have *: "\<forall>x\<in>d. g x = g (x \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (x \<inter> {x. c \<le> x \<bullet> k})" + unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>] + using g kc(4) by blast + have "F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d \<^bold>* F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d = F g d" + using * by (simp add: distrib) + finally show ?thesis by auto + qed + qed + qed +qed + +lemma (in comm_monoid_set) operative_tagged_division: + assumes f: "operative g" and d: "d tagged_division_of (cbox a b)" + shows "F (\<lambda>(x, l). g l) d = g (cbox a b)" + unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric] + by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d]) + +lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> setsum content d = content (cbox a b)" + by (metis operative_content setsum.operative_division) + +lemma additive_content_tagged_division: + "d tagged_division_of (cbox a b) \<Longrightarrow> setsum (\<lambda>(x,l). content l) d = content (cbox a b)" + unfolding setsum.operative_tagged_division[OF operative_content, symmetric] by blast + +lemma + shows real_inner_1_left: "inner 1 x = x" + and real_inner_1_right: "inner x 1 = x" + by simp_all + +lemma content_real_eq_0: "content {a .. b::real} = 0 \<longleftrightarrow> a \<ge> b" + by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) + +lemma interval_real_split: + "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}" + "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}" + apply (metis Int_atLeastAtMostL1 atMost_def) + apply (metis Int_atLeastAtMostL2 atLeast_def) + done + +lemma (in comm_monoid) operative_1_lt: + "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow> + ((\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1) \<and> (\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))" + apply (simp add: operative_def content_real_eq_0 atMost_def[symmetric] atLeast_def[symmetric] + del: content_real_if) +proof safe + fix a b c :: real + assume *: "\<forall>a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" + assume "a < c" "c < b" + with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}" + by (simp add: less_imp_le min.absorb2 max.absorb2) +next + fix a b c :: real + assume as: "\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1" + "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" + from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2) + have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1" + "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" + by auto + show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" + by (auto simp: min_def max_def le_less) +qed + +lemma (in comm_monoid) operative_1_le: + "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow> + ((\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1) \<and> (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))" + unfolding operative_1_lt +proof safe + fix a b c :: real + assume as: "\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b" + show "g {a..c} \<^bold>* g {c..b} = g {a..b}" + apply (rule as(1)[rule_format]) + using as(2-) + apply auto + done +next + fix a b c :: real + assume "\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1" + and "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" + and "a \<le> c" + and "c \<le> b" + note as = this[rule_format] + show "g {a..c} \<^bold>* g {c..b} = g {a..b}" + proof (cases "c = a \<or> c = b") + case False + then show ?thesis + apply - + apply (subst as(2)) + using as(3-) + apply auto + done + next + case True + then show ?thesis + proof + assume *: "c = a" + then have "g {a .. c} = \<^bold>1" + apply - + apply (rule as(1)[rule_format]) + apply auto + done + then show ?thesis + unfolding * by auto + next + assume *: "c = b" + then have "g {c .. b} = \<^bold>1" + apply - + apply (rule as(1)[rule_format]) + apply auto + done + then show ?thesis + unfolding * by auto + qed + qed +qed + +subsection \<open>Fine-ness of a partition w.r.t. a gauge.\<close> + +definition fine (infixr "fine" 46) + where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)" + +lemma fineI: + assumes "\<And>x k. (x, k) \<in> s \<Longrightarrow> k \<subseteq> d x" + shows "d fine s" + using assms unfolding fine_def by auto + +lemma fineD[dest]: + assumes "d fine s" + shows "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> d x" + using assms unfolding fine_def by auto + +lemma fine_inter: "(\<lambda>x. d1 x \<inter> d2 x) fine p \<longleftrightarrow> d1 fine p \<and> d2 fine p" + unfolding fine_def by auto + +lemma fine_inters: + "(\<lambda>x. \<Inter>{f d x | d. d \<in> s}) fine p \<longleftrightarrow> (\<forall>d\<in>s. (f d) fine p)" + unfolding fine_def by blast + +lemma fine_union: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)" + unfolding fine_def by blast + +lemma fine_unions: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)" + unfolding fine_def by auto + +lemma fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p" + unfolding fine_def by blast + + +subsection \<open>Gauge integral. Define on compact intervals first, then use a limit.\<close> + +definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46) + where "(f has_integral_compact_interval y) i \<longleftrightarrow> + (\<forall>e>0. \<exists>d. gauge d \<and> + (\<forall>p. p tagged_division_of i \<and> d fine p \<longrightarrow> + norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - y) < e))" + +definition has_integral :: + "('n::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> 'n set \<Rightarrow> bool" + (infixr "has'_integral" 46) + where "(f has_integral y) i \<longleftrightarrow> + (if \<exists>a b. i = cbox a b + then (f has_integral_compact_interval y) i + else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> + (\<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral_compact_interval z) (cbox a b) \<and> + norm (z - y) < e)))" + +lemma has_integral: + "(f has_integral y) (cbox a b) \<longleftrightarrow> + (\<forall>e>0. \<exists>d. gauge d \<and> + (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> + norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))" + unfolding has_integral_def has_integral_compact_interval_def + by auto + +lemma has_integral_real: + "(f has_integral y) {a .. b::real} \<longleftrightarrow> + (\<forall>e>0. \<exists>d. gauge d \<and> + (\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow> + norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - 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 d where "gauge d" + and "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d fine p \<Longrightarrow> + norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f(x)) p - y) < e" + using assms unfolding has_integral by auto + +lemma has_integral_alt: + "(f has_integral y) i \<longleftrightarrow> + (if \<exists>a b. i = cbox a b + then (f has_integral y) i + else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> + (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e)))" + unfolding has_integral + unfolding has_integral_compact_interval_def has_integral_def + by auto + +lemma has_integral_altD: + assumes "(f has_integral y) i" + and "\<not> (\<exists>a b. i = cbox a b)" + and "e>0" + obtains B where "B > 0" + and "\<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> + (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> norm(z - y) < e)" + using assms + unfolding has_integral + unfolding has_integral_compact_interval_def has_integral_def + by auto + +definition integrable_on (infixr "integrable'_on" 46) + where "f integrable_on i \<longleftrightarrow> (\<exists>y. (f has_integral y) i)" + +definition "integral i f = (SOME y. (f has_integral y) i \<or> ~ f integrable_on i \<and> y=0)" + +lemma integrable_integral[dest]: "f integrable_on i \<Longrightarrow> (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 \<Longrightarrow> integral i f = 0" + unfolding integrable_on_def integral_def by blast + +lemma has_integral_integrable[intro]: "(f has_integral i) s \<Longrightarrow> f integrable_on s" + unfolding integrable_on_def by auto + +lemma has_integral_integral: "f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s" + by auto + +lemma setsum_content_null: + assumes "content (cbox a b) = 0" + and "p tagged_division_of (cbox a b)" + shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" +proof (rule setsum.neutral, rule) + fix y + assume y: "y \<in> p" + obtain x k where xk: "y = (x, k)" + using surj_pair[of y] by blast + note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] + from this(2) obtain c d where k: "k = cbox c d" by blast + have "(\<lambda>(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" + unfolding xk by auto + also have "\<dots> = 0" + using content_subset[OF assm(1)[unfolded k]] content_pos_le[of c d] + unfolding assms(1) k + by auto + finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" . +qed + + +subsection \<open>Some basic combining lemmas.\<close> + +lemma tagged_division_unions_exists: + assumes "finite iset" + and "\<forall>i\<in>iset. \<exists>p. p tagged_division_of i \<and> d fine p" + and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}" + and "\<Union>iset = i" + obtains p where "p tagged_division_of i" and "d fine p" +proof - + obtain pfn where pfn: + "\<And>x. x \<in> iset \<Longrightarrow> pfn x tagged_division_of x" + "\<And>x. x \<in> iset \<Longrightarrow> d fine pfn x" + using bchoice[OF assms(2)] by auto + show thesis + apply (rule_tac p="\<Union>(pfn ` iset)" in that) + using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force + by (metis (mono_tags, lifting) fine_unions imageE pfn(2)) +qed + + +subsection \<open>The set we're concerned with must be closed.\<close> + +lemma division_of_closed: + fixes i :: "'n::euclidean_space set" + shows "s division_of i \<Longrightarrow> closed i" + unfolding division_of_def by fastforce + +subsection \<open>General bisection principle for intervals; might be useful elsewhere.\<close> + +lemma interval_bisection_step: + fixes type :: "'a::euclidean_space" + assumes "P {}" + and "\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P (s \<union> t)" + and "\<not> P (cbox a (b::'a))" + obtains c d where "\<not> P (cbox c d)" + and "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i" +proof - + have "cbox a b \<noteq> {}" + using assms(1,3) by metis + then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" + by (force simp: mem_box) + { fix f + have "\<lbrakk>finite f; + \<And>s. s\<in>f \<Longrightarrow> P s; + \<And>s. s\<in>f \<Longrightarrow> \<exists>a b. s = cbox a b; + \<And>s t. s\<in>f \<Longrightarrow> t\<in>f \<Longrightarrow> s \<noteq> t \<Longrightarrow> interior s \<inter> interior t = {}\<rbrakk> \<Longrightarrow> P (\<Union>f)" + proof (induct f rule: finite_induct) + case empty + show ?case + using assms(1) by auto + next + case (insert x f) + show ?case + unfolding Union_insert + apply (rule assms(2)[rule_format]) + using inter_interior_unions_intervals [of f "interior x"] + apply (auto simp: insert) + by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff) + qed + } note UN_cases = this + let ?A = "{cbox c d | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<or> + (c\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<and> (d\<bullet>i = b\<bullet>i)}" + let ?PP = "\<lambda>c d. \<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i" + { + presume "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d) \<Longrightarrow> False" + then show thesis + unfolding atomize_not not_all + by (blast intro: that) + } + assume as: "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d)" + have "P (\<Union>?A)" + proof (rule UN_cases) + let ?B = "(\<lambda>s. cbox (\<Sum>i\<in>Basis. (if i \<in> s then a\<bullet>i else (a\<bullet>i + b\<bullet>i) / 2) *\<^sub>R i::'a) + (\<Sum>i\<in>Basis. (if i \<in> s then (a\<bullet>i + b\<bullet>i) / 2 else b\<bullet>i) *\<^sub>R i)) ` {s. s \<subseteq> Basis}" + have "?A \<subseteq> ?B" + proof + fix x + assume "x \<in> ?A" + then obtain c d + where x: "x = cbox c d" + "\<And>i. i \<in> Basis \<Longrightarrow> + c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or> + c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" by blast + show "x \<in> ?B" + unfolding image_iff x + apply (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in bexI) + apply (rule arg_cong2 [where f = cbox]) + using x(2) ab + apply (auto simp add: euclidean_eq_iff[where 'a='a]) + by fastforce + qed + then show "finite ?A" + by (rule finite_subset) auto + next + fix s + assume "s \<in> ?A" + then obtain c d + where s: "s = cbox c d" + "\<And>i. i \<in> Basis \<Longrightarrow> + c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or> + c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" + by blast + show "P s" + unfolding s + apply (rule as[rule_format]) + using ab s(2) by force + show "\<exists>a b. s = cbox a b" + unfolding s by auto + fix t + assume "t \<in> ?A" + then obtain e f where t: + "t = cbox e f" + "\<And>i. i \<in> Basis \<Longrightarrow> + e \<bullet> i = a \<bullet> i \<and> f \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or> + e \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> f \<bullet> i = b \<bullet> i" + by blast + assume "s \<noteq> t" + then have "\<not> (c = e \<and> d = f)" + unfolding s t by auto + then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i': "i \<in> Basis" + unfolding euclidean_eq_iff[where 'a='a] by auto + then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i" + using s(2) t(2) apply fastforce + using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce + have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}" + by auto + show "interior s \<inter> interior t = {}" + unfolding s t interior_cbox + proof (rule *) + fix x + assume "x \<in> box c d" "x \<in> box e f" + then have x: "c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i" + unfolding mem_box using i' + by force+ + show False using s(2)[OF i'] + proof safe + assume as: "c \<bullet> i = a \<bullet> i" "d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" + show False + using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps) + next + assume as: "c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" "d \<bullet> i = b \<bullet> i" + show False + using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps) + qed + qed + qed + also have "\<Union>?A = cbox a b" + proof (rule set_eqI,rule) + fix x + assume "x \<in> \<Union>?A" + then obtain c d where x: + "x \<in> cbox c d" + "\<And>i. i \<in> Basis \<Longrightarrow> + c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or> + c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" + by blast + show "x\<in>cbox a b" + unfolding mem_box + proof safe + fix i :: 'a + assume i: "i \<in> Basis" + then show "a \<bullet> i \<le> x \<bullet> i" "x \<bullet> i \<le> b \<bullet> i" + using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto + qed + next + fix x + assume x: "x \<in> cbox a b" + have "\<forall>i\<in>Basis. + \<exists>c d. (c = a\<bullet>i \<and> d = (a\<bullet>i + b\<bullet>i) / 2 \<or> c = (a\<bullet>i + b\<bullet>i) / 2 \<and> d = b\<bullet>i) \<and> c\<le>x\<bullet>i \<and> x\<bullet>i \<le> d" + (is "\<forall>i\<in>Basis. \<exists>c d. ?P i c d") + unfolding mem_box + proof + fix i :: 'a + assume i: "i \<in> Basis" + have "?P i (a\<bullet>i) ((a \<bullet> i + b \<bullet> i) / 2) \<or> ?P i ((a \<bullet> i + b \<bullet> i) / 2) (b\<bullet>i)" + using x[unfolded mem_box,THEN bspec, OF i] by auto + then show "\<exists>c d. ?P i c d" + by blast + qed + then show "x\<in>\<Union>?A" + unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff + apply auto + apply (rule_tac x="cbox xa xaa" in exI) + unfolding mem_box + apply auto + done + qed + finally show False + using assms by auto +qed + +lemma interval_bisection: + fixes type :: "'a::euclidean_space" + assumes "P {}" + and "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" + and "\<not> P (cbox a (b::'a))" + obtains x where "x \<in> cbox a b" + and "\<forall>e>0. \<exists>c d. x \<in> cbox c d \<and> cbox c d \<subseteq> ball x e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)" +proof - + have "\<forall>x. \<exists>y. \<not> P (cbox (fst x) (snd x)) \<longrightarrow> (\<not> P (cbox (fst y) (snd y)) \<and> + (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and> + 2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" (is "\<forall>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 as: False + obtain c d where "\<not> P (cbox c d)" + "\<forall>i\<in>Basis. + fst x \<bullet> i \<le> c \<bullet> i \<and> + c \<bullet> i \<le> d \<bullet> i \<and> + d \<bullet> i \<le> snd x \<bullet> i \<and> + 2 * (d \<bullet> i - c \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i" + by (rule interval_bisection_step[of P, OF assms(1-2) as]) + then show ?thesis + apply - + apply (rule_tac x="(c,d)" in exI) + apply auto + done + qed + qed + then obtain f where f: + "\<forall>x. + \<not> P (cbox (fst x) (snd x)) \<longrightarrow> + \<not> P (cbox (fst (f x)) (snd (f x))) \<and> + (\<forall>i\<in>Basis. + fst x \<bullet> i \<le> fst (f x) \<bullet> i \<and> + fst (f x) \<bullet> i \<le> snd (f x) \<bullet> i \<and> + snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and> + 2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)" + apply - + apply (drule choice) + apply blast + done + 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 "A 0 = a" "B 0 = b" "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and> + (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and> + 2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>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 + apply (rule f[rule_format]) using assms(3) + apply auto + done + next + case (Suc n) + show ?case + unfolding S + apply (rule f[rule_format]) + using Suc + unfolding S + apply auto + done + qed + qed + note AB = this(1-2) conjunctD2[OF this(3),rule_format] + + have interv: "\<exists>n. \<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e" + if e: "0 < e" for e + proof - + obtain n where n: "(\<Sum>i\<in>Basis. b \<bullet> i - a \<bullet> i) / e < 2 ^ n" + using real_arch_pow[of 2 "(setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto + show ?thesis + proof (rule exI [where x=n], clarify) + fix x y + assume xy: "x\<in>cbox (A n) (B n)" "y\<in>cbox (A n) (B n)" + have "dist x y \<le> setsum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis" + unfolding dist_norm by(rule norm_le_l1) + also have "\<dots> \<le> setsum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis" + proof (rule setsum_mono) + fix i :: 'a + assume i: "i \<in> Basis" + show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i" + using xy[unfolded mem_box,THEN bspec, OF i] + by (auto simp: inner_diff_left) + qed + also have "\<dots> \<le> setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n" + unfolding setsum_divide_distrib + proof (rule setsum_mono) + show "B n \<bullet> i - A n \<bullet> i \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ n" if i: "i \<in> Basis" for i + proof (induct n) + case 0 + then show ?case + unfolding AB by auto + next + case (Suc n) + have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2" + using AB(4)[of i n] using i by auto + also have "\<dots> \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ Suc n" + using Suc by (auto simp add: field_simps) + finally show ?case . + qed + qed + also have "\<dots> < 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 \<le> n" then have "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)" + proof (induction rule: inc_induct) + case (step i) + show ?case + using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto + qed simp + } note ABsubset = this + have "\<exists>a. \<forall>n. a\<in> cbox (A n) (B n)" + by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv]) + (metis nat.exhaust AB(1-3) assms(1,3)) + then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)" + by blast + show thesis + proof (rule that[rule_format, of x0]) + show "x0\<in>cbox a b" + using x0[of 0] unfolding AB . + fix e :: real + assume "e > 0" + from interv[OF this] obtain n + where n: "\<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e" .. + have "\<not> P (cbox (A n) (B n))" + apply (cases "0 < n") + using AB(3)[of "n - 1"] assms(3) AB(1-2) + apply auto + done + moreover have "cbox (A n) (B n) \<subseteq> ball x0 e" + using n using x0[of n] by auto + moreover have "cbox (A n) (B n) \<subseteq> cbox a b" + unfolding AB(1-2)[symmetric] by (rule ABsubset) auto + ultimately show "\<exists>c d. x0 \<in> cbox c d \<and> cbox c d \<subseteq> ball x0 e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> 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 \<open>Cousin's lemma.\<close> + +lemma fine_division_exists: + fixes a b :: "'a::euclidean_space" + assumes "gauge g" + obtains p where "p tagged_division_of (cbox a b)" "g fine p" +proof - + presume "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p) \<Longrightarrow> False" + then obtain p where "p tagged_division_of (cbox a b)" "g fine p" + by blast + then show thesis .. +next + assume as: "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p)" + obtain x where x: + "x \<in> (cbox a b)" + "\<And>e. 0 < e \<Longrightarrow> + \<exists>c d. + x \<in> cbox c d \<and> + cbox c d \<subseteq> ball x e \<and> + cbox c d \<subseteq> (cbox a b) \<and> + \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)" + apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ as]) + apply (simp add: fine_def) + apply (metis tagged_division_union fine_union) + apply (auto simp: ) + done + obtain e where e: "e > 0" "ball x e \<subseteq> 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 \<in> cbox c d" + "cbox c d \<subseteq> ball x e" + "cbox c d \<subseteq> cbox a b" + "\<not> (\<exists>p. p tagged_division_of cbox c d \<and> 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 False + 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 \<open>Basic theorems about integrals.\<close> + +lemma has_integral_unique: + fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" + assumes "(f has_integral k1) i" + and "(f has_integral k2) i" + shows "k1 = k2" +proof (rule ccontr) + let ?e = "norm (k1 - k2) / 2" + assume as: "k1 \<noteq> k2" + then have e: "?e > 0" + by auto + have lem: False + if f_k1: "(f has_integral k1) (cbox a b)" + and f_k2: "(f has_integral k2) (cbox a b)" + and "k1 \<noteq> k2" + for f :: "'n \<Rightarrow> 'a" and a b k1 k2 + proof - + let ?e = "norm (k1 - k2) / 2" + from \<open>k1 \<noteq> k2\<close> have e: "?e > 0" by auto + obtain d1 where d1: + "gauge d1" + "\<And>p. p tagged_division_of cbox a b \<Longrightarrow> + d1 fine p \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k1) < norm (k1 - k2) / 2" + by (rule has_integralD[OF f_k1 e]) blast + obtain d2 where d2: + "gauge d2" + "\<And>p. p tagged_division_of cbox a b \<Longrightarrow> + d2 fine p \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k2) < norm (k1 - k2) / 2" + by (rule has_integralD[OF f_k2 e]) blast + obtain p where p: + "p tagged_division_of cbox a b" + "(\<lambda>x. d1 x \<inter> d2 x) fine p" + by (rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)]]) + let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" + have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)" + using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] + by (auto simp add:algebra_simps norm_minus_commute) + also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" + apply (rule add_strict_mono) + apply (rule_tac[!] d2(2) d1(2)) + using p unfolding fine_def + apply auto + done + finally show False by auto + qed + { + presume "\<not> (\<exists>a b. i = cbox a b) \<Longrightarrow> False" + then show False + using as assms lem by blast + } + assume as: "\<not> (\<exists>a b. i = cbox a b)" + obtain B1 where B1: + "0 < B1" + "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow> + \<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b) \<and> + norm (z - k1) < norm (k1 - k2) / 2" + by (rule has_integral_altD[OF assms(1) as,OF e]) blast + obtain B2 where B2: + "0 < B2" + "\<And>a b. ball 0 B2 \<subseteq> cbox a b \<Longrightarrow> + \<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b) \<and> + norm (z - k2) < norm (k1 - k2) / 2" + by (rule has_integral_altD[OF assms(2) as,OF e]) blast + have "\<exists>a b::'n. ball 0 B1 \<union> ball 0 B2 \<subseteq> cbox a b" + apply (rule bounded_subset_cbox) + using bounded_Un bounded_ball + apply auto + done + then obtain a b :: 'n where ab: "ball 0 B1 \<subseteq> cbox a b" "ball 0 B2 \<subseteq> cbox a b" + by blast + obtain w where w: + "((\<lambda>x. if x \<in> i then f x else 0) 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: + "((\<lambda>x. if x \<in> i then f x else 0) 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 lem[OF w(1) z(1)] by auto + then have "norm (k1 - k2) \<le> norm (z - k2) + norm (w - k1)" + using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] + by (auto simp add: norm_minus_commute) + also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" + apply (rule add_strict_mono) + apply (rule_tac[!] z(2) w(2)) + done + finally show False by auto +qed + +lemma integral_unique [intro]: "(f has_integral y) k \<Longrightarrow> integral k f = y" + unfolding integral_def + by (rule some_equality) (auto intro: has_integral_unique) + +lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> ~ f integrable_on k \<and> y=0" + unfolding integral_def integrable_on_def + apply (erule subst) + apply (rule someI_ex) + by blast + +lemma has_integral_is_0: + fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" + assumes "\<forall>x\<in>s. f x = 0" + shows "(f has_integral 0) s" +proof - + have lem: "\<And>a b. \<And>f::'n \<Rightarrow> 'a. + (\<forall>x\<in>cbox a b. f(x) = 0) \<Longrightarrow> (f has_integral 0) (cbox a b)" + unfolding has_integral + proof clarify + fix a b e + fix f :: "'n \<Rightarrow> 'a" + assume as: "\<forall>x\<in>cbox a b. f x = 0" "0 < (e::real)" + have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e" + if p: "p tagged_division_of cbox a b" for p + proof - + have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) = 0" + proof (rule setsum.neutral, rule) + fix x + assume x: "x \<in> p" + have "f (fst x) = 0" + using tagged_division_ofD(2-3)[OF p, of "fst x" "snd x"] using as x by auto + then show "(\<lambda>(x, k). content k *\<^sub>R f x) x = 0" + apply (subst surjective_pairing[of x]) + unfolding split_conv + apply auto + done + qed + then show ?thesis + using as by auto + qed + then show "\<exists>d. gauge d \<and> + (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e)" + by auto + qed + { + presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis" + with assms lem show ?thesis + by blast + } + have *: "(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. 0)" + apply (rule ext) + using assms + apply auto + done + assume "\<not> (\<exists>a b. s = cbox a b)" + then show ?thesis + using lem + by (subst has_integral_alt) (force simp add: *) +qed + +lemma has_integral_0[simp]: "((\<lambda>x::'n::euclidean_space. 0) has_integral 0) s" + by (rule has_integral_is_0) auto + +lemma has_integral_0_eq[simp]: "((\<lambda>x. 0) has_integral i) s \<longleftrightarrow> i = 0" + using has_integral_unique[OF has_integral_0] by auto + +lemma has_integral_linear: + fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" + assumes "(f has_integral y) s" + and "bounded_linear h" + shows "((h \<circ> f) has_integral ((h y))) s" +proof - + interpret bounded_linear h + using assms(2) . + from pos_bounded obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B" + by blast + have lem: "\<And>(f :: 'n \<Rightarrow> 'a) y a b. + (f has_integral y) (cbox a b) \<Longrightarrow> ((h \<circ> f) has_integral h y) (cbox a b)" + unfolding has_integral + proof (clarify, goal_cases) + case prems: (1 f y a b e) + from pos_bounded + obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B" + by blast + have "e / B > 0" using prems(2) B by simp + then obtain g + where g: "gauge g" + "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> g fine p \<Longrightarrow> + norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e / B" + using prems(1) by auto + { + fix p + assume as: "p tagged_division_of (cbox a b)" "g fine p" + have hc: "\<And>x k. h ((\<lambda>(x, k). content k *\<^sub>R f x) x) = (\<lambda>(x, k). h (content k *\<^sub>R f x)) x" + by auto + then have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = setsum (h \<circ> (\<lambda>(x, k). content k *\<^sub>R f x)) p" + unfolding o_def unfolding scaleR[symmetric] hc by simp + also have "\<dots> = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" + using setsum[of "\<lambda>(x,k). content k *\<^sub>R f x" p] using as by auto + finally have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" . + then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) - h y) < e" + apply (simp add: diff[symmetric]) + apply (rule le_less_trans[OF B(2)]) + using g(2)[OF as] B(1) + apply (auto simp add: field_simps) + done + } + with g show ?case + by (rule_tac x=g in exI) auto + qed + { + presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis" + then show ?thesis + using assms(1) lem by blast + } + assume as: "\<not> (\<exists>a b. s = cbox a b)" + then show ?thesis + proof (subst has_integral_alt, clarsimp) + fix e :: real + assume e: "e > 0" + have *: "0 < e/B" using e B(1) by simp + obtain M where M: + "M > 0" + "\<And>a b. ball 0 M \<subseteq> cbox a b \<Longrightarrow> + \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e / B" + using has_integral_altD[OF assms(1) as *] by blast + show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> + (\<exists>z. ((\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) has_integral z) (cbox a b) \<and> norm (z - h y) < e)" + proof (rule_tac x=M in exI, clarsimp simp add: M, goal_cases) + case prems: (1 a b) + obtain z where z: + "((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b)" + "norm (z - y) < e / B" + using M(2)[OF prems(1)] by blast + have *: "(\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) = h \<circ> (\<lambda>x. if x \<in> s then f x else 0)" + using zero by auto + show ?case + apply (rule_tac x="h z" in exI) + apply (simp add: * lem z(1)) + apply (metis B diff le_less_trans pos_less_divide_eq z(2)) + done + qed + qed +qed + +lemma has_integral_scaleR_left: + "(f has_integral y) s \<Longrightarrow> ((\<lambda>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 has_integral_mult_left: + fixes c :: "_ :: real_normed_algebra" + shows "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x * c) has_integral (y * c)) s" + using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) + +text\<open>The case analysis eliminates the condition @{term "f integrable_on s"} at the cost + of the type class constraint \<open>division_ring\<close>\<close> +corollary integral_mult_left [simp]: + fixes c:: "'a::{real_normed_algebra,division_ring}" + shows "integral s (\<lambda>x. f x * c) = integral s f * c" +proof (cases "f integrable_on s \<or> c = 0") + case True then show ?thesis + by (force intro: has_integral_mult_left) +next + case False then have "~ (\<lambda>x. f x * c) integrable_on s" + using has_integral_mult_left [of "(\<lambda>x. f x * c)" _ s "inverse c"] + by (force 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 (\<lambda>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 (\<lambda>x. f x / z) = integral S (\<lambda>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 \<Longrightarrow> ((\<lambda>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 \<Longrightarrow> ((\<lambda>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 \<noteq> 0 \<Longrightarrow> (f has_integral x) A" + shows "((\<lambda>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 \<Longrightarrow> ((\<lambda>x. -(f x)) has_integral -k) s" + by (drule_tac c="-1" in has_integral_cmul) auto + +lemma has_integral_add: + fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" + assumes "(f has_integral k) s" + and "(g has_integral l) s" + shows "((\<lambda>x. f x + g x) has_integral (k + l)) s" +proof - + have lem: "((\<lambda>x. f x + g x) has_integral (k + l)) (cbox a b)" + if f_k: "(f has_integral k) (cbox a b)" + and g_l: "(g has_integral l) (cbox a b)" + for f :: "'n \<Rightarrow> 'a" and g a b k l + unfolding has_integral + proof clarify + fix e :: real + assume e: "e > 0" + then have *: "e / 2 > 0" + by auto + obtain d1 where d1: + "gauge d1" + "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d1 fine p \<Longrightarrow> + norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) < e / 2" + using has_integralD[OF f_k *] by blast + obtain d2 where d2: + "gauge d2" + "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d2 fine p \<Longrightarrow> + norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l) < e / 2" + using has_integralD[OF g_l *] by blast + show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> + norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)" + proof (rule exI [where x="\<lambda>x. (d1 x) \<inter> (d2 x)"], clarsimp simp add: gauge_inter[OF d1(1) d2(1)]) + fix p + assume as: "p tagged_division_of (cbox a b)" "(\<lambda>x. d1 x \<inter> d2 x) fine p" + have *: "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) = + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p. content k *\<^sub>R g x)" + unfolding scaleR_right_distrib setsum.distrib[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,symmetric] + by (rule setsum.cong) auto + from as have fine: "d1 fine p" "d2 fine p" + unfolding fine_inter by auto + have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) = + norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))" + unfolding * by (auto simp add: algebra_simps) + also have "\<dots> < e/2 + e/2" + apply (rule le_less_trans[OF norm_triangle_ineq]) + using as d1 d2 fine + apply (blast intro: add_strict_mono) + done + finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e" + by auto + qed + qed + { + presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis" + then show ?thesis + using assms lem by force + } + assume as: "\<not> (\<exists>a b. s = cbox a b)" + then show ?thesis + proof (subst has_integral_alt, clarsimp, goal_cases) + case (1 e) + then have *: "e / 2 > 0" + by auto + from has_integral_altD[OF assms(1) as *] + obtain B1 where B1: + "0 < B1" + "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow> + \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - k) < e / 2" + by blast + from has_integral_altD[OF assms(2) as *] + obtain B2 where B2: + "0 < B2" + "\<And>a b. ball 0 B2 \<subseteq> (cbox a b) \<Longrightarrow> + \<exists>z. ((\<lambda>x. if x \<in> s then g x else 0) has_integral z) (cbox a b) \<and> norm (z - l) < e / 2" + by blast + show ?case + proof (rule_tac x="max B1 B2" in exI, clarsimp simp add: max.strict_coboundedI1 B1) + fix a b + assume "ball 0 (max B1 B2) \<subseteq> cbox a (b::'n)" + then have *: "ball 0 B1 \<subseteq> cbox a (b::'n)" "ball 0 B2 \<subseteq> cbox a (b::'n)" + by auto + obtain w where w: + "((\<lambda>x. if x \<in> s then f x else 0) has_integral w) (cbox a b)" + "norm (w - k) < e / 2" + using B1(2)[OF *(1)] by blast + obtain z where z: + "((\<lambda>x. if x \<in> s then g x else 0) has_integral z) (cbox a b)" + "norm (z - l) < e / 2" + using B2(2)[OF *(2)] by blast + have *: "\<And>x. (if x \<in> s then f x + g x else 0) = + (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0)" + by auto + show "\<exists>z. ((\<lambda>x. if x \<in> s then f x + g x else 0) has_integral z) (cbox a b) \<and> norm (z - (k + l)) < e" + apply (rule_tac x="w + z" in exI) + apply (simp add: lem[OF w(1) z(1), unfolded *[symmetric]]) + using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) + apply (auto simp add: field_simps) + done + qed + qed +qed + +lemma has_integral_sub: + "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow> + ((\<lambda>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 (\<lambda>x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" + by (rule integral_unique has_integral_0)+ + +lemma integral_add: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> + integral s (\<lambda>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 (\<lambda>x. c *\<^sub>R f x) = c *\<^sub>R integral s f" +proof (cases "f integrable_on s \<or> c = 0") + case True with has_integral_cmul show ?thesis by force +next + case False then have "~ (\<lambda>x. c *\<^sub>R f x) integrable_on s" + using has_integral_cmul [of "(\<lambda>x. c *\<^sub>R f x)" _ s "inverse c"] + by force + with False show ?thesis by (simp add: not_integrable_integral) +qed + +lemma integral_neg [simp]: "integral s (\<lambda>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 "~ (\<lambda>x. - f x) integrable_on s" + using has_integral_neg [of "(\<lambda>x. - f x)" _ s ] + by force + with False show ?thesis by (simp add: not_integrable_integral) +qed + +lemma integral_diff: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> + integral s (\<lambda>x. f x - g x) = integral s f - integral s g" + by (rule integral_unique) (metis integrable_integral has_integral_sub) + +lemma integrable_0: "(\<lambda>x. 0) integrable_on s" + unfolding integrable_on_def using has_integral_0 by auto + +lemma integrable_add: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>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 \<Longrightarrow> (\<lambda>x. c *\<^sub>R f(x)) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_cmul) + +lemma integrable_on_cmult_iff: + fixes c :: real + assumes "c \<noteq> 0" + shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> f integrable_on s" + using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] \<open>c \<noteq> 0\<close> + by auto + +lemma integrable_on_cmult_left: + assumes "f integrable_on s" + shows "(\<lambda>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 \<Longrightarrow> (\<lambda>x. -f(x)) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_neg) + +lemma integrable_diff: + "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x - g x) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_sub) + +lemma integrable_linear: + "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h \<circ> f) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_linear) + +lemma integral_linear: + "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> integral s (h \<circ> f) = h (integral s f)" + apply (rule has_integral_unique [where i=s and f = "h \<circ> f"]) + apply (simp_all add: integrable_integral integrable_linear has_integral_linear ) + done + +lemma integral_component_eq[simp]: + fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" + assumes "f integrable_on s" + shows "integral s (\<lambda>x. f x \<bullet> k) = integral s f \<bullet> k" + unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] .. + +lemma has_integral_setsum: + assumes "finite t" + and "\<forall>a\<in>t. ((f a) has_integral (i a)) s" + shows "((\<lambda>x. setsum (\<lambda>a. f a x) t) has_integral (setsum 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_setsum: + "\<lbrakk>finite t; \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow> + integral s (\<lambda>x. setsum (\<lambda>a. f a x) t) = setsum (\<lambda>a. integral s (f a)) t" + by (auto intro: has_integral_setsum integrable_integral) + +lemma integrable_setsum: + "\<lbrakk>finite t; \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) t) integrable_on s" + unfolding integrable_on_def + apply (drule bchoice) + using has_integral_setsum[of t] + apply auto + done + +lemma has_integral_eq: + assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x" + and "(f has_integral k) s" + shows "(g has_integral k) s" + using has_integral_sub[OF assms(2), of "\<lambda>x. f x - g x" 0] + using has_integral_is_0[of s "\<lambda>x. f x - g x"] + using assms(1) + by auto + +lemma integrable_eq: "(\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f integrable_on s \<Longrightarrow> 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 "\<And>x. x \<in> s \<Longrightarrow> 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 "\<And>x. x \<in> s \<Longrightarrow> 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 \<noteq> 0" + shows "(\<lambda>x. of_real c * f x) integrable_on s \<longleftrightarrow> f integrable_on s" + (is "?lhs = ?rhs") +proof + assume ?lhs + then have "(\<lambda>x. of_real (1 / c) * (of_real c * f x)) integrable_on s" + using integrable_cmul[of "\<lambda>x. of_real c * f x" s "1 / of_real c"] + by (simp add: scaleR_conv_of_real) + then have "(\<lambda>x. (of_real (1 / c) * of_real c * f x)) integrable_on s" + by (simp add: algebra_simps) + with \<open>c \<noteq> 0\<close> 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 :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}" + assumes "f integrable_on s" + shows "(\<lambda>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 :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}" + assumes "c \<noteq> 0" + shows "(\<lambda>x. f x * of_real c) integrable_on s \<longleftrightarrow> f integrable_on s" +using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) + +lemma integrable_on_cdivide: + fixes f :: "_ \<Rightarrow> 'b :: real_normed_field" + assumes "f integrable_on s" + shows "(\<lambda>x. f x / of_real c) integrable_on s" +by (simp add: integrable_on_cmult_right divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) + +lemma integrable_on_cdivide_iff [simp]: + fixes f :: "_ \<Rightarrow> 'b :: real_normed_field" + assumes "c \<noteq> 0" + shows "(\<lambda>x. f x / of_real c) integrable_on s \<longleftrightarrow> f integrable_on s" +by (simp add: divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) + +lemma has_integral_null [intro]: + assumes "content(cbox a b) = 0" + shows "(f has_integral 0) (cbox a b)" +proof - + have "gauge (\<lambda>x. ball x 1)" + by auto + moreover + { + fix e :: real + fix p + assume e: "e > 0" + assume p: "p tagged_division_of (cbox a b)" + have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) = 0" + unfolding norm_eq_zero diff_0_right + using setsum_content_null[OF assms(1) p, of f] . + then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e" + using e by auto + } + ultimately show ?thesis + by (auto simp: has_integral) +qed + +lemma has_integral_null_real [intro]: + assumes "content {a .. b::real} = 0" + shows "(f has_integral 0) {a .. b}" + by (metis assms box_real(2) has_integral_null) + +lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \<Longrightarrow> (f has_integral i) (cbox a b) \<longleftrightarrow> i = 0" + by (auto simp add: has_integral_null dest!: integral_unique) + +lemma integral_null [simp]: "content (cbox a b) = 0 \<Longrightarrow> integral (cbox a b) f = 0" + by (metis has_integral_null integral_unique) + +lemma integrable_on_null [intro]: "content (cbox a b) = 0 \<Longrightarrow> f integrable_on (cbox a b)" + by (simp add: has_integral_integrable) + +lemma has_integral_empty[intro]: "(f has_integral 0) {}" + by (simp add: has_integral_is_0) + +lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \<longleftrightarrow> 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 - + have *: "{a} = cbox a a" + apply (rule set_eqI) + unfolding mem_box singleton_iff euclidean_eq_iff[where 'a='a] + apply safe + prefer 3 + apply (erule_tac x=b in ballE) + apply (auto simp add: field_simps) + done + show "(f has_integral 0) (cbox a a)" "(f has_integral 0) {a}" + unfolding * + apply (rule_tac[!] has_integral_null) + unfolding content_eq_0_interior + unfolding interior_cbox + using box_sing + apply auto + done +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 (\<lambda>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 (\<lambda>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 \<Rightarrow> 'b :: euclidean_space" + shows "(f has_integral y) A \<longleftrightarrow> (\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)" +proof safe + fix b :: 'b assume "(f has_integral y) A" + from has_integral_linear[OF this(1) bounded_linear_component, of b] + show "((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A" by (simp add: o_def) +next + assume "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)" + hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral ((y \<bullet> b) *\<^sub>R b)) A" + by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) + hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. (y \<bullet> b) *\<^sub>R b)) A" + by (intro has_integral_setsum) (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 \<Rightarrow> 'b :: euclidean_space" + shows "(\<And>b. b \<in> Basis \<Longrightarrow> ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A) \<Longrightarrow> (f has_integral y) A" + by (subst has_integral_componentwise_iff) blast + +lemma integrable_componentwise_iff: + fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" + shows "f integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> 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 "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)" + by (subst (asm) has_integral_componentwise_iff) + thus "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)" by (auto simp: integrable_on_def) +next + assume "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)" + then obtain y where "\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral y b) A" + unfolding integrable_on_def by (subst (asm) bchoice_iff) blast + hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral (y b *\<^sub>R b)) A" + by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) + hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. y b *\<^sub>R b)) A" + by (intro has_integral_setsum) (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 \<Rightarrow> 'b :: euclidean_space" + shows "(\<And>b. b \<in> Basis \<Longrightarrow> (\<lambda>x. f x \<bullet> b) integrable_on A) \<Longrightarrow> f integrable_on A" + by (subst integrable_componentwise_iff) blast + +lemma integral_componentwise: + fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" + assumes "f integrable_on A" + shows "integral A f = (\<Sum>b\<in>Basis. integral A (\<lambda>x. (f x \<bullet> b) *\<^sub>R b))" +proof - + from assms have integrable: "\<forall>b\<in>Basis. (\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. (f x \<bullet> 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 (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b)" + by (simp add: euclidean_representation) + also from integrable have "\<dots> = (\<Sum>a\<in>Basis. integral A (\<lambda>x. (f x \<bullet> a) *\<^sub>R a))" + by (subst integral_setsum) (simp_all add: o_def) + finally show ?thesis . +qed + +lemma integrable_component: + "f integrable_on A \<Longrightarrow> (\<lambda>x. f x \<bullet> (y :: 'b :: euclidean_space)) integrable_on A" + by (drule integrable_linear[OF _ bounded_linear_component[of y]]) (simp add: o_def) + + + +subsection \<open>Cauchy-type criterion for integrability.\<close> + +(* XXXXXXX *) +lemma integrable_cauchy: + fixes f :: "'n::euclidean_space \<Rightarrow> 'a::{real_normed_vector,complete_space}" + shows "f integrable_on cbox a b \<longleftrightarrow> + (\<forall>e>0.\<exists>d. gauge d \<and> + (\<forall>p1 p2. p1 tagged_division_of (cbox a b) \<and> d fine p1 \<and> + p2 tagged_division_of (cbox a b) \<and> d fine p2 \<longrightarrow> + norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 - + setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) < e))" + (is "?l = (\<forall>e>0. \<exists>d. ?P e d)") +proof + assume ?l + then guess y unfolding integrable_on_def has_integral .. note y=this + show "\<forall>e>0. \<exists>d. ?P e d" + proof (clarify, goal_cases) + case (1 e) + then have "e/2 > 0" by auto + then guess d + apply - + apply (drule y[rule_format]) + apply (elim exE conjE) + done + note d=this[rule_format] + show ?case + proof (rule_tac x=d in exI, clarsimp simp: d) + fix p1 p2 + assume as: "p1 tagged_division_of (cbox a b)" "d fine p1" + "p2 tagged_division_of (cbox a b)" "d fine p2" + show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e" + apply (rule dist_triangle_half_l[where y=y,unfolded dist_norm]) + using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] . + qed + qed +next + assume "\<forall>e>0. \<exists>d. ?P e d" + then have "\<forall>n::nat. \<exists>d. ?P (inverse(of_nat (n + 1))) d" + by auto + from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format] + have "\<And>n. gauge (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}})" + apply (rule gauge_inters) + using d(1) + apply auto + done + then have "\<forall>n. \<exists>p. p tagged_division_of (cbox a b) \<and> (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}}) fine p" + by (meson fine_division_exists) + from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]] + have dp: "\<And>i n. i\<le>n \<Longrightarrow> d i fine p n" + using p(2) unfolding fine_inters by auto + have "Cauchy (\<lambda>n. setsum (\<lambda>(x,k). content k *\<^sub>R (f x)) (p n))" + proof (rule CauchyI, goal_cases) + case (1 e) + then guess N unfolding real_arch_inverse[of e] .. note N=this + show ?case + apply (rule_tac x=N in exI) + proof clarify + fix m n + assume mn: "N \<le> m" "N \<le> n" + have *: "N = (N - 1) + 1" using N by auto + show "norm ((\<Sum>(x, k)\<in>p m. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p n. content k *\<^sub>R f x)) < e" + apply (rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) + apply(subst *) + using dp p(1) mn d(2) by auto + qed + qed + then guess y unfolding convergent_eq_cauchy[symmetric] .. note y=this[THEN 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 *:"e/2 > 0" by auto + then guess N1 unfolding real_arch_inverse[of "e/2"] .. note N1=this + then have N1': "N1 = N1 - 1 + 1" + by auto + guess N2 using y[OF *] .. note N2=this + have "gauge (d (N1 + N2))" + using d by auto + moreover + { + fix q + assume as: "q tagged_division_of (cbox a b)" "d (N1 + N2) fine q" + have *: "inverse (of_nat (N1 + N2 + 1)) < e / 2" + apply (rule less_trans) + using N1 + apply auto + done + have "norm ((\<Sum>(x, k)\<in>q. content k *\<^sub>R f x) - y) < e" + apply (rule norm_triangle_half_r) + apply (rule less_trans[OF _ *]) + apply (subst N1', rule d(2)[of "p (N1+N2)"]) + using N1' as(1) as(2) dp + apply (simp add: \<open>\<forall>x. p x tagged_division_of cbox a b \<and> (\<lambda>xa. \<Inter>{d i xa |i. i \<in> {0..x}}) fine p x\<close>) + using N2 le_add2 by blast + } + ultimately show "\<exists>d. gauge d \<and> + (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> + norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e)" + by (rule_tac x="d (N1 + N2)" in exI) auto + qed +qed + + +subsection \<open>Additivity of integral on abutting intervals.\<close> + +lemma tagged_division_split_left_inj: + fixes x1 :: "'a::euclidean_space" + assumes d: "d tagged_division_of i" + and k12: "(x1, k1) \<in> d" + "(x2, k2) \<in> d" + "k1 \<noteq> k2" + "k1 \<inter> {x. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}" + "k \<in> Basis" + shows "content (k1 \<inter> {x. x\<bullet>k \<le> c}) = 0" +proof - + have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" + by force + show ?thesis + using k12 + by (fastforce intro!: division_split_left_inj[OF division_of_tagged_division[OF d]] *) +qed + +lemma tagged_division_split_right_inj: + fixes x1 :: "'a::euclidean_space" + assumes d: "d tagged_division_of i" + and k12: "(x1, k1) \<in> d" + "(x2, k2) \<in> d" + "k1 \<noteq> k2" + "k1 \<inter> {x. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}" + "k \<in> Basis" + shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0" +proof - + have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" + by force + show ?thesis + using k12 + by (fastforce intro!: division_split_right_inj[OF division_of_tagged_division[OF d]] *) +qed + +lemma has_integral_split: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes fi: "(f has_integral i) (cbox a b \<inter> {x. x\<bullet>k \<le> c})" + and fj: "(f has_integral j) (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" + and k: "k \<in> Basis" + shows "(f has_integral (i + j)) (cbox a b)" +proof (unfold has_integral, rule, rule, goal_cases) + case (1 e) + then have e: "e/2 > 0" + by auto + obtain d1 + where d1: "gauge d1" + and d1norm: + "\<And>p. \<lbrakk>p tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c}; + d1 fine p\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x, k) \<in> p. 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 d2 + where d2: "gauge d2" + and d2norm: + "\<And>p. \<lbrakk>p tagged_division_of cbox a b \<inter> {x. c \<le> x \<bullet> k}; + d2 fine p\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x, k) \<in> p. 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 ?d = "\<lambda>x. if x\<bullet>k = c then (d1 x \<inter> d2 x) else ball x \<bar>x\<bullet>k - c\<bar> \<inter> d1 x \<inter> d2 x" + have "gauge ?d" + using d1 d2 unfolding gauge_def by auto + then show ?case + proof (rule_tac x="?d" in exI, safe) + fix p + assume "p tagged_division_of (cbox a b)" "?d fine p" + note p = this tagged_division_ofD[OF this(1)] + have xk_le_c: "\<And>x kk. (x, kk) \<in> p \<Longrightarrow> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {} \<Longrightarrow> x\<bullet>k \<le> c" + proof - + fix x kk + assume as: "(x, kk) \<in> p" and kk: "kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}" + show "x\<bullet>k \<le> c" + proof (rule ccontr) + assume **: "\<not> ?thesis" + from this[unfolded not_le] + have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>" + using p(2)[unfolded fine_def, rule_format,OF as] by auto + with kk obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<le> c" + by blast + then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>" + 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 + qed + have xk_ge_c: "\<And>x kk. (x, kk) \<in> p \<Longrightarrow> kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {} \<Longrightarrow> x\<bullet>k \<ge> c" + proof - + fix x kk + assume as: "(x, kk) \<in> p" and kk: "kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}" + show "x\<bullet>k \<ge> c" + proof (rule ccontr) + assume **: "\<not> ?thesis" + from this[unfolded not_le] have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>" + using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto + with kk obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<ge> c" + by blast + then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>" + 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 + qed + + have lem1: "\<And>f P Q. (\<forall>x k. (x, k) \<in> {(x, f k) | x k. P x k} \<longrightarrow> Q x k) \<longleftrightarrow> + (\<forall>x k. P x k \<longrightarrow> Q x (f k))" + by auto + have fin_finite: "finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}" if "finite s" for f s P + proof - + from that have "finite ((\<lambda>(x, k). (x, f k)) ` s)" + by auto + then show ?thesis + by (rule rev_finite_subset) auto + qed + { fix g :: "'a set \<Rightarrow> 'a set" + fix i :: "'a \<times> 'a set" + assume "i \<in> (\<lambda>(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}" + then obtain x k where xk: + "i = (x, g k)" "(x, k) \<in> p" + "(x, g k) \<notin> {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}" + by auto + have "content (g k) = 0" + using xk using content_empty by auto + then have "(\<lambda>(x, k). content k *\<^sub>R f x) i = 0" + unfolding xk split_conv by auto + } note [simp] = this + have lem3: "\<And>g :: 'a set \<Rightarrow> 'a set. finite p \<Longrightarrow> + setsum (\<lambda>(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> g k \<noteq> {}} = + setsum (\<lambda>(x, k). content k *\<^sub>R f x) ((\<lambda>(x, k). (x, g k)) ` p)" + by (rule setsum.mono_neutral_left) auto + let ?M1 = "{(x, kk \<inter> {x. x\<bullet>k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}" + have d1_fine: "d1 fine ?M1" + by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm) + have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2" + proof (rule d1norm [OF tagged_division_ofI d1_fine]) + show "finite ?M1" + by (rule fin_finite p(3))+ + show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = cbox a b \<inter> {x. x\<bullet>k \<le> c}" + unfolding p(8)[symmetric] by auto + fix x l + assume xl: "(x, l) \<in> ?M1" + then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this + show "x \<in> l" "l \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<le> c}" + unfolding xl' + using p(4-6)[OF xl'(3)] using xl'(4) + using xk_le_c[OF xl'(3-4)] by auto + show "\<exists>a b. l = cbox a b" + unfolding xl' + using p(6)[OF xl'(3)] + by (fastforce simp add: interval_split[OF k,where c=c]) + fix y r + let ?goal = "interior l \<inter> interior r = {}" + assume yr: "(y, r) \<in> ?M1" + then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this + assume as: "(x, l) \<noteq> (y, r)" + show "interior l \<inter> interior r = {}" + proof (cases "l' = r' \<longrightarrow> x' = y'") + case False + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + next + case True + then have "l' \<noteq> r'" + using as unfolding xl' yr' by auto + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + qed + qed + moreover + let ?M2 = "{(x,kk \<inter> {x. x\<bullet>k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}" + have d2_fine: "d2 fine ?M2" + by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm) + have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2" + proof (rule d2norm [OF tagged_division_ofI d2_fine]) + show "finite ?M2" + by (rule fin_finite p(3))+ + show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = cbox a b \<inter> {x. x\<bullet>k \<ge> c}" + unfolding p(8)[symmetric] by auto + fix x l + assume xl: "(x, l) \<in> ?M2" + then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this + show "x \<in> l" "l \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<ge> c}" + unfolding xl' + using p(4-6)[OF xl'(3)] xl'(4) xk_ge_c[OF xl'(3-4)] + by auto + show "\<exists>a b. l = cbox a b" + unfolding xl' + using p(6)[OF xl'(3)] + by (fastforce simp add: interval_split[OF k, where c=c]) + fix y r + let ?goal = "interior l \<inter> interior r = {}" + assume yr: "(y, r) \<in> ?M2" + then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this + assume as: "(x, l) \<noteq> (y, r)" + show "interior l \<inter> interior r = {}" + proof (cases "l' = r' \<longrightarrow> x' = y'") + case False + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + next + case True + then have "l' \<noteq> r'" + using as unfolding xl' yr' by auto + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + qed + qed + ultimately + have "norm (((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2" + using norm_add_less by blast + also { + have eq0: "\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'b) = 0" + using scaleR_zero_left by auto + have cont_eq: "\<And>g. (\<lambda>(x,l). content l *\<^sub>R f x) \<circ> (\<lambda>(x,l). (x,g l)) = (\<lambda>(x,l). content (g l) *\<^sub>R f x)" + by auto + have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) = + (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)" + by auto + also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) + + (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)" + unfolding lem3[OF p(3)] + by (subst setsum.reindex_nontrivial[OF p(3)], auto intro!: k eq0 tagged_division_split_left_inj[OF p(1)] tagged_division_split_right_inj[OF p(1)] + simp: cont_eq)+ + also note setsum.distrib[symmetric] + also have "\<And>x. x \<in> p \<Longrightarrow> + (\<lambda>(x,ka). content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) x + + (\<lambda>(x,ka). content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) x = + (\<lambda>(x,ka). content ka *\<^sub>R f x) x" + proof clarify + fix a b + assume "(a, b) \<in> p" + from p(6)[OF this] guess u v by (elim exE) note uv=this + then show "content (b \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f a = + content b *\<^sub>R f a" + unfolding scaleR_left_distrib[symmetric] + unfolding uv content_split[OF k,of u v c] + by auto + qed + note setsum.cong [OF _ this] + finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x \<bullet> k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i + + ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x \<bullet> k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}}. content k *\<^sub>R f x) - j) = + (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)" + by auto + } + finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e" + by auto + qed +qed + + +subsection \<open>A sort of converse, integrability on subintervals.\<close> + +lemma tagged_division_union_interval: + fixes a :: "'a::euclidean_space" + assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})" + and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" + and k: "k \<in> Basis" + shows "(p1 \<union> p2) tagged_division_of (cbox a b)" +proof - + have *: "cbox a b = (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<union> (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" + by auto + show ?thesis + apply (subst *) + apply (rule tagged_division_union[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_union_interval_real: + fixes a :: real + assumes "p1 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<le> (c::real)})" + and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})" + and k: "k \<in> Basis" + shows "(p1 \<union> p2) tagged_division_of {a .. b}" + using assms + unfolding box_real[symmetric] + by (rule tagged_division_union_interval) + +lemma has_integral_separate_sides: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "(f has_integral i) (cbox a b)" + and "e > 0" + and k: "k \<in> Basis" + obtains d where "gauge d" + "\<forall>p1 p2. p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<and> d fine p1 \<and> + p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) \<and> d fine p2 \<longrightarrow> + norm ((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 + setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e" +proof - + guess d using has_integralD[OF assms(1-2)] . note d=this + { fix p1 p2 + assume "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p1" + note p1=tagged_division_ofD[OF this(1)] this + assume "p2 tagged_division_of (cbox a b) \<inter> {x. c \<le> x \<bullet> k}" "d fine p2" + note p2=tagged_division_ofD[OF this(1)] this + note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this + { fix a b + assume ab: "(a, b) \<in> p1 \<inter> p2" + have "(a, b) \<in> p1" + using ab by auto + with p1 obtain u v where uv: "b = cbox u v" by auto + have "b \<subseteq> {x. x\<bullet>k = c}" + using ab p1(3)[of a b] p2(3)[of a b] by fastforce + moreover + have "interior {x::'a. x \<bullet> k = c} = {}" + proof (rule ccontr) + assume "\<not> ?thesis" + then obtain x where x: "x \<in> interior {x::'a. x\<bullet>k = c}" + by auto + then guess e unfolding mem_interior .. note e=this + have x: "x\<bullet>k = c" + using x interior_subset by fastforce + have *: "\<And>i. i \<in> Basis \<Longrightarrow> \<bar>(x - (x + (e / 2) *\<^sub>R k)) \<bullet> i\<bar> = (if i = k then e/2 else 0)" + using e k by (auto simp: inner_simps inner_not_same_Basis) + have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (e / 2 ) *\<^sub>R k)) \<bullet> i\<bar>) = + (\<Sum>i\<in>Basis. (if i = k then e / 2 else 0))" + using "*" by (blast intro: setsum.cong) + also have "\<dots> < e" + apply (subst setsum.delta) + using e + apply auto + done + finally have "x + (e/2) *\<^sub>R k \<in> ball x e" + unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) + then have "x + (e/2) *\<^sub>R k \<in> {x. x\<bullet>k = c}" + using e by auto + then show False + unfolding mem_Collect_eq using e 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 ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) = + norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)" + by (subst setsum.union_inter_neutral) (auto simp: p1 p2) + also have "\<dots> < e" + by (rule k d(2) p12 fine_union p1 p2)+ + finally have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" . + } + then show ?thesis + by (auto intro: that[of d] d elim: ) +qed + +lemma integrable_split[intro]: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}" + assumes "f integrable_on cbox a b" + and k: "k \<in> Basis" + shows "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<le> c})" (is ?t1) + and "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" (is ?t2) +proof - + guess y using assms(1) unfolding integrable_on_def .. note y=this + define b' where "b' = (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i)" + define a' where "a' = (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i)*\<^sub>R i)" + show ?t1 ?t2 + unfolding interval_split[OF k] integrable_cauchy + unfolding interval_split[symmetric,OF k] + proof (rule_tac[!] allI impI)+ + fix e :: real + assume "e > 0" + then have "e/2>0" + by auto + from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format] + let ?P = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of (cbox a b) \<inter> A \<and> d fine p1 \<and> + p2 tagged_division_of (cbox a b) \<inter> A \<and> d fine p2 \<longrightarrow> + norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)" + show "?P {x. x \<bullet> k \<le> c}" + proof (rule_tac x=d in exI, clarsimp simp add: d) + fix p1 p2 + assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p1" + "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p2" + show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e" + proof (rule fine_division_exists[OF d(1), 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(2)[of p1 p] d(2)[of p2 p]] + unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] + by (auto simp add: algebra_simps) + qed + qed + show "?P {x. x \<bullet> k \<ge> c}" + proof (rule_tac x=d in exI, clarsimp simp add: d) + fix p1 p2 + assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p1" + "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p2" + show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e" + proof (rule fine_division_exists[OF d(1), 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(2)[of p p1] d(2)[of p p2]] + unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] + by (auto simp add: algebra_simps) + qed + qed + qed +qed + +lemma operative_integral: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" + shows "comm_monoid.operative (lift_option op +) (Some 0) (\<lambda>i. if f integrable_on i then Some (integral i f) else None)" + unfolding comm_monoid.operative_def[OF add.comm_monoid_lift_option] +proof safe + fix a b c + fix k :: 'a + assume k: "k \<in> Basis" + show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = + lift_option op + (if f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c} then Some (integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f) else None) + (if f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k} then Some (integral (cbox a b \<inter> {x. c \<le> x \<bullet> k}) f) else None)" + proof (cases "f integrable_on cbox a b") + case True + with k show ?thesis + apply (simp add: integrable_split) + apply (rule integral_unique [OF has_integral_split[OF _ _ k]]) + apply (auto intro: integrable_integral) + done + next + case False + have "\<not> (f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<or> \<not> ( f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k})" + proof (rule ccontr) + assume "\<not> ?thesis" + then have "f integrable_on cbox a b" + unfolding integrable_on_def + apply (rule_tac x="integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f + integral (cbox a b \<inter> {x. x \<bullet> k \<ge> c}) f" in exI) + apply (rule has_integral_split[OF _ _ k]) + apply (auto intro: integrable_integral) + done + then show False + using False by auto + qed + then show ?thesis + using False by auto + qed +next + fix a b :: 'a + assume "content (cbox a b) = 0" + 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) +qed + +subsection \<open>Finally, the integral of a constant\<close> + +lemma has_integral_const [intro]: + fixes a b :: "'a::euclidean_space" + shows "((\<lambda>x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" + apply (auto intro!: exI [where x="\<lambda>x. ball x 1"] simp: split_def has_integral) + apply (subst scaleR_left.setsum[symmetric, unfolded o_def]) + apply (subst additive_content_tagged_division[unfolded split_def]) + apply auto + done + +lemma has_integral_const_real [intro]: + fixes a b :: real + shows "((\<lambda>x. c) has_integral (content {a .. b} *\<^sub>R c)) {a .. b}" + by (metis box_real(2) has_integral_const) + +lemma integral_const [simp]: + fixes a b :: "'a::euclidean_space" + shows "integral (cbox a b) (\<lambda>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} (\<lambda>x. c) = content {a .. b} *\<^sub>R c" + by (metis box_real(2) integral_const) + + +subsection \<open>Bounds on the norm of Riemann sums and the integral itself.\<close> + +lemma dsum_bound: + assumes "p division_of (cbox a b)" + and "norm c \<le> e" + shows "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content(cbox a b)" +proof - + have sumeq: "(\<Sum>i\<in>p. \<bar>content i\<bar>) = setsum content p" + apply (rule setsum.cong) + using assms + apply simp + apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4)) + done + have e: "0 \<le> e" + using assms(2) norm_ge_zero order_trans by blast + have "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))" + using norm_setsum by blast + also have "... \<le> e * (\<Sum>i\<in>p. \<bar>content i\<bar>)" + apply (simp add: setsum_right_distrib[symmetric] mult.commute) + using assms(2) mult_right_mono by blast + also have "... \<le> e * content (cbox a b)" + apply (rule mult_left_mono [OF _ e]) + apply (simp add: sumeq) + using additive_content_division assms(1) eq_iff apply blast + done + finally show ?thesis . +qed + +lemma rsum_bound: + assumes p: "p tagged_division_of (cbox a b)" + and "\<forall>x\<in>cbox a b. norm (f x) \<le> e" + shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> 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 \<ge> 0" + by (meson ex_in_conv assms(2) norm_ge_zero order_trans) + have setsum_le: "setsum (content \<circ> snd) p \<le> content (cbox a b)" + unfolding additive_content_tagged_division[OF p, symmetric] split_def + by (auto intro: eq_refl) + have con: "\<And>xk. xk \<in> p \<Longrightarrow> 0 \<le> content (snd xk)" + using tagged_division_ofD(4) [OF p] content_pos_le + by force + have norm: "\<And>xk. xk \<in> p \<Longrightarrow> norm (f (fst xk)) \<le> e" + unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms + by (metis prod.collapse subset_eq) + have "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> (\<Sum>i\<in>p. norm (case i of (x, k) \<Rightarrow> content k *\<^sub>R f x))" + by (rule norm_setsum) + also have "... \<le> e * content (cbox a b)" + unfolding split_def norm_scaleR + apply (rule order_trans[OF setsum_mono]) + apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) + apply (metis norm) + unfolding setsum_left_distrib[symmetric] + using con setsum_le + apply (auto simp: mult.commute intro: mult_left_mono [OF _ e]) + done + finally show ?thesis . +qed + +lemma rsum_diff_bound: + assumes "p tagged_division_of (cbox a b)" + and "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" + shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - setsum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le> + e * content (cbox a b)" + apply (rule order_trans[OF _ rsum_bound[OF assms]]) + apply (simp add: split_def scaleR_diff_right setsum_subtractf eq_refl) + done + +lemma has_integral_bound: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "0 \<le> B" + and "(f has_integral i) (cbox a b)" + and "\<forall>x\<in>cbox a b. norm (f x) \<le> B" + shows "norm i \<le> B * content (cbox a b)" +proof (rule ccontr) + assume "\<not> ?thesis" + then have *: "norm i - B * content (cbox a b) > 0" + by auto + from assms(2)[unfolded has_integral,rule_format,OF *] + guess d by (elim exE conjE) note d=this[rule_format] + from fine_division_exists[OF this(1), of a b] guess p . note p=this + have *: "\<And>s B. norm s \<le> B \<Longrightarrow> \<not> norm (s - i) < norm i - B" + unfolding not_less + by (metis norm_triangle_sub[of i] add.commute le_less_trans less_diff_eq linorder_not_le norm_minus_commute) + show False + using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto +qed + +corollary has_integral_bound_real: + fixes f :: "real \<Rightarrow> 'b::real_normed_vector" + assumes "0 \<le> B" + and "(f has_integral i) {a .. b}" + and "\<forall>x\<in>{a .. b}. norm (f x) \<le> B" + shows "norm i \<le> B * content {a .. b}" + by (metis assms box_real(2) has_integral_bound) + +corollary integrable_bound: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" + assumes "0 \<le> B" + and "f integrable_on (cbox a b)" + and "\<And>x. x\<in>cbox a b \<Longrightarrow> norm (f x) \<le> B" + shows "norm (integral (cbox a b) f) \<le> B * content (cbox a b)" +by (metis integrable_integral has_integral_bound assms) + + +subsection \<open>Similar theorems about relationship among components.\<close> + +lemma rsum_component_le: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "p tagged_division_of (cbox a b)" + and "\<forall>x\<in>cbox a b. (f x)\<bullet>i \<le> (g x)\<bullet>i" + shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)\<bullet>i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)\<bullet>i" +unfolding inner_setsum_left +proof (rule setsum_mono, clarify) + fix a b + assume ab: "(a, b) \<in> p" + note tagged = tagged_division_ofD(2-4)[OF assms(1) ab] + from this(3) guess u v by (elim exE) note b=this + show "(content b *\<^sub>R f a) \<bullet> i \<le> (content b *\<^sub>R g a) \<bullet> i" + unfolding b inner_simps real_scaleR_def + apply (rule mult_left_mono) + using assms(2) tagged + by (auto simp add: content_pos_le) +qed + +lemma has_integral_component_le: + fixes f g :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes k: "k \<in> Basis" + assumes "(f has_integral i) s" "(g has_integral j) s" + and "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k" + shows "i\<bullet>k \<le> j\<bullet>k" +proof - + have lem: "i\<bullet>k \<le> j\<bullet>k" + if f_i: "(f has_integral i) (cbox a b)" + and g_j: "(g has_integral j) (cbox a b)" + and le: "\<forall>x\<in>cbox a b. (f x)\<bullet>k \<le> (g x)\<bullet>k" + for a b i and j :: 'b and f g :: "'a \<Rightarrow> 'b" + proof (rule ccontr) + assume "\<not> ?thesis" + then have *: "0 < (i\<bullet>k - j\<bullet>k) / 3" + by auto + guess d1 using f_i[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d1=this[rule_format] + guess d2 using g_j[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d2=this[rule_format] + obtain p where p: "p tagged_division_of cbox a b" "d1 fine p" "d2 fine p" + using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter + by metis + note le_less_trans[OF Basis_le_norm[OF k]] + then have "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3" + "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - j) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3" + using k norm_bound_Basis_lt d1 d2 p + by blast+ + then show False + unfolding inner_simps + using rsum_component_le[OF p(1) le] + by (simp add: abs_real_def split: if_split_asm) + qed + show ?thesis + proof (cases "\<exists>a b. s = cbox a b") + case True + with lem assms show ?thesis + by auto + next + case False + show ?thesis + proof (rule ccontr) + assume "\<not> i\<bullet>k \<le> j\<bullet>k" + then have ij: "(i\<bullet>k - j\<bullet>k) / 3 > 0" + by auto + note has_integral_altD[OF _ False this] + from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format] + have "bounded (ball 0 B1 \<union> ball (0::'a) B2)" + unfolding bounded_Un by(rule conjI bounded_ball)+ + from bounded_subset_cbox[OF this] guess a b by (elim exE) + note ab = conjunctD2[OF this[unfolded Un_subset_iff]] + guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this] + guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this] + have *: "\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False" + by (simp add: abs_real_def split: if_split_asm) + note le_less_trans[OF Basis_le_norm[OF k]] + note this[OF w1(2)] this[OF w2(2)] + moreover + have "w1\<bullet>k \<le> w2\<bullet>k" + by (rule lem[OF w1(1) w2(1)]) (simp add: assms(4)) + ultimately show False + unfolding inner_simps by(rule *) + qed + qed +qed + +lemma integral_component_le: + fixes g f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "k \<in> Basis" + and "f integrable_on s" "g integrable_on s" + and "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k" + shows "(integral s f)\<bullet>k \<le> (integral s g)\<bullet>k" + apply (rule has_integral_component_le) + using integrable_integral assms + apply auto + done + +lemma has_integral_component_nonneg: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "k \<in> Basis" + and "(f has_integral i) s" + and "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k" + shows "0 \<le> i\<bullet>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 \<Rightarrow> 'b::euclidean_space" + assumes "k \<in> Basis" + and "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k" + shows "0 \<le> (integral s f)\<bullet>k" +proof (cases "f integrable_on s") + case True show ?thesis + apply (rule has_integral_component_nonneg) + using assms True + apply auto + done +next + case False then show ?thesis by (simp add: not_integrable_integral) +qed + +lemma has_integral_component_neg: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "k \<in> Basis" + and "(f has_integral i) s" + and "\<forall>x\<in>s. (f x)\<bullet>k \<le> 0" + shows "i\<bullet>k \<le> 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 \<Rightarrow> 'b::euclidean_space" + assumes "(f has_integral i) (cbox a b)" + and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k" + and "k \<in> Basis" + shows "B * content (cbox a b) \<le> i\<bullet>k" + using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\<Sum>i\<in>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 "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B" + and "k \<in> Basis" + shows "i\<bullet>k \<le> B * content (cbox a b)" + using has_integral_component_le[OF assms(3,1) has_integral_const, of "\<Sum>i\<in>Basis. B *\<^sub>R i"] assms(2-) + by (auto simp add: field_simps) + +lemma integral_component_lbound: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "f integrable_on cbox a b" + and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k" + and "k \<in> Basis" + shows "B * content (cbox a b) \<le> (integral(cbox a b) f)\<bullet>k" + apply (rule has_integral_component_lbound) + using assms + unfolding has_integral_integral + apply auto + done + +lemma integral_component_lbound_real: + assumes "f integrable_on {a ::real .. b}" + and "\<forall>x\<in>{a .. b}. B \<le> f(x)\<bullet>k" + and "k \<in> Basis" + shows "B * content {a .. b} \<le> (integral {a .. b} f)\<bullet>k" + using assms + by (metis box_real(2) integral_component_lbound) + +lemma integral_component_ubound: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" + assumes "f integrable_on cbox a b" + and "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B" + and "k \<in> Basis" + shows "(integral (cbox a b) f)\<bullet>k \<le> B * content (cbox a b)" + apply (rule has_integral_component_ubound) + using assms + unfolding has_integral_integral + apply auto + done + +lemma integral_component_ubound_real: + fixes f :: "real \<Rightarrow> 'a::euclidean_space" + assumes "f integrable_on {a .. b}" + and "\<forall>x\<in>{a .. b}. f x\<bullet>k \<le> B" + and "k \<in> Basis" + shows "(integral {a .. b} f)\<bullet>k \<le> B * content {a .. b}" + using assms + by (metis box_real(2) integral_component_ubound) + +subsection \<open>Uniform limit of integrable functions is integrable.\<close> + +lemma real_arch_invD: + "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)" + by (subst(asm) real_arch_inverse) + +lemma integrable_uniform_limit: + fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" + assumes "\<forall>e>0. \<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> 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 *: "\<And>P. \<forall>e>(0::real). P e \<Longrightarrow> \<forall>n::nat. P (inverse (real n + 1))" + by auto + from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format] + from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\<lambda>x. x"]] + obtain i where i: "\<And>x. (g x has_integral i x) (cbox a b)" + by auto + have "Cauchy i" + 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 add: field_simps) + then obtain M :: nat + where M: "M \<noteq> 0" "0 < inverse (real_of_nat M)" "inverse (of_nat M) < e / 4 / content (cbox a b)" + by (subst (asm) real_arch_inverse) auto + show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (i m) (i n) < e" + proof (rule exI [where x=M], clarify) + fix m n + assume m: "M \<le> m" and n: "M \<le> n" + have "e/4>0" using \<open>e>0\<close> by auto + note * = i[unfolded has_integral,rule_format,OF this] + from *[of m] guess gm by (elim conjE exE) note gm=this[rule_format] + from *[of n] guess gn by (elim conjE exE) note gn=this[rule_format] + from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b] + obtain p where p: "p tagged_division_of cbox a b" "(\<lambda>x. gm x \<inter> gn x) fine p" + by auto + { fix s1 s2 i1 and i2::'b + assume no: "norm(s2 - s1) \<le> e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" + have "norm (i1 - i2) \<le> 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 add: algebra_simps) + also have "\<dots> < e" + using no + unfolding norm_minus_commute + by (auto simp add: algebra_simps) + finally have "norm (i1 - i2) < e" . + } note triangle3 = this + have finep: "gm fine p" "gn fine p" + using fine_inter p by auto + { fix x + assume x: "x \<in> cbox a b" + have "norm (f x - g n x) + norm (f x - g m x) \<le> inverse (real n + 1) + inverse (real m + 1)" + using g(1)[OF x, of n] g(1)[OF x, of m] by auto + also have "\<dots> \<le> inverse (real M) + inverse (real M)" + apply (rule add_mono) + using M(2) m n by auto + also have "\<dots> = 2 / real M" + unfolding divide_inverse by auto + finally have "norm (g n x - g m x) \<le> 2 / real M" + using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] + by (auto simp add: algebra_simps simp add: norm_minus_commute) + } note norm_le = this + have le_e2: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g n x) - (\<Sum>(x, k)\<in>p. content k *\<^sub>R g m x)) \<le> e / 2" + apply (rule order_trans [OF rsum_diff_bound[OF p(1), where e="2 / real M"]]) + apply (blast intro: norm_le) + using M True + by (auto simp add: field_simps) + then show "dist (i m) (i n) < e" + unfolding dist_norm + using gm gn p finep + by (auto intro!: triangle3) + qed + qed + then obtain s where s: "i \<longlonglongrightarrow> 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: "\<forall>n\<ge>N1. norm (i 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 add: field_simps) + from real_arch_invD[OF this] guess N2 by (elim exE conjE) note N2=this + from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format] + { fix sf sg i + assume no: "norm (sf - sg) \<le> e / 3" + "norm(i - s) < e / 3" + "norm (sg - i) < e / 3" + have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)" + using norm_triangle_ineq[of "sf - sg" "sg - s"] + using norm_triangle_ineq[of "sg - i" " i - s"] + by (auto simp add: algebra_simps) + also have "\<dots> < e" + using no + unfolding norm_minus_commute + by (auto simp add: algebra_simps) + finally have "norm (sf - s) < e" . + } note lem = this + { fix p + assume p: "p tagged_division_of (cbox a b) \<and> g' fine p" + then have norm_less: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g (N1 + N2) x) - i (N1 + N2)) < e / 3" + using g' by blast + have "content (cbox a b) < e / 3 * (of_nat N2)" + using N2 unfolding inverse_eq_divide using True by (auto simp add: field_simps) + moreover have "e / 3 * of_nat N2 \<le> e / 3 * (of_nat (N1 + N2) + 1)" + using \<open>e>0\<close> by auto + ultimately have "content (cbox a b) < e / 3 * (of_nat (N1 + N2) + 1)" + by linarith + then have le_e3: "inverse (real (N1 + N2) + 1) * content (cbox a b) \<le> e / 3" + unfolding inverse_eq_divide + by (auto simp add: field_simps) + have ne3: "norm (i (N1 + N2) - s) < e / 3" + using N1 by auto + have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e" + apply (rule lem[OF order_trans [OF _ le_e3] ne3 norm_less]) + apply (rule rsum_diff_bound[OF p[THEN conjunct1]]) + apply (blast intro: g) + done } + then show "\<exists>d. gauge d \<and> + (\<forall>p. p tagged_division_of cbox a b \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e)" + by (blast intro: g') + qed +qed + +lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] + + +subsection \<open>Negligible sets.\<close> + +definition "negligible (s:: 'a::euclidean_space set) \<longleftrightarrow> + (\<forall>a b. ((indicator s :: 'a\<Rightarrow>real) has_integral 0) (cbox a b))" + + +subsection \<open>Negligibility of hyperplane.\<close> + +lemma interval_doublesplit: + fixes a :: "'a::euclidean_space" + assumes "k \<in> Basis" + shows "cbox a b \<inter> {x . \<bar>x\<bullet>k - c\<bar> \<le> (e::real)} = + cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) (c - e) else a\<bullet>i) *\<^sub>R i) + (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) (c + e) else b\<bullet>i) *\<^sub>R i)" +proof - + have *: "\<And>x c e::real. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e" + by auto + have **: "\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {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 \<in> Basis" + shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}} + division_of (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})" +proof - + have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e" + by auto + have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> 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]) + apply (rule equalityI) + apply blast + apply clarsimp + apply (rule_tac x="l \<inter> {x. c + e \<ge> x \<bullet> k}" in exI) + apply auto + done + by (simp add: interval_split k interval_doublesplit) +qed + +lemma content_doublesplit: + fixes a :: "'a::euclidean_space" + assumes "0 < e" + and k: "k \<in> Basis" + obtains d where "0 < d" and "content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) < e" +proof (cases "content (cbox a b) = 0") + case True + then have ce: "content (cbox a b) < e" + by (metis \<open>0 < e\<close>) + show ?thesis + apply (rule that[of 1]) + apply simp + unfolding interval_doublesplit[OF k] + apply (rule le_less_trans[OF content_subset ce]) + apply (auto simp: interval_doublesplit[symmetric] k) + done +next + case False + define d where "d = e / 3 / setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})" + note False[unfolded content_eq_0 not_ex not_le, rule_format] + then have "\<And>x. x \<in> Basis \<Longrightarrow> b\<bullet>x > a\<bullet>x" + by (auto simp add:not_le) + then have prod0: "0 < setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})" + by (force simp add: setprod_pos field_simps) + then have "d > 0" + using assms + by (auto simp add: d_def field_simps) + then show ?thesis + proof (rule that[of d]) + have *: "Basis = insert k (Basis - {k})" + using k by auto + have less_e: "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) * (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i) < e" + proof - + have "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) \<le> 2 * d" + by auto + also have "\<dots> < e / (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i)" + unfolding d_def + using assms prod0 + by (auto simp add: field_simps) + finally show "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) * (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i) < e" + unfolding pos_less_divide_eq[OF prod0] . + qed + show "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e" + proof (cases "cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = {}") + case True + then show ?thesis + using assms by simp + next + case False + then have + "(\<Prod>i\<in>Basis - {k}. interval_upperbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i - + interval_lowerbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i) + = (\<Prod>i\<in>Basis - {k}. b\<bullet>i - a\<bullet>i)" + by (simp add: box_eq_empty interval_doublesplit[OF k]) + then show "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e" + unfolding content_def + using assms False + apply (subst *) + apply (subst setprod.insert) + apply (simp_all add: interval_doublesplit[OF k] box_eq_empty not_less less_e) + done + qed + qed +qed + +lemma negligible_standard_hyperplane[intro]: + fixes k :: "'a::euclidean_space" + assumes k: "k \<in> Basis" + shows "negligible {x. x\<bullet>k = c}" + unfolding negligible_def has_integral +proof (clarify, goal_cases) + case (1 a b e) + from this and k obtain d where d: "0 < d" "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e" + by (rule content_doublesplit) + let ?i = "indicator {x::'a. x\<bullet>k = c} :: 'a\<Rightarrow>real" + show ?case + apply (rule_tac x="\<lambda>x. ball x d" in exI) + apply rule + apply (rule gauge_ball) + apply (rule d) + proof (rule, rule) + fix p + assume p: "p tagged_division_of (cbox a b) \<and> (\<lambda>x. ball x d) fine p" + have *: "(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = + (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) *\<^sub>R ?i x)" + apply (rule setsum.cong) + apply (rule refl) + unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv + apply cases + apply (rule disjI1) + apply assumption + apply (rule disjI2) + proof - + fix x l + assume as: "(x, l) \<in> p" "?i x \<noteq> 0" + then have xk: "x\<bullet>k = c" + unfolding indicator_def + apply - + apply (rule ccontr) + apply auto + done + show "content l = content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})" + apply (rule arg_cong[where f=content]) + apply (rule set_eqI) + apply rule + apply rule + unfolding mem_Collect_eq + proof - + fix y + assume y: "y \<in> l" + note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv] + note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] + note le_less_trans[OF Basis_le_norm[OF k] this] + then show "\<bar>y \<bullet> k - c\<bar> \<le> d" + unfolding inner_simps xk by auto + qed auto + qed + note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]] + show "norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) - 0) < e" + unfolding diff_0_right * + unfolding real_scaleR_def real_norm_def + apply (subst abs_of_nonneg) + apply (rule setsum_nonneg) + apply rule + unfolding split_paired_all split_conv + apply (rule mult_nonneg_nonneg) + apply (drule p'(4)) + apply (erule exE)+ + apply(rule_tac b=b in back_subst) + prefer 2 + apply (subst(asm) eq_commute) + apply assumption + apply (subst interval_doublesplit[OF k]) + apply (rule content_pos_le) + apply (rule indicator_pos_le) + proof - + have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) \<le> + (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}))" + apply (rule setsum_mono) + unfolding split_paired_all split_conv + apply (rule mult_right_le_one_le) + apply (drule p'(4)) + apply (auto simp add:interval_doublesplit[OF k]) + done + also have "\<dots> < e" + proof (subst setsum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases) + case prems: (1 u v) + have "content (cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<le> content (cbox u v)" + unfolding interval_doublesplit[OF k] + apply (rule content_subset) + unfolding interval_doublesplit[symmetric,OF k] + apply auto + done + then show ?case + unfolding prems interval_doublesplit[OF k] + by (blast intro: antisym) + next + have "(\<Sum>l\<in>snd ` p. content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) = + setsum content ((\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}})" + proof (subst (2) setsum.reindex_nontrivial) + fix x y assume "x \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}" "y \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}" + "x \<noteq> y" and eq: "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}" + then obtain x' y' where "(x', x) \<in> p" "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}" "(y', y) \<in> p" "y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}" + by (auto) + from p'(5)[OF \<open>(x', x) \<in> p\<close> \<open>(y', y) \<in> p\<close>] \<open>x \<noteq> y\<close> have "interior (x \<inter> y) = {}" + by auto + moreover have "interior ((x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<inter> (y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) \<subseteq> interior (x \<inter> y)" + by (auto intro: interior_mono) + ultimately have "interior (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = {}" + by (auto simp: eq) + then show "content (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0" + using p'(4)[OF \<open>(x', x) \<in> p\<close>] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) + qed (insert p'(1), auto intro!: setsum.mono_neutral_right) + also have "\<dots> \<le> norm (\<Sum>l\<in>(\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}. content l *\<^sub>R 1::real)" + by simp + also have "\<dots> \<le> 1 * content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})" + using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] + unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto + also have "\<dots> < e" + using d(2) by simp + finally show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) < e" . + qed + finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) < e" . + qed + qed +qed + + +subsection \<open>A technical lemma about "refinement" of division.\<close> + +lemma tagged_division_finer: + fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set" + assumes "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 "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q" +proof - + let ?P = "\<lambda>p. p tagged_partial_division_of (cbox a b) \<longrightarrow> gauge d \<longrightarrow> + (\<exists>q. q tagged_division_of (\<Union>{k. \<exists>x. (x,k) \<in> p}) \<and> d fine q \<and> + (\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q))" + { + have *: "finite p" "p tagged_partial_division_of (cbox a b)" + using assms(1) + unfolding tagged_division_of_def + by auto + presume "\<And>p. finite p \<Longrightarrow> ?P p" + from this[rule_format,OF * assms(2)] guess q .. note q=this + then show ?thesis + apply - + apply (rule that[of q]) + unfolding tagged_division_ofD[OF assms(1)] + apply auto + done + } + fix p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set" + assume as: "finite p" + show "?P p" + apply rule + apply rule + using as + proof (induct p) + case empty + show ?case + apply (rule_tac x="{}" in exI) + unfolding fine_def + apply auto + done + next + case (insert xk p) + guess x k using surj_pair[of xk] by (elim exE) note xk=this + note tagged_partial_division_subset[OF insert(4) subset_insertI] + from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this] + have *: "\<Union>{l. \<exists>y. (y,l) \<in> insert xk p} = k \<union> \<Union>{l. \<exists>y. (y,l) \<in> p}" + unfolding xk by auto + note p = tagged_partial_division_ofD[OF insert(4)] + from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this + + have "finite {k. \<exists>x. (x, k) \<in> p}" + apply (rule finite_subset[of _ "snd ` p"]) + using p + apply safe + apply (metis image_iff snd_conv) + apply auto + done + then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}" + apply (rule inter_interior_unions_intervals) + apply (rule open_interior) + apply (rule_tac[!] ballI) + unfolding mem_Collect_eq + apply (erule_tac[!] exE) + apply (drule p(4)[OF insertI2]) + apply assumption + apply (rule p(5)) + unfolding uv xk + apply (rule insertI1) + apply (rule insertI2) + apply assumption + using insert(2) + unfolding uv xk + apply auto + done + show ?case + proof (cases "cbox u v \<subseteq> d x") + case True + then show ?thesis + apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI) + apply rule + unfolding * uv + apply (rule tagged_division_union) + apply (rule tagged_division_of_self) + apply (rule p[unfolded xk uv] insertI1)+ + apply (rule q1) + apply (rule int) + apply rule + apply (rule fine_union) + apply (subst fine_def) + defer + apply (rule q1) + unfolding Ball_def split_paired_All split_conv + apply rule + apply rule + apply rule + apply rule + apply (erule insertE) + apply (simp add: uv xk) + apply (rule UnI2) + apply (drule q1(3)[rule_format]) + unfolding xk uv + apply auto + done + next + case False + from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this + show ?thesis + apply (rule_tac x="q2 \<union> q1" in exI) + apply rule + unfolding * uv + apply (rule tagged_division_union q2 q1 int fine_union)+ + unfolding Ball_def split_paired_All split_conv + apply rule + apply (rule fine_union) + apply (rule q1 q2)+ + apply rule + apply rule + apply rule + apply rule + apply (erule insertE) + apply (rule UnI2) + apply (simp add: False uv xk) + apply (drule q1(3)[rule_format]) + using False + unfolding xk uv + apply auto + done + qed + qed +qed + + +subsection \<open>Hence the main theorem about negligible sets.\<close> + +lemma finite_product_dependent: + assumes "finite s" + and "\<And>x. x \<in> s \<Longrightarrow> finite (t x)" + shows "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}" + using assms +proof induct + case (insert x s) + have *: "{(i, j) |i j. i \<in> insert x s \<and> j \<in> t i} = + (\<lambda>y. (x,y)) ` (t x) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto + show ?case + unfolding * + apply (rule finite_UnI) + using insert + apply auto + done +qed auto + +lemma sum_sum_product: + assumes "finite s" + and "\<forall>i\<in>s. finite (t i)" + shows "setsum (\<lambda>i. setsum (x i) (t i)::real) s = + setsum (\<lambda>(i,j). x i j) {(i,j) | i j. i \<in> s \<and> j \<in> t i}" + using assms +proof induct + case (insert a s) + have *: "{(i, j) |i j. i \<in> insert a s \<and> j \<in> t i} = + (\<lambda>y. (a,y)) ` (t a) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto + show ?case + unfolding * + apply (subst setsum.union_disjoint) + unfolding setsum.insert[OF insert(1-2)] + prefer 4 + apply (subst insert(3)) + unfolding add_right_cancel + proof - + show "setsum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)" + apply (subst setsum.reindex) + unfolding inj_on_def + apply auto + done + show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}" + apply (rule finite_product_dependent) + using insert + apply auto + done + qed (insert insert, auto) +qed auto + +lemma has_integral_negligible: + fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector" + assumes "negligible s" + and "\<forall>x\<in>(t - s). f x = 0" + shows "(f has_integral 0) t" +proof - + presume P: "\<And>f::'b::euclidean_space \<Rightarrow> 'a. + \<And>a b. \<forall>x. x \<notin> s \<longrightarrow> f x = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)" + let ?f = "(\<lambda>x. if x \<in> t then f x else 0)" + show ?thesis + apply (rule_tac f="?f" in has_integral_eq) + unfolding if_P + apply (rule refl) + apply (subst has_integral_alt) + apply cases + apply (subst if_P, assumption) + unfolding if_not_P + proof - + assume "\<exists>a b. t = cbox a b" + then guess a b apply - by (erule exE)+ note t = this + show "(?f has_integral 0) t" + unfolding t + apply (rule P) + using assms(2) + unfolding t + apply auto + done + next + show "\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> + (\<exists>z. ((\<lambda>x. if x \<in> t then ?f x else 0) has_integral z) (cbox a b) \<and> norm (z - 0) < e)" + apply safe + apply (rule_tac x=1 in exI) + apply rule + apply (rule zero_less_one) + apply safe + apply (rule_tac x=0 in exI) + apply rule + apply (rule P) + using assms(2) + apply auto + done + qed +next + fix f :: "'b \<Rightarrow> 'a" + fix a b :: 'b + assume assm: "\<forall>x. x \<notin> s \<longrightarrow> f x = 0" + show "(f has_integral 0) (cbox a b)" + unfolding has_integral + proof (safe, goal_cases) + case prems: (1 e) + then have "\<And>n. e / 2 / ((real n+1) * (2 ^ n)) > 0" + apply - + apply (rule divide_pos_pos) + defer + apply (rule mult_pos_pos) + apply (auto simp add:field_simps) + done + note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b] + note allI[OF this,of "\<lambda>x. x"] + from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]] + show ?case + apply (rule_tac x="\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x" in exI) + proof safe + show "gauge (\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x)" + using d(1) unfolding gauge_def by auto + fix p + assume as: "p tagged_division_of (cbox a b)" "(\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x) fine p" + let ?goal = "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e" + { + presume "p \<noteq> {} \<Longrightarrow> ?goal" + then show ?goal + apply (cases "p = {}") + using prems + apply auto + done + } + assume as': "p \<noteq> {}" + from real_arch_simple[of "Max((\<lambda>(x,k). norm(f x)) ` p)"] guess N .. + then have N: "\<forall>x\<in>(\<lambda>(x, k). norm (f x)) ` p. x \<le> real N" + by (meson Max_ge as(1) dual_order.trans finite_imageI tagged_division_of_finite) + have "\<forall>i. \<exists>q. q tagged_division_of (cbox a b) \<and> (d i) fine q \<and> (\<forall>(x, k)\<in>p. k \<subseteq> (d i) x \<longrightarrow> (x, k) \<in> q)" + by (auto intro: tagged_division_finer[OF as(1) d(1)]) + from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]] + have *: "\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> (0::real)" + apply (rule setsum_nonneg) + apply safe + unfolding real_scaleR_def + apply (drule tagged_division_ofD(4)[OF q(1)]) + apply (auto intro: mult_nonneg_nonneg) + done + have **: "finite s \<Longrightarrow> finite t \<Longrightarrow> (\<forall>(x,y) \<in> t. (0::real) \<le> g(x,y)) \<Longrightarrow> + (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> setsum f s \<le> setsum g t" for f g s t + apply (rule setsum_le_included[of s t g snd f]) + prefer 4 + apply safe + apply (erule_tac x=x in ballE) + apply (erule exE) + apply (rule_tac x="(xa,x)" in bexI) + apply auto + done + have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) * + norm (setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}" + unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right + apply (rule order_trans) + apply (rule norm_setsum) + apply (subst sum_sum_product) + prefer 3 + proof (rule **, safe) + show "finite {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i}" + apply (rule finite_product_dependent) + using q + apply auto + done + fix i a b + assume as'': "(a, b) \<in> q i" + show "0 \<le> (real i + 1) * (content b *\<^sub>R indicator s a)" + unfolding real_scaleR_def + using tagged_division_ofD(4)[OF q(1) as''] + by (auto intro!: mult_nonneg_nonneg) + next + fix i :: nat + show "finite (q i)" + using q by auto + next + fix x k + assume xk: "(x, k) \<in> p" + define n where "n = nat \<lfloor>norm (f x)\<rfloor>" + have *: "norm (f x) \<in> (\<lambda>(x, k). norm (f x)) ` p" + using xk by auto + have nfx: "real n \<le> norm (f x)" "norm (f x) \<le> real n + 1" + unfolding n_def by auto + then have "n \<in> {0..N + 1}" + using N[rule_format,OF *] by auto + moreover + note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv] + note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this] + note this[unfolded n_def[symmetric]] + moreover + have "norm (content k *\<^sub>R f x) \<le> (real n + 1) * (content k * indicator s x)" + proof (cases "x \<in> s") + case False + then show ?thesis + using assm by auto + next + case True + have *: "content k \<ge> 0" + using tagged_division_ofD(4)[OF as(1) xk] by auto + moreover + have "content k * norm (f x) \<le> content k * (real n + 1)" + apply (rule mult_mono) + using nfx * + apply auto + done + ultimately + show ?thesis + unfolding abs_mult + using nfx True + by (auto simp add: field_simps) + qed + ultimately show "\<exists>y. (y, x, k) \<in> {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i} \<and> norm (content k *\<^sub>R f x) \<le> + (real y + 1) * (content k *\<^sub>R indicator s x)" + apply (rule_tac x=n in exI) + apply safe + apply (rule_tac x=n in exI) + apply (rule_tac x="(x,k)" in exI) + apply safe + apply auto + done + qed (insert as, auto) + also have "\<dots> \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {..N+1}" + proof (rule setsum_mono, goal_cases) + case (1 i) + then show ?case + apply (subst mult.commute, subst pos_le_divide_eq[symmetric]) + using d(2)[rule_format, of "q i" i] + using q[rule_format] + apply (auto simp add: field_simps) + done + qed + also have "\<dots> < e * inverse 2 * 2" + unfolding divide_inverse setsum_right_distrib[symmetric] + apply (rule mult_strict_left_mono) + unfolding power_inverse [symmetric] lessThan_Suc_atMost[symmetric] + apply (subst geometric_sum) + using prems + apply auto + done + finally show "?goal" by auto + qed + qed +qed + +lemma has_integral_spike: + fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector" + assumes "negligible s" + and "(\<forall>x\<in>(t - s). g x = f x)" + and "(f has_integral y) t" + shows "(g has_integral y) t" +proof - + { + fix a b :: 'b + fix f g :: "'b \<Rightarrow> 'a" + fix y :: 'a + assume as: "\<forall>x \<in> cbox a b - s. g x = f x" "(f has_integral y) (cbox a b)" + have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" + apply (rule has_integral_add[OF as(2)]) + apply (rule has_integral_negligible[OF assms(1)]) + using as + apply auto + done + then have "(g has_integral y) (cbox a b)" + by auto + } note * = this + show ?thesis + apply (subst has_integral_alt) + using assms(2-) + apply - + apply (rule cond_cases) + apply safe + apply (rule *) + apply assumption+ + apply (subst(asm) has_integral_alt) + unfolding if_not_P + apply (erule_tac x=e in allE) + apply safe + apply (rule_tac x=B in exI) + apply safe + apply (erule_tac x=a in allE) + apply (erule_tac x=b in allE) + apply safe + apply (rule_tac x=z in exI) + apply safe + apply (rule *[where fa2="\<lambda>x. if x\<in>t then f x else 0"]) + apply auto + done +qed + +lemma has_integral_spike_eq: + assumes "negligible s" + and "\<forall>x\<in>(t - s). g x = f x" + shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)" + apply rule + apply (rule_tac[!] has_integral_spike[OF assms(1)]) + using assms(2) + apply auto + done + +lemma integrable_spike: + assumes "negligible s" + and "\<forall>x\<in>(t - s). g x = f x" + and "f integrable_on t" + shows "g integrable_on t" + using assms + unfolding integrable_on_def + apply - + apply (erule exE) + apply rule + apply (rule has_integral_spike) + apply fastforce+ + done + +lemma integral_spike: + assumes "negligible s" + and "\<forall>x\<in>(t - s). g x = f x" + shows "integral t f = integral t g" + using has_integral_spike_eq[OF assms] by (simp add: integral_def integrable_on_def) + + +subsection \<open>Some other trivialities about negligible sets.\<close> + +lemma negligible_subset[intro]: + assumes "negligible s" + and "t \<subseteq> s" + shows "negligible t" + unfolding negligible_def +proof (safe, goal_cases) + case (1 a b) + show ?case + using assms(1)[unfolded negligible_def,rule_format,of a b] + apply - + apply (rule has_integral_spike[OF assms(1)]) + defer + apply assumption + using assms(2) + unfolding indicator_def + apply auto + done +qed + +lemma negligible_diff[intro?]: + assumes "negligible s" + shows "negligible (s - t)" + using assms by auto + +lemma negligible_Int: + assumes "negligible s \<or> negligible t" + shows "negligible (s \<inter> t)" + using assms by auto + +lemma negligible_Un: + assumes "negligible s" + and "negligible t" + shows "negligible (s \<union> t)" + unfolding negligible_def +proof (safe, goal_cases) + case (1 a b) + note assm = assms[unfolded negligible_def,rule_format,of a b] + then show ?case + apply (subst has_integral_spike_eq[OF assms(2)]) + defer + apply assumption + unfolding indicator_def + apply auto + done +qed + +lemma negligible_Un_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> negligible s \<and> negligible t" + using negligible_Un by auto + +lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" + using negligible_standard_hyperplane[OF SOME_Basis, of "a \<bullet> (SOME i. i \<in> Basis)"] by auto + +lemma negligible_insert[simp]: "negligible (insert a s) \<longleftrightarrow> negligible s" + apply (subst insert_is_Un) + unfolding negligible_Un_eq + apply auto + done + +lemma negligible_empty[iff]: "negligible {}" + by auto + +lemma negligible_finite[intro]: + assumes "finite s" + shows "negligible s" + using assms by (induct s) auto + +lemma negligible_Union[intro]: + assumes "finite s" + and "\<forall>t\<in>s. negligible t" + shows "negligible(\<Union>s)" + using assms by induct auto + +lemma negligible: + "negligible s \<longleftrightarrow> (\<forall>t::('a::euclidean_space) set. ((indicator s::'a\<Rightarrow>real) has_integral 0) t)" + apply safe + defer + apply (subst negligible_def) +proof - + fix t :: "'a set" + assume as: "negligible s" + have *: "(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)" + by auto + show "((indicator s::'a\<Rightarrow>real) has_integral 0) t" + apply (subst has_integral_alt) + apply cases + apply (subst if_P,assumption) + unfolding if_not_P + apply safe + apply (rule as[unfolded negligible_def,rule_format]) + apply (rule_tac x=1 in exI) + apply safe + apply (rule zero_less_one) + apply (rule_tac x=0 in exI) + using negligible_subset[OF as,of "s \<inter> t"] + unfolding negligible_def indicator_def [abs_def] + unfolding * + apply auto + done +qed auto + + +subsection \<open>Finite case of the spike theorem is quite commonly needed.\<close> + +lemma has_integral_spike_finite: + assumes "finite s" + and "\<forall>x\<in>t-s. g x = f x" + and "(f has_integral y) t" + shows "(g has_integral y) t" + apply (rule has_integral_spike) + using assms + apply auto + done + +lemma has_integral_spike_finite_eq: + assumes "finite s" + and "\<forall>x\<in>t-s. g x = f x" + shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)" + apply rule + apply (rule_tac[!] has_integral_spike_finite) + using assms + apply auto + done + +lemma integrable_spike_finite: + assumes "finite s" + and "\<forall>x\<in>t-s. g x = f x" + and "f integrable_on t" + shows "g integrable_on t" + using assms + unfolding integrable_on_def + apply safe + apply (rule_tac x=y in exI) + apply (rule has_integral_spike_finite) + apply auto + done + + +subsection \<open>In particular, the boundary of an interval is negligible.\<close> + +lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" +proof - + let ?A = "\<Union>((\<lambda>k. {x. x\<bullet>k = a\<bullet>k} \<union> {x::'a. x\<bullet>k = b\<bullet>k}) ` Basis)" + have "cbox a b - box a b \<subseteq> ?A" + apply rule unfolding Diff_iff mem_box + apply simp + apply(erule conjE bexE)+ + apply(rule_tac x=i in bexI) + apply auto + done + then show ?thesis + apply - + apply (rule negligible_subset[of ?A]) + apply (rule negligible_Union[OF finite_imageI]) + apply auto + done +qed + +lemma has_integral_spike_interior: + assumes "\<forall>x\<in>box a b. g x = f x" + and "(f has_integral y) (cbox a b)" + shows "(g has_integral y) (cbox a b)" + apply (rule has_integral_spike[OF negligible_frontier_interval _ assms(2)]) + using assms(1) + apply auto + done + +lemma has_integral_spike_interior_eq: + assumes "\<forall>x\<in>box a b. g x = f x" + shows "(f has_integral y) (cbox a b) \<longleftrightarrow> (g has_integral y) (cbox a b)" + apply rule + apply (rule_tac[!] has_integral_spike_interior) + using assms + apply auto + done + +lemma integrable_spike_interior: + assumes "\<forall>x\<in>box a b. g x = f x" + and "f integrable_on cbox a b" + shows "g integrable_on cbox a b" + using assms + unfolding integrable_on_def + using has_integral_spike_interior[OF assms(1)] + by auto + + +subsection \<open>Integrability of continuous functions.\<close> + +lemma operative_approximable: + fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" + assumes "0 \<le> e" + shows "comm_monoid.operative op \<and> True (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i)" + unfolding comm_monoid.operative_def[OF comm_monoid_and] +proof safe + fix a b :: 'b + show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" + if "content (cbox a b) = 0" + apply (rule_tac x=f in exI) + using assms that + apply (auto intro!: integrable_on_null) + done + { + fix c g + fix k :: 'b + assume as: "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" "g integrable_on cbox a b" + assume k: "k \<in> Basis" + show "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" + "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}" + apply (rule_tac[!] x=g in exI) + using as(1) integrable_split[OF as(2) k] + apply auto + done + } + fix c k g1 g2 + assume as: "\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" + "\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g2 x) \<le> e" "g2 integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}" + assume k: "k \<in> Basis" + let ?g = "\<lambda>x. if x\<bullet>k = c then f x else if x\<bullet>k \<le> c then g1 x else g2 x" + show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" + apply (rule_tac x="?g" in exI) + apply safe + proof goal_cases + case (1 x) + then show ?case + apply - + apply (cases "x\<bullet>k=c") + apply (case_tac "x\<bullet>k < c") + using as assms + apply auto + done + next + case 2 + presume "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" + and "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}" + then guess h1 h2 unfolding integrable_on_def by auto + from has_integral_split[OF this k] show ?case + unfolding integrable_on_def by auto + next + show "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}" + apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) + using k as(2,4) + apply auto + done + qed +qed + +lemma comm_monoid_set_F_and: "comm_monoid_set.F op \<and> True f s \<longleftrightarrow> (finite s \<longrightarrow> (\<forall>x\<in>s. f x))" +proof - + interpret bool: comm_monoid_set "op \<and>" True + proof qed auto + show ?thesis + by (induction s rule: infinite_finite_induct) auto +qed + +lemma approximable_on_division: + fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" + assumes "0 \<le> e" + and "d division_of (cbox a b)" + and "\<forall>i\<in>d. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i" + obtains g where "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" "g integrable_on cbox a b" +proof - + note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_approximable[OF assms(1)] assms(2)] + from assms(3) this[unfolded comm_monoid_set_F_and, of f] division_of_finite[OF assms(2)] + guess g by auto + then show thesis + apply - + apply (rule that[of g]) + apply auto + done +qed + +lemma integrable_continuous: + fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" + assumes "continuous_on (cbox a b) f" + shows "f integrable_on cbox a b" +proof (rule integrable_uniform_limit, safe) + fix e :: real + assume e: "e > 0" + from compact_uniformly_continuous[OF assms compact_cbox,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d .. + note d=conjunctD2[OF this,rule_format] + from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this + note p' = tagged_division_ofD[OF p(1)] + have *: "\<forall>i\<in>snd ` p. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i" + proof (safe, unfold snd_conv) + fix x l + assume as: "(x, l) \<in> p" + from p'(4)[OF this] guess a b by (elim exE) note l=this + show "\<exists>g. (\<forall>x\<in>l. norm (f x - g x) \<le> e) \<and> g integrable_on l" + apply (rule_tac x="\<lambda>y. f x" in exI) + proof safe + show "(\<lambda>y. f x) integrable_on l" + unfolding integrable_on_def l + apply rule + apply (rule has_integral_const) + done + fix y + assume y: "y \<in> l" + note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this] + note d(2)[OF _ _ this[unfolded mem_ball]] + then show "norm (f y - f x) \<le> e" + using y p'(2-3)[OF as] unfolding dist_norm l norm_minus_commute by fastforce + qed + qed + from e have "e \<ge> 0" + by auto + from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g . + then show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" + by auto +qed + +lemma integrable_continuous_real: + fixes f :: "real \<Rightarrow> 'a::banach" + assumes "continuous_on {a .. b} f" + shows "f integrable_on {a .. b}" + by (metis assms box_real(2) integrable_continuous) + +subsection \<open>Specialization of additivity to one dimension.\<close> + +subsection \<open>Special case of additivity we need for the FTC.\<close> + +lemma additive_tagged_division_1: + fixes f :: "real \<Rightarrow> 'a::real_normed_vector" + assumes "a \<le> b" + and "p tagged_division_of {a..b}" + shows "setsum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a" +proof - + let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" + have ***: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" + using assms by auto + have *: "add.operative ?f" + unfolding add.operative_1_lt box_eq_empty + by auto + have **: "cbox a b \<noteq> {}" + using assms(1) by auto + note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]] + note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric] + show ?thesis + unfolding * + apply (rule setsum.cong) + unfolding split_paired_all split_conv + using assms(2) + apply auto + done +qed + + +subsection \<open>A useful lemma allowing us to factor out the content size.\<close> + +lemma has_integral_factor_content: + "(f has_integral i) (cbox a b) \<longleftrightarrow> + (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> + norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content (cbox a b)))" +proof (cases "content (cbox a b) = 0") + case True + show ?thesis + unfolding has_integral_null_eq[OF True] + apply safe + apply (rule, rule, rule gauge_trivial, safe) + unfolding setsum_content_null[OF True] True + defer + apply (erule_tac x=1 in allE) + apply safe + defer + apply (rule fine_division_exists[of _ a b]) + apply assumption + apply (erule_tac x=p in allE) + unfolding setsum_content_null[OF True] + apply auto + done +next + case False + note F = this[unfolded content_lt_nz[symmetric]] + let ?P = "\<lambda>e opp. \<exists>d. gauge d \<and> + (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> opp (norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i)) e)" + show ?thesis + apply (subst has_integral) + proof safe + fix e :: real + assume e: "e > 0" + { + assume "\<forall>e>0. ?P e op <" + then show "?P (e * content (cbox a b)) op \<le>" + apply (erule_tac x="e * content (cbox a b)" in allE) + apply (erule impE) + defer + apply (erule exE,rule_tac x=d in exI) + using F e + apply (auto simp add:field_simps) + done + } + { + assume "\<forall>e>0. ?P (e * content (cbox a b)) op \<le>" + then show "?P e op <" + apply (erule_tac x="e / 2 / content (cbox a b)" in allE) + apply (erule impE) + defer + apply (erule exE,rule_tac x=d in exI) + using F e + apply (auto simp add: field_simps) + done + } + qed +qed + +lemma has_integral_factor_content_real: + "(f has_integral i) {a .. b::real} \<longleftrightarrow> + (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow> + norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a .. b} ))" + unfolding box_real[symmetric] + by (rule has_integral_factor_content) + + +subsection \<open>Fundamental theorem of calculus.\<close> + +lemma interval_bounds_real: + fixes q b :: real + assumes "a \<le> b" + shows "Sup {a..b} = b" + and "Inf {a..b} = a" + using assms by auto + +lemma fundamental_theorem_of_calculus: + fixes f :: "real \<Rightarrow> 'a::banach" + assumes "a \<le> b" + and "\<forall>x\<in>{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: "e > 0" + note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt] + have *: "\<And>P Q. \<forall>x\<in>{a .. b}. P x \<and> (\<forall>e>0. \<exists>d>0. Q x e d) \<Longrightarrow> \<forall>x. \<exists>(d::real)>0. x\<in>{a .. b} \<longrightarrow> Q x e d" + using e by blast + note this[OF assm,unfolded gauge_existence_lemma] + from choice[OF this,unfolded Ball_def[symmetric]] guess d .. + note d=conjunctD2[OF this[rule_format],rule_format] + show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> + norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b))" + apply (rule_tac x="\<lambda>x. ball x (d x)" in exI) + apply safe + apply (rule gauge_ball_dependent) + apply rule + apply (rule d(1)) + proof - + fix p + assume as: "p tagged_division_of cbox a b" "(\<lambda>x. ball x (d x)) fine p" + show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b)" + unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric] + unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\<lambda>x. x",symmetric] + unfolding setsum_right_distrib + defer + unfolding setsum_subtractf[symmetric] + proof (rule setsum_norm_le,safe) + fix x k + assume "(x, k) \<in> p" + note xk = tagged_division_ofD(2-4)[OF as(1) this] + from this(3) guess u v by (elim exE) note k=this + have *: "u \<le> v" + using xk unfolding k by auto + have ball: "\<forall>xa\<in>k. xa \<in> ball x (d x)" + using as(2)[unfolded fine_def,rule_format,OF \<open>(x,k)\<in>p\<close>,unfolded split_conv subset_eq] . + have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \<le> + norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" + apply (rule order_trans[OF _ norm_triangle_ineq4]) + apply (rule eq_refl) + apply (rule arg_cong[where f=norm])